From 638269af0d6149dccc5058bf76c09d80c0ad9fe1 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 3 Jul 2024 16:24:30 -0700 Subject: [PATCH 01/89] pts_to_serialized, validator, jumper, peek, CPS reader, read --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 192 +++++++++++++++++++++ src/lowparse/pulse/Makefile | 13 ++ 2 files changed, 205 insertions(+) create mode 100644 src/lowparse/pulse/LowParse.Pulse.Base.fst create mode 100644 src/lowparse/pulse/Makefile diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst new file mode 100644 index 000000000..567da813a --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -0,0 +1,192 @@ +module LowParse.Pulse.Base +open FStar.Tactics.V2 +open Pulse.Lib.Pervasives +open Pulse.Lib.Slice +open LowParse.Spec.Base + +module SZ = FStar.SizeT + +let parser = tot_parser +let serializer #k = tot_serializer #k + +let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : vprop = + pts_to input #pm (bare_serialize s v) + +let validator_post (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (res: option SZ.t) : Tot prop = + SZ.v offset <= Seq.length v /\ + begin match res, parse p (Seq.slice v (SZ.v offset) (Seq.length v)) with + | None, None -> True + | Some off, Some (_, consumed) -> SZ.v off == consumed + | _ -> False + end + +inline_for_extraction +let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = + (input: slice byte) -> + (offset: SZ.t) -> + (#pm: perm) -> + (#v: Ghost.erased bytes) -> + stt (option SZ.t) + (pts_to input #pm v ** pure (SZ.v offset <= Seq.length v)) + (fun res -> pts_to input #pm v ** pure (validator_post p offset v res)) + +inline_for_extraction +```pulse +fn validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + v1 input offset; +} +``` + +inline_for_extraction +let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = + (input: slice byte) -> + (offset: SZ.t) -> + (#pm: perm) -> + (#v: Ghost.erased bytes) -> + stt SZ.t + (pts_to input #pm v ** pure (SZ.v offset <= Seq.length v /\ Some? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))))) + (fun res -> pts_to input #pm v ** pure (validator_post p offset v (Some res))) + +inline_for_extraction +```pulse +fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + v1 input offset; +} +``` + +let peek_post' + (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) + (input: slice byte) + (pm: perm) + (v: bytes) + (consumed: SZ.t) + (left right: slice byte) +: Tot vprop += exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** is_split input pm consumed left right ** pure ( + bare_serialize s v1 `Seq.append` v2 == v /\ + Seq.length (bare_serialize s v1) == SZ.v consumed + ) + +let peek_post + (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) + (input: slice byte) + (pm: perm) + (v: bytes) + (consumed: SZ.t) + (res: (slice byte & slice byte)) +: Tot vprop += let (left, right) = res in + peek_post' s input pm v consumed left right + +```pulse +fn peek + (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased bytes) + (consumed: SZ.t) + requires (pts_to input #pm v ** pure (validator_post #k #t p 0sz v (Some consumed))) + returns res: (slice byte & slice byte) + ensures peek_post s input pm v consumed res +{ + let s1s2 = split false input #pm #v consumed; + match s1s2 { + Mktuple2 s1 s2 -> { + Seq.lemma_split v (SZ.v consumed); + let v1 = Ghost.hide (fst (Some?.v (parse p v))); + parse_injective #k p (bare_serialize s v1) v; + unfold (split_post input pm v consumed (s1, s2)); + unfold (split_post' input pm v consumed s1 s2); + with v1' . assert (pts_to s1 #pm v1'); + rewrite (pts_to s1 #pm v1') as (pts_to_serialized s s1 #pm v1); + fold (peek_post' s input pm v consumed s1 s2); + fold (peek_post s input pm v consumed (s1, s2)); + (s1, s2) + } + } +} +``` + +inline_for_extraction +let reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) +: Tot Type += (input: slice byte) -> + (#pm: perm) -> + (#v: Ghost.erased t) -> + (pre: vprop) -> + (t': Type0) -> + (post: (t' -> vprop)) -> + (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) post) -> + stt t' (pts_to_serialized s input #pm v ** pre) post + +inline_for_extraction +let call_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun t' -> post t')) +: stt t' (pts_to_serialized s input #pm v ** pre) (fun t' -> post t') += r input pre t' (fun x -> post x) (fun x -> cont x) + +#push-options "--print_implicits --query_stats" + +inline_for_extraction +```pulse +fn read_cont + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: reader s) + (input: slice byte) + (pm: perm) + (v: Ghost.erased t) + (v': t { Ghost.reveal v == v' }) + requires (pts_to_serialized s input #pm v ** emp) + returns v' : t + ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) +{ + v' +} +``` + +inline_for_extraction +```pulse +fn read + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + requires (pts_to_serialized s input #pm v) + returns v' : t + ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) +{ + call_reader r input #pm #v emp t (fun v' -> pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) (read_cont r input pm v) +} +``` diff --git a/src/lowparse/pulse/Makefile b/src/lowparse/pulse/Makefile new file mode 100644 index 000000000..8267bae2b --- /dev/null +++ b/src/lowparse/pulse/Makefile @@ -0,0 +1,13 @@ +all: verify + +EVERCBOR_SRC_PATH = $(realpath ../..) +INCLUDE_PATHS += $(realpath ..) + +include $(EVERCBOR_SRC_PATH)/pulse.Makefile +include $(EVERCBOR_SRC_PATH)/karamel.Makefile +include $(EVERCBOR_SRC_PATH)/everparse.Makefile + +ALREADY_CACHED := *,-LowParse.Pulse, +OUTPUT_DIRECTORY := _output + +include $(EVERCBOR_SRC_PATH)/common.Makefile From 9777ccf0455729d53b54ab5ae745873525949b94 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Sun, 7 Jul 2024 23:27:39 -0700 Subject: [PATCH 02/89] WIP LowParse.Pulse.Combinators --- .../pulse/LowParse.Pulse.Combinators.fst | 114 +++++++++ src/lowparse/pulse/LowParse.Pulse.Util.fst | 230 ++++++++++++++++++ 2 files changed, 344 insertions(+) create mode 100644 src/lowparse/pulse/LowParse.Pulse.Combinators.fst create mode 100644 src/lowparse/pulse/LowParse.Pulse.Util.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst new file mode 100644 index 000000000..2a868ef88 --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -0,0 +1,114 @@ +module LowParse.Pulse.Combinators +include LowParse.Spec.Combinators +include LowParse.Pulse.Base +open FStar.Tactics.V2 +open LowParse.Pulse.Util +open Pulse.Lib.Stick +open Pulse.Lib.Slice + +module SZ = FStar.SizeT + +let parse_serialize_strong_prefix + (#t: Type) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (v: t) + (suff: bytes) +: Lemma + (requires (k.parser_kind_subkind == Some ParserStrong)) + (ensures ( + let sv = bare_serialize s v in + parse p (sv `Seq.append` suff) == Some (v, Seq.length sv) + )) += let sv = bare_serialize s v in + parse_strong_prefix #k p sv (sv `Seq.append` suff) + +let split_nondep_then_post' + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (#k2: parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2) + (input: slice byte) + (pm: perm) + (v: Ghost.erased (t1 & t2)) + (left right: slice byte) +: Tot vprop += pts_to_serialized s1 left #pm (fst v) ** + pts_to_serialized s2 right #pm (snd v) ** + ((pts_to_serialized s1 left #pm (fst v) ** + pts_to_serialized s2 right #pm (snd v)) @==> + pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + +let split_nondep_then_post + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (#k2: parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2) + (input: slice byte) + (pm: perm) + (v: Ghost.erased (t1 & t2)) + (res: (slice byte & slice byte)) +: Tot vprop += let (left, right) = res in + split_nondep_then_post' s1 s2 input pm v left right + +```pulse +fn split_nondep_then + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (t1 & t2)) + requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + returns res: (slice byte & slice byte) + ensures split_nondep_then_post s1 s2 input pm v res +{ + tot_serialize_nondep_then_eq s1 s2 v; + rewrite (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + as (pts_to input #pm (bare_serialize s1 (fst v) `Seq.append` bare_serialize s2 (snd v))); + parse_serialize_strong_prefix s1 (fst v) (bare_serialize s2 (snd v)); + let i = j1 input 0sz; + let res = slice_append_split false input i; + match res { + Mktuple2 input1 input2 -> { + unfold (slice_append_split_post input pm (bare_serialize s1 (fst v)) (bare_serialize s2 (snd v)) i res); + unfold (slice_append_split_post' input pm (bare_serialize s1 (fst v)) (bare_serialize s2 (snd v)) i input1 input2); + fold (pts_to_serialized s1 input1 #pm (fst v)); + fold (pts_to_serialized s2 input2 #pm (snd v)); + ghost + fn aux + (_foo: unit) + requires is_split input pm i input1 input2 ** (pts_to_serialized s1 input1 #pm (fst v) ** pts_to_serialized s2 input2 #pm (snd v)) + ensures pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + { + unfold (pts_to_serialized s1 input1 #pm (fst v)); + unfold (pts_to_serialized s2 input2 #pm (snd v)); + join input1 input2 input; + rewrite (pts_to input #pm (bare_serialize s1 (fst v) `Seq.append` bare_serialize s2 (snd v))) + as (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + }; + intro_stick + (pts_to_serialized s1 input1 #pm (fst v) ** pts_to_serialized s2 input2 #pm (snd v)) + (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + (is_split input pm i input1 input2) + aux; + fold (split_nondep_then_post' s1 s2 input pm v input1 input2); + fold (split_nondep_then_post s1 s2 input pm v (input1, input2)); + (input1, input2) + } + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst new file mode 100644 index 000000000..ddfdf854c --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -0,0 +1,230 @@ +module LowParse.Pulse.Util +include Pulse.Lib.Pervasives +include Pulse.Lib.Stick + +module S = Pulse.Lib.Slice +module SZ = FStar.SizeT + +```pulse +ghost +fn stick_trans + (p1 p2 p3: vprop) + requires (p1 @==> p2) ** (p2 @==> p3) + ensures (p1 @==> p3) +{ + ghost + fn aux + (_foo: unit) + requires ((p1 @==> p2) ** (p2 @==> p3)) ** p1 + ensures p3 + { + elim_stick p1 p2; + elim_stick p2 p3 + }; + intro_stick p1 p3 ((p1 @==> p2) ** (p2 @==> p3)) aux +} +``` + +```pulse +ghost +fn stick_reg_l + (p p1 p2: vprop) + requires (p1 @==> p2) + ensures ((p ** p1) @==> (p ** p2)) +{ + ghost + fn aux + (_foo: unit) + requires ((p1 @==> p2) ** (p ** p1)) + ensures (p ** p2) + { + elim_stick p1 p2 + }; + intro_stick (p ** p1) (p ** p2) (p1 @==> p2) aux +} +``` + +```pulse +ghost +fn stick_eq + (p1 p2: vprop) + requires pure (p1 == p2) // ideally with `vprop_equivs ()` + ensures (p1 @==> p2) +{ + ghost + fn aux + (_foo: unit) + requires pure (p1 == p2) ** p1 + ensures p2 + { + rewrite p1 as p2 + }; + intro_stick p1 p2 (pure (p1 == p2)) aux +} +``` + +```pulse +ghost +fn stick_rewrite_l + (l1 l2 r: vprop) + requires (l1 @==> r) ** pure (l1 == l2) + ensures l2 @==> r +{ + rewrite (l1 @==> r) as (l2 @==> r) +} +``` + +```pulse +ghost +fn stick_rewrite_r + (l r1 r2: vprop) + requires (l @==> r1) ** pure (r1 == r2) + ensures l @==> r2 +{ + rewrite (l @==> r1) as (l @==> r2) +} +``` + +```pulse +ghost +fn stick_reg_r + (p1 p2 p: vprop) + requires (p1 @==> p2) + ensures ((p1 ** p) @==> (p2 ** p)) +{ + stick_reg_l p p1 p2; + vprop_equivs (); + rewrite ((p ** p1) @==> (p ** p2)) + as ((p1 ** p) @==> (p2 ** p)) +} +``` + +```pulse +ghost +fn stick_weak_concl_l + (p1 p2 p: vprop) + requires (p1 @==> p2) ** p + ensures (p1 @==> (p ** p2)) +{ + ghost + fn aux + (_foo: unit) + requires ((p1 @==> p2) ** p) ** p1 + ensures p ** p2 + { + elim_stick p1 p2 + }; + intro_stick p1 (p ** p2) ((p1 @==> p2) ** p) aux +} +``` + +```pulse +ghost +fn stick_weak_concl_r + (p1 p2 p: vprop) + requires (p1 @==> p2) ** p + ensures (p1 @==> (p2 ** p)) +{ + stick_weak_concl_l p1 p2 p; + vprop_equivs (); + stick_eq (p ** p2) (p2 ** p); + stick_trans p1 _ _ +} +``` + +```pulse +ghost +fn stick_prod + (l1 r1 l2 r2: vprop) + requires ((l1 @==> r1) ** (l2 @==> r2)) + ensures ((l1 ** l2) @==> (r1 ** r2)) +{ + ghost + fn aux + (_foo: unit) + requires ((l1 @==> r1) ** (l2 @==> r2)) ** (l1 ** l2) + ensures r1 ** r2 + { + elim_stick l1 r1; + elim_stick l2 r2 + }; + intro_stick (l1 ** l2) (r1 ** r2) ((l1 @==> r1) ** (l2 @==> r2)) aux +} +``` + +```pulse +ghost +fn stick_elim_partial_l + (p p1 p2: vprop) + requires p ** ((p ** p1) @==> p2) + ensures p1 @==> p2 +{ + ghost + fn aux + (_foo: unit) + requires (p ** ((p ** p1) @==> p2)) ** p1 + ensures p2 + { + elim_stick (p ** p1) p2 + }; + intro_stick p1 p2 (p ** ((p ** p1) @==> p2)) aux +} +``` + +```pulse +ghost +fn stick_elim_partial_r + (p1 p p2: vprop) + requires ((p1 ** p) @==> p2) ** p + ensures p1 @==> p2 +{ + vprop_equivs (); + stick_rewrite_l (p1 ** p) (p ** p1) p2; + stick_elim_partial_l p p1 p2 +} +``` + +let slice_append_split_precond + (#t: Type) (mutb: bool) (p: perm) (v1: Ghost.erased (Seq.seq t)) (i: SZ.t) +: Tot prop += SZ.v i == Seq.length v1 /\ (mutb == true ==> p == 1.0R) + +let slice_append_split_post' + (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + (s1: S.slice t) + (s2: S.slice t) +: Tot vprop += + S.pts_to s1 #p v1 ** + S.pts_to s2 #p v2 ** + S.is_split s p i s1 s2 + +let slice_append_split_post + (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + (res: (S.slice t & S.slice t)) +: Tot vprop += let (s1, s2) = res in + slice_append_split_post' s p v1 v2 i s1 s2 + +inline_for_extraction +```pulse +fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + requires S.pts_to s #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) + returns res: (S.slice t & S.slice t) + ensures slice_append_split_post s p v1 v2 i res +{ + let vs = Ghost.hide (Seq.split (Seq.append v1 v2) (SZ.v i)); + assert (pure (fst vs `Seq.equal` v1)); + assert (pure (snd vs `Seq.equal` v2)); + let res = S.split mutb s i; + match res { + Mktuple2 s1 s2 -> { + unfold (S.split_post s p (Seq.append v1 v2) i res); + unfold (S.split_post' s p (Seq.append v1 v2) i s1 s2); + fold (slice_append_split_post' s p v1 v2 i s1 s2); + fold (slice_append_split_post s p v1 v2 i (s1, s2)); + (s1, s2) + } + } +} +``` From 573761d310470498a28f5e12580d0bda601a5696 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 11:43:12 -0700 Subject: [PATCH 03/89] integer readers --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 69 +++++- .../pulse/LowParse.Pulse.Endianness.fst | 226 ++++++++++++++++++ src/lowparse/pulse/LowParse.Pulse.Int.fst | 97 ++++++++ 3 files changed, 390 insertions(+), 2 deletions(-) create mode 100644 src/lowparse/pulse/LowParse.Pulse.Endianness.fst create mode 100644 src/lowparse/pulse/LowParse.Pulse.Int.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 567da813a..43b75cacb 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -130,8 +130,8 @@ let reader (pre: vprop) -> (t': Type0) -> (post: (t' -> vprop)) -> - (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) post) -> - stt t' (pts_to_serialized s input #pm v ** pre) post + (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x)) -> + stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x) inline_for_extraction let call_reader @@ -190,3 +190,68 @@ fn read call_reader r input #pm #v emp t (fun v' -> pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) (read_cont r input pm v) } ``` + +inline_for_extraction +let leaf_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) +: Tot Type += (input: slice byte) -> + (#pm: perm) -> + (#v: Ghost.erased t) -> + stt t (pts_to_serialized s input #pm v) (fun res -> + pts_to_serialized s input #pm v ** + pure (res == Ghost.reveal v) + ) + +inline_for_extraction +let leaf_read + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: leaf_reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) +: stt t (pts_to_serialized s input #pm v) (fun res -> + pts_to_serialized s input #pm v ** + pure (res == Ghost.reveal v) + ) += r input #pm #v + +inline_for_extraction +```pulse +fn reader_of_leaf_reader' + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: leaf_reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x)) + requires (pts_to_serialized s input #pm v ** pre) + returns res: t' + ensures post res +{ + let res = leaf_read r input #pm #v; + cont res +} +``` + +inline_for_extraction +let reader_of_leaf_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: leaf_reader s) +: reader s += reader_of_leaf_reader' r diff --git a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst new file mode 100644 index 000000000..2124abe49 --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst @@ -0,0 +1,226 @@ +module LowParse.Pulse.Endianness +open LowParse.Pulse.Util +include LowParse.Spec.Endianness + +module U8 = FStar.UInt8 +module E = LowParse.Endianness +module SZ = FStar.SizeT +module S = Pulse.Lib.Slice + +inline_for_extraction +noextract +let be_to_n_t + (#t: Type0) + (#tot: nat) + (u: uinttype u#0 t tot) + (len: nat { len <= tot }) +: Tot Type += (x: S.slice U8.t) -> + (#pm: perm) -> + (#v: Ghost.erased (Seq.seq U8.t)) -> + (pos: SZ.t) -> + stt t + (S.pts_to x #pm v ** pure ( + SZ.v pos == len /\ + len <= Seq.length v + )) + (fun res -> S.pts_to x #pm v ** pure ( + SZ.v pos == len /\ + len <= Seq.length v /\ + u.v res == E.be_to_n (Seq.slice v 0 len) + )) + +inline_for_extraction +noextract +```pulse +fn be_to_n_0 + (#t: Type0) + (#tot: nat) + (u: uinttype t tot) +: be_to_n_t #t #tot u 0 += (x: S.slice U8.t) + (#pm: perm) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + E.reveal_be_to_n (Seq.slice (v) 0 0); + u.zero +} +``` + +open FStar.Math.Lemmas +open FStar.Mul + +inline_for_extraction +noextract +```pulse +fn be_to_n_1 + (#t: Type) + (#tot: nat) + (u: uinttype t tot { tot > 0 }) +: (be_to_n_t #t #tot u 1) += (x: S.slice U8.t) + (#pm: perm) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + E.reveal_be_to_n (Seq.slice (v) 0 1); + E.reveal_be_to_n (Seq.slice (v) 0 0); + let last = S.op_Array_Access x 0sz; + (u.from_byte last) +} +``` + +inline_for_extraction +noextract +```pulse +fn be_to_n_S + (#t: Type) + (#tot: nat) + (#u: uinttype t tot) + (#len: nat { len + 1 <= tot }) + (ih: be_to_n_t #t #tot u len) +: (be_to_n_t #t #tot u (len + 1)) += (x: S.slice U8.t) + (#pm: perm) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + assert_norm (pow2 8 == 256); + E.reveal_be_to_n (Seq.slice (v) 0 (len + 1)); + E.lemma_be_to_n_is_bounded (Seq.slice (v) 0 len); + pow2_le_compat (8 * tot) (8 * (len + 1)); + pow2_le_compat (8 * (len + 1)) (8 * len); + pow2_plus (8 * len) 8; + let pos' = pos `SZ.sub` 1sz; + let last = S.op_Array_Access x pos'; + let n = ih x #pm #v pos'; + let blast = u.from_byte last; + u.add blast (u.mul256 n) +} +``` + +// attribute for use with delta_attr +noextract +noeq +type must_reduce = | MustReduce_dummy_do_not_use + +[@must_reduce] +noextract +let rec mk_be_to_n + (#t: Type) + (#tot: nat) + (u: uinttype t tot) + (len: nat {len <= tot}) +: Tot (be_to_n_t u len) + (decreases len) += if len = 0 + then be_to_n_0 u + else if len = 1 + then be_to_n_1 u + else be_to_n_S (mk_be_to_n u (len - 1)) + +// Disclaimer: a function of type n_to_be_t is ultimately meant to be called with a buffer of size len, so we do not care about any output bytes beyond position len. + +inline_for_extraction +noextract +let n_to_be_t + (#t: Type0) + (#tot: nat) + (u: uinttype t tot) + (len: nat { len <= tot }) +: Tot Type += (n: t) -> + (x: S.slice U8.t) -> + (#v: Ghost.erased (Seq.seq U8.t)) -> + (pos: SZ.t) -> + stt unit + (S.pts_to x v ** pure ( + SZ.v pos == len /\ + len <= Seq.length v /\ + u.v n < pow2 (8 * len) + )) + (fun _ -> exists* v' . S.pts_to x v' ** pure ( + SZ.v pos == len /\ + len <= Seq.length v /\ + Seq.length v' == Seq.length v /\ + u.v n < pow2 (8 * len) /\ + Seq.slice (v') 0 len `Seq.equal` n_to_be len (u.v n) + )) + +inline_for_extraction +noextract +```pulse +fn n_to_be_0 + (#t: Type0) + (#tot: nat) + (u: uinttype t tot) +: (n_to_be_t #t #tot u 0) += (n: t) + (x: S.slice U8.t) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + E.reveal_be_to_n (Seq.slice (v) 0 0); + () +} +``` + +inline_for_extraction +noextract +```pulse +fn n_to_be_1 + (#t: Type) + (#tot: nat) + (u: uinttype t tot { tot > 0 }) +: (n_to_be_t #t #tot u 1) += (n: t) + (x: S.slice U8.t) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + E.reveal_n_to_be 1 (u.v n); + E.reveal_n_to_be 0 (u.v n / pow2 8); + let n' = u.to_byte n; + S.op_Array_Assignment x 0sz n' +} +``` + +inline_for_extraction +noextract +```pulse +fn n_to_be_S + (#t: Type) + (#tot: nat) + (#u: uinttype t tot) + (#len: nat {len + 1 <= tot}) + (ih: n_to_be_t #t #tot u len) +: (n_to_be_t #t #tot u (len + 1)) += (n: t) + (x: S.slice U8.t) + (#v: Ghost.erased (Seq.seq U8.t)) + (pos: SZ.t) +{ + reveal_n_to_be (len + 1) (u.v n); + let lo = u.to_byte n; + let hi = u.div256 n; + let pos' = pos `SZ.sub` 1sz; + let _ = ih hi x pos'; + S.op_Array_Assignment x pos' lo +} +``` + +[@must_reduce] +noextract +let rec mk_n_to_be + (#t: Type) + (#tot: nat) + (u: uinttype t tot) + (len: nat {len <= tot}) +: Tot (n_to_be_t u len) + (decreases len) += if len = 0 + then n_to_be_0 u + else if len = 1 + then n_to_be_1 u + else n_to_be_S (mk_n_to_be u (len - 1)) diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst new file mode 100644 index 000000000..db5f8727a --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -0,0 +1,97 @@ +module LowParse.Pulse.Int +include LowParse.Spec.Int +include LowParse.Pulse.Base +open LowParse.Pulse.Util + +module E = LowParse.Pulse.Endianness +module EI = LowParse.Spec.Endianness.Instances +module SZ = FStar.SizeT +module S = Pulse.Lib.Slice + +inline_for_extraction +noextract +let be_to_n_1 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint8 1) + +inline_for_extraction +```pulse +fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 tot_serialize_u8 += (input: S.slice byte) + (#pm: perm) + (#v: Ghost.erased FStar.UInt8.t) +{ + unfold (pts_to_serialized tot_serialize_u8 input #pm v); + serialize_u8_spec_be v; + let res = be_to_n_1 input 1sz; + fold (pts_to_serialized tot_serialize_u8 input #pm v); + res +} +``` + +inline_for_extraction +let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) + +inline_for_extraction +noextract +let be_to_n_2 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint16 2) + +inline_for_extraction +```pulse +fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_u16 tot_serialize_u16 += (input: S.slice byte) + (#pm: perm) + (#v: Ghost.erased FStar.UInt16.t) +{ + unfold (pts_to_serialized tot_serialize_u16 input #pm v); + serialize_u16_spec_be v; + let res = be_to_n_2 input 2sz; + fold (pts_to_serialized tot_serialize_u16 input #pm v); + res +} +``` + +inline_for_extraction +let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) + +inline_for_extraction +noextract +let be_to_n_4 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint32 4) + +inline_for_extraction +```pulse +fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_u32 tot_serialize_u32 += (input: S.slice byte) + (#pm: perm) + (#v: Ghost.erased FStar.UInt32.t) +{ + unfold (pts_to_serialized tot_serialize_u32 input #pm v); + serialize_u32_spec_be v; + let res = be_to_n_4 input 4sz; + fold (pts_to_serialized tot_serialize_u32 input #pm v); + res +} +``` + +inline_for_extraction +let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) + +inline_for_extraction +noextract +let be_to_n_8 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint64 8) + +inline_for_extraction +```pulse +fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_u64 tot_serialize_u64 += (input: S.slice byte) + (#pm: perm) + (#v: Ghost.erased FStar.UInt64.t) +{ + unfold (pts_to_serialized tot_serialize_u64 input #pm v); + serialize_u64_spec_be v; + let res = be_to_n_8 input 8sz; + fold (pts_to_serialized tot_serialize_u64 input #pm v); + res +} +``` + +inline_for_extraction +let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) From b08443989e37a5475707ad66ba53da5dc2c98725 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 13:23:44 -0700 Subject: [PATCH 04/89] validate, jump int --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 59 ++++++++++++++++++++++ src/lowparse/pulse/LowParse.Pulse.Int.fst | 32 ++++++++++++ 2 files changed, 91 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 43b75cacb..d3b3e6cab 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -20,6 +20,24 @@ let validator_post (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) ( | _ -> False end +let validator_post_intro_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : Lemma + (requires ( + SZ.v offset <= Seq.length v /\ ( + let pv = parse p (Seq.slice v (SZ.v offset) (Seq.length v)) in + Some? pv /\ + snd (Some?.v pv) == SZ.v off + ))) + (ensures (validator_post p offset v (Some off))) += () + +let validator_post_intro_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : Lemma + (requires ( + SZ.v offset <= Seq.length v /\ + parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == None + )) + (ensures (validator_post p offset v None)) += () + inline_for_extraction let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (input: slice byte) -> @@ -42,6 +60,30 @@ fn validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator } ``` +inline_for_extraction +```pulse +fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { + k.parser_kind_high == Some k.parser_kind_low /\ + k.parser_kind_low == SZ.v sz /\ + k.parser_kind_metadata == Some ParserKindMetadataTotal +}) +: validator #_ #k p = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + parser_kind_prop_equiv k p; + pts_to_len input; + if SZ.lt (SZ.sub (len input) offset) sz + { + None #SZ.t + } else { + Some (sz <: SZ.t) + } +} +``` + inline_for_extraction let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (input: slice byte) -> @@ -64,6 +106,23 @@ fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (# } ``` +inline_for_extraction +```pulse +fn jump_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { + k.parser_kind_high == Some k.parser_kind_low /\ + k.parser_kind_low == SZ.v sz +}) +: jumper #_ #k p = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + parser_kind_prop_equiv k p; + (sz <: SZ.t) +} +``` + let peek_post' (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index db5f8727a..30452d13f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -8,6 +8,14 @@ module EI = LowParse.Spec.Endianness.Instances module SZ = FStar.SizeT module S = Pulse.Lib.Slice +inline_for_extraction +let validate_u8 : validator tot_parse_u8 = + validate_total_constant_size tot_parse_u8 1sz + +inline_for_extraction +let jump_u8 : jumper tot_parse_u8 = + jump_constant_size tot_parse_u8 1sz + inline_for_extraction noextract let be_to_n_1 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint8 1) @@ -30,6 +38,14 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 inline_for_extraction let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) +inline_for_extraction +let validate_u16 : validator tot_parse_u16 = + validate_total_constant_size tot_parse_u16 2sz + +inline_for_extraction +let jump_u16 : jumper tot_parse_u16 = + jump_constant_size tot_parse_u16 2sz + inline_for_extraction noextract let be_to_n_2 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint16 2) @@ -52,6 +68,14 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_ inline_for_extraction let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) +inline_for_extraction +let validate_u32 : validator tot_parse_u32 = + validate_total_constant_size tot_parse_u32 4sz + +inline_for_extraction +let jump_u32 : jumper tot_parse_u32 = + jump_constant_size tot_parse_u32 4sz + inline_for_extraction noextract let be_to_n_4 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint32 4) @@ -74,6 +98,14 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_ inline_for_extraction let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) +inline_for_extraction +let validate_u64 : validator tot_parse_u64 = + validate_total_constant_size tot_parse_u64 8sz + +inline_for_extraction +let jump_u64 : jumper tot_parse_u64 = + jump_constant_size tot_parse_u64 8sz + inline_for_extraction noextract let be_to_n_8 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint64 8) From de8c9001bf72a0f4a85a64b2191f61edec3f3217 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 13:36:10 -0700 Subject: [PATCH 05/89] switch to CPS validators --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 52 +++++++++++----------- 1 file changed, 26 insertions(+), 26 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index d3b3e6cab..c0d1aa1d2 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -12,31 +12,16 @@ let serializer #k = tot_serializer #k let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : vprop = pts_to input #pm (bare_serialize s v) -let validator_post (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (res: option SZ.t) : Tot prop = - SZ.v offset <= Seq.length v /\ - begin match res, parse p (Seq.slice v (SZ.v offset) (Seq.length v)) with - | None, None -> True - | Some off, Some (_, consumed) -> SZ.v off == consumed - | _ -> False - end - -let validator_post_intro_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : Lemma - (requires ( +let validator_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : Tot prop = SZ.v offset <= Seq.length v /\ ( let pv = parse p (Seq.slice v (SZ.v offset) (Seq.length v)) in Some? pv /\ snd (Some?.v pv) == SZ.v off - ))) - (ensures (validator_post p offset v (Some off))) -= () + ) -let validator_post_intro_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : Lemma - (requires ( +let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : Tot prop = SZ.v offset <= Seq.length v /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == None - )) - (ensures (validator_post p offset v None)) -= () inline_for_extraction let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = @@ -44,9 +29,14 @@ let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (offset: SZ.t) -> (#pm: perm) -> (#v: Ghost.erased bytes) -> - stt (option SZ.t) - (pts_to input #pm v ** pure (SZ.v offset <= Seq.length v)) - (fun res -> pts_to input #pm v ** pure (validator_post p offset v res)) + (pre: vprop) -> + (t': Type0) -> + (post: (t' -> vprop)) -> + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) -> + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -> + stt t' + (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) + (fun x -> post x) inline_for_extraction ```pulse @@ -55,8 +45,13 @@ fn validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p2 offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p2 offset v)) (fun x -> post x))) { - v1 input offset; + v1 input offset pre t' post ksucc kfail; } ``` @@ -72,14 +67,19 @@ fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (s (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) { parser_kind_prop_equiv k p; pts_to_len input; if SZ.lt (SZ.sub (len input) offset) sz { - None #SZ.t + kfail () } else { - Some (sz <: SZ.t) + ksucc (sz <: SZ.t) } } ``` @@ -92,7 +92,7 @@ let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (#v: Ghost.erased bytes) -> stt SZ.t (pts_to input #pm v ** pure (SZ.v offset <= Seq.length v /\ Some? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))))) - (fun res -> pts_to input #pm v ** pure (validator_post p offset v (Some res))) + (fun res -> pts_to input #pm v ** pure (validator_success p offset v res)) inline_for_extraction ```pulse @@ -154,7 +154,7 @@ fn peek (#pm: perm) (#v: Ghost.erased bytes) (consumed: SZ.t) - requires (pts_to input #pm v ** pure (validator_post #k #t p 0sz v (Some consumed))) + requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) returns res: (slice byte & slice byte) ensures peek_post s input pm v consumed res { From 2e9bf08a1da896826151b06ac363c747bd69f92e Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 17:47:20 -0700 Subject: [PATCH 06/89] synth, filter --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 128 +++++-- .../pulse/LowParse.Pulse.Combinators.fst | 347 ++++++++++++++++++ 2 files changed, 440 insertions(+), 35 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index c0d1aa1d2..c468f8a27 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -1,6 +1,6 @@ module LowParse.Pulse.Base open FStar.Tactics.V2 -open Pulse.Lib.Pervasives +open LowParse.Pulse.Util open Pulse.Lib.Slice open LowParse.Spec.Base @@ -12,16 +12,16 @@ let serializer #k = tot_serializer #k let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : vprop = pts_to input #pm (bare_serialize s v) -let validator_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : Tot prop = - SZ.v offset <= Seq.length v /\ ( +let validator_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : GTot bool = + SZ.v offset <= Seq.length v && ( let pv = parse p (Seq.slice v (SZ.v offset) (Seq.length v)) in - Some? pv /\ - snd (Some?.v pv) == SZ.v off + Some? pv && + snd (Some?.v pv) = SZ.v off ) -let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : Tot prop = - SZ.v offset <= Seq.length v /\ - parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == None +let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : GTot bool = + SZ.v offset <= Seq.length v && + None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))) inline_for_extraction let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = @@ -39,21 +39,8 @@ let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (fun x -> post x) inline_for_extraction -```pulse -fn validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p2 offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p2 offset v)) (fun x -> post x))) -{ - v1 input offset pre t' post ksucc kfail; -} -``` +let validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = + v1 inline_for_extraction ```pulse @@ -84,6 +71,26 @@ fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (s } ``` +let jumper_pre + (#t: Type0) + (#k: parser_kind) + (p: parser k t) + (offset: SZ.t) + (v: Ghost.erased bytes) +: GTot bool += + SZ.v offset <= Seq.length v && Some? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))) + +let jumper_pre' + (#t: Type0) + (#k: parser_kind) + (p: parser k t) + (offset: SZ.t) + (v: Ghost.erased bytes) +: Tot prop += + jumper_pre p offset v + inline_for_extraction let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (input: slice byte) -> @@ -91,20 +98,12 @@ let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (#pm: perm) -> (#v: Ghost.erased bytes) -> stt SZ.t - (pts_to input #pm v ** pure (SZ.v offset <= Seq.length v /\ Some? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))))) + (pts_to input #pm v ** pure (jumper_pre' p offset v)) (fun res -> pts_to input #pm v ** pure (validator_success p offset v res)) inline_for_extraction -```pulse -fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) -{ - v1 input offset; -} -``` +let jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = + v1 inline_for_extraction ```pulse @@ -133,7 +132,8 @@ let peek_post' : Tot vprop = exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** is_split input pm consumed left right ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ - Seq.length (bare_serialize s v1) == SZ.v consumed + Seq.length (bare_serialize s v1) == SZ.v consumed /\ + parse p v == Some (v1, SZ.v consumed) ) let peek_post @@ -176,6 +176,64 @@ fn peek } ``` +let peek_stick_post' + (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) + (input: slice byte) + (pm: perm) + (v: bytes) + (consumed: SZ.t) + (left right: slice byte) +: Tot vprop += exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** ((pts_to_serialized s left #pm v1 ** pts_to right #pm v2) @==> pts_to input #pm v) ** pure ( + bare_serialize s v1 `Seq.append` v2 == v /\ + Seq.length (bare_serialize s v1) == SZ.v consumed /\ + parse p v == Some (v1, SZ.v consumed) + ) + +let peek_stick_post + (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) + (input: slice byte) + (pm: perm) + (v: bytes) + (consumed: SZ.t) + (res: (slice byte & slice byte)) +: Tot vprop += let (left, right) = res in + peek_stick_post' s input pm v consumed left right + +```pulse +fn peek_stick + (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased bytes) + (consumed: SZ.t) + requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) + returns res: (slice byte & slice byte) + ensures peek_stick_post s input pm v consumed res +{ + let res = peek s input consumed; + match res { Mktuple2 left right -> { + unfold (peek_post s input pm v consumed res); + unfold (peek_post' s input pm v consumed left right); + with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); + ghost + fn aux + (_foo: unit) + requires (is_split input pm consumed left right ** (pts_to_serialized s left #pm v1 ** pts_to right #pm v2)) + ensures pts_to input #pm v + { + unfold (pts_to_serialized s left #pm v1); + join left right input + }; + intro_stick (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) aux; + fold (peek_stick_post' s input pm v consumed left right); + fold (peek_stick_post s input pm v consumed (left, right)); + (left, right) + }} +} +``` + inline_for_extraction let reader (#t: Type0) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 2a868ef88..9103bcc62 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -24,6 +24,352 @@ let parse_serialize_strong_prefix = let sv = bare_serialize s v in parse_strong_prefix #k p sv (sv `Seq.append` suff) +let tot_parse_synth_eq' + (#k: parser_kind) + (#t1: Type) + (#t2: Type) + (p1: tot_parser k t1) + (f2: (t1 -> Tot t2) {synth_injective f2}) + (b: bytes) +: Lemma + (ensures (parse (tot_parse_synth p1 f2) b == parse_synth' #k p1 f2 b)) += parse_synth_eq #k p1 f2 b + +inline_for_extraction +let validate_synth + (#t #t': Type) + (#k: parser_kind) + (#p: tot_parser k t) + (w: validator p) + (f: (t -> t') { synth_injective f }) +: Tot (validator (tot_parse_synth p f)) += Classical.forall_intro (tot_parse_synth_eq' p f); + w + +inline_for_extraction +let jump_synth + (#t #t': Type) + (#k: parser_kind) + (#p: tot_parser k t) + (w: jumper p) + (f: (t -> t') { synth_injective f }) +: Tot (jumper (tot_parse_synth p f)) += Classical.forall_intro (tot_parse_synth_eq' p f); + w + +```pulse +ghost +fn pts_to_serialized_synth_intro + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (#v: t) + requires pts_to_serialized s input #pm v + ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) +{ + tot_parse_synth_eq p f (bare_serialize s v); + parse_serialize #k #t' #(tot_parse_synth p f) (tot_serialize_synth p f s f' ()) (f v); + parse_injective #k #t' (tot_parse_synth p f) (bare_serialize s v) (bare_serialize (tot_serialize_synth p f s f' ()) (f v)); + unfold (pts_to_serialized s input #pm v); + rewrite (pts_to input #pm (bare_serialize s v)) + as (pts_to input #pm (bare_serialize (tot_serialize_synth p f s f' ()) (f v))); + fold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)) +} +``` + +```pulse +ghost +fn pts_to_serialized_synth_elim + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (v: t) + requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) + ensures pts_to_serialized s input #pm v +{ + tot_parse_synth_eq p f (bare_serialize s v); + parse_serialize #k #t' #(tot_parse_synth p f) (tot_serialize_synth p f s f' ()) (f v); + parse_injective #k #t' (tot_parse_synth p f) (bare_serialize s v) (bare_serialize (tot_serialize_synth p f s f' ()) (f v)); + unfold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)); + rewrite (pts_to input #pm (bare_serialize (tot_serialize_synth p f s f' ()) (f v))) + as (pts_to input #pm (bare_serialize s v)); + fold (pts_to_serialized s input #pm v) +} +``` + +```pulse +ghost +fn pts_to_serialized_synth_l2r + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (#v: t') + requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + ensures pts_to_serialized s input #pm (f' v) +{ + tot_serialize_synth_eq p f s f' () v; + unfold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v); + fold (pts_to_serialized s input #pm (f' v)) +} +``` + +```pulse +ghost +fn pts_to_serialized_synth_r2l + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (v: t') + requires pts_to_serialized s input #pm (f' v) + ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v +{ + tot_serialize_synth_eq p f s f' () v; + unfold (pts_to_serialized s input #pm (f' v)); + fold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) +} +``` + +inline_for_extraction +```pulse +fn validate_filter_gen_cont_success + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: serializer p) + (f: (t -> bool)) + (f': (input: slice byte) -> (pm: perm) -> (v: Ghost.erased t) -> stt bool (pts_to_serialized s input #pm v) (fun res -> pts_to_serialized s input #pm v ** pure (res == f v))) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) + (off: SZ.t) + requires (pts_to input #pm v ** pre ** pure (validator_success (p) offset v off)) + returns res: t' + ensures post res +{ + Seq.lemma_split v (SZ.v offset); + let split123 = split false input offset; + match split123 { Mktuple2 input1 input23 -> { + unfold (split_post input pm v offset split123); + unfold (split_post' input pm v offset input1 input23); + with v1 v23 . assert (pts_to input1 #pm v1 ** pts_to input23 #pm v23); + tot_parse_filter_eq p f v23; + let split23 = peek_stick s input23 off; + match split23 { Mktuple2 input2 input3 -> { + unfold (peek_stick_post s input23 pm v23 off split23); + unfold (peek_stick_post' s input23 pm v23 off input2 input3); + let cond = f' input2 pm _; + elim_stick (pts_to_serialized s input2 #pm _ ** _) _; + join input1 input23 input; + rewrite (pts_to input #pm (Seq.append v1 v23)) as (pts_to input #pm v); + if cond { + ksucc off + } else { + kfail () + } + }} + }} +} +``` + +inline_for_extraction +```pulse +fn validate_filter_gen_cont_failure + (#t: Type0) + (#k: parser_kind) + (p: tot_parser k t) + (f: (t -> bool)) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) + (_: unit) + requires (pts_to input #pm v ** pre ** pure (validator_failure (p) offset v)) + returns res: t' + ensures post res +{ + tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () +} +``` + +inline_for_extraction +```pulse +fn validate_filter_gen + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: serializer p) + (w: validator p) + (f: (t -> bool)) + (f': (input: slice byte) -> (pm: perm) -> (v: Ghost.erased t) -> stt bool (pts_to_serialized s input #pm v) (fun res -> pts_to_serialized s input #pm v ** pure (res == f v))) +: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) +{ + w input offset pre t' post + (validate_filter_gen_cont_success s f f' input offset #pm #v pre t' post ksucc kfail) + (validate_filter_gen_cont_failure p f input offset #pm #v pre t' post kfail) +} +``` + +inline_for_extraction +```pulse +fn validate_filter_read_cond_cont + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: serializer p) + (f: (t -> bool)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased t) + (x: t { x == Ghost.reveal v }) + requires (pts_to_serialized s input #pm v ** emp) + returns res: bool + ensures pts_to_serialized s input #pm v ** pure (res == f v) +{ + f x +} +``` + +inline_for_extraction +```pulse +fn validate_filter_read_cond + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (#s: serializer p) + (r: reader s) + (f: (t -> bool)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased t) + requires (pts_to_serialized s input #pm v) + returns res: bool + ensures pts_to_serialized s input #pm v ** pure (res == f v) +{ + r input #pm #v emp bool (fun res -> pts_to_serialized s input #pm v ** pure (res == f v)) (validate_filter_read_cond_cont s f input pm v) +} +``` + +inline_for_extraction +let validate_filter + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (#s: serializer p) + (r: reader s) + (w: validator p) + (f: (t -> bool)) +: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) += validate_filter_gen s w f (validate_filter_read_cond r f) + +inline_for_extraction +```pulse +fn jump_filter + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (w: jumper p) + (f: (t -> bool)) +: jumper #_ #(parse_filter_kind k) (tot_parse_filter p f) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + Classical.forall_intro (tot_parse_filter_eq p f); + w input offset #pm #v +} +``` + +let parse_filter_refine_intro + (#t: Type) + (f: (t -> GTot bool)) + (v: t) + (sq: squash (f v == true)) +: Tot (parse_filter_refine f) += v + +```pulse +ghost +fn pts_to_serialized_filter_intro + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: serializer p) + (f: (t -> bool)) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + requires (pts_to_serialized s input #pm v ** pure (f v == true)) + ensures exists* (v': parse_filter_refine f) . pts_to_serialized (tot_serialize_filter s f) input #pm v' ** pure ((v' <: t) == Ghost.reveal v) +{ + unfold (pts_to_serialized s input #pm v); + let sq: squash (f v == true) = (); + let v' : Ghost.erased (parse_filter_refine f) = Ghost.hide (parse_filter_refine_intro #t f (Ghost.reveal v) sq); + fold (pts_to_serialized (tot_serialize_filter s f) input #pm v'); +} +``` + +```pulse +ghost +fn pts_to_serialized_filter_elim + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: serializer p) + (f: (t -> bool)) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (parse_filter_refine f)) + requires (pts_to_serialized (tot_serialize_filter s f) input #pm v) + ensures pts_to_serialized s input #pm v +{ + unfold (pts_to_serialized (tot_serialize_filter s f) input #pm v); + fold (pts_to_serialized s input #pm v); +} +``` + let split_nondep_then_post' (#t1 #t2: Type0) (#k1: parser_kind) @@ -59,6 +405,7 @@ let split_nondep_then_post = let (left, right) = res in split_nondep_then_post' s1 s2 input pm v left right +inline_for_extraction ```pulse fn split_nondep_then (#t1 #t2: Type0) From cf0534cd399a8d001c1167b0d0ef046a97bd21f0 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 18:04:41 -0700 Subject: [PATCH 07/89] validate_nonempty --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 77 ++++++++++++++++++++++ 1 file changed, 77 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index c468f8a27..8f2f40276 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -38,6 +38,83 @@ let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) (fun x -> post x) +inline_for_extraction +let validate + (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) +: stt t' + (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) + (fun x -> post x) += w input offset #pm #v pre t' post ksucc kfail + +let validate_nonempty_post + (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) +: Tot prop += SZ.v offset <= Seq.length v /\ + (off == 0sz <==> None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v)))) /\ + (if off = 0sz then validator_failure p offset v else validator_success p offset v off) + +inline_for_extraction +```pulse +fn validate_nonempty_success (#t: Type0) (#k: parser_kind) (p: parser k t { k.parser_kind_low > 0 }) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (sq: squash (SZ.v offset <= Seq.length v)) + (off: SZ.t) + requires pts_to input #pm v ** emp ** pure (validator_success p offset v off) + returns off: SZ.t + ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) +{ + parser_kind_prop_equiv k p; + off +} +``` + +inline_for_extraction +```pulse +fn validate_nonempty_failure (#t: Type0) (#k: parser_kind) (p: parser k t { k.parser_kind_low > 0 }) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (sq: squash (SZ.v offset <= Seq.length v)) + (_: unit) + requires pts_to input #pm v ** emp ** pure (validator_failure p offset v) + returns off: SZ.t + ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) +{ + parser_kind_prop_equiv k p; + 0sz +} +``` + +inline_for_extraction +```pulse +fn validate_nonempty (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p { k.parser_kind_low > 0 }) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + requires pts_to input #pm v ** pure (SZ.v offset <= Seq.length v) + returns off: SZ.t + ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) +{ + validate w input offset #pm #v emp SZ.t (fun off -> pts_to input #pm v ** pure (validate_nonempty_post p offset v off)) + (validate_nonempty_success p input offset #pm #v ()) + (validate_nonempty_failure p input offset #pm #v ()) +} +``` + inline_for_extraction let validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = v1 From 7209ceb2cf02cad5a007c22209a671980352d0d4 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 8 Jul 2024 20:38:13 -0700 Subject: [PATCH 08/89] validate_nondep_then --- .../pulse/LowParse.Pulse.Combinators.fst | 166 ++++++++++++++++++ 1 file changed, 166 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 9103bcc62..b3c230740 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -370,6 +370,172 @@ fn pts_to_serialized_filter_elim } ``` + + +let pair_of_dtuple2 + (#t1 #t2: Type) + (x: dtuple2 t1 (fun _ -> t2)) +: Tot (t1 & t2) += match x with + | (| x1, x2 |) -> (x1, x2) + +let dtuple2_of_pair + (#t1 #t2: Type) + (x: (t1 & t2)) +: Tot (dtuple2 t1 (fun _ -> t2)) += match x with + | (x1, x2) -> (| x1, x2 |) + +let nondep_then_eq_dtuple2 + (#t1 #t2: Type) + (#k1 #k2: parser_kind) + (p1: tot_parser k1 t1) + (p2: tot_parser k2 t2) + (input: bytes) +: Lemma + (parse (tot_nondep_then p1 p2) input == parse (tot_parse_synth (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2) input) += tot_parse_synth_eq (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 input + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_success2 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (off': SZ.t) + requires pts_to input #pm v ** pre ** pure (validator_success p2 (offset `SZ.add` off) v off') + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (off + off') + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + ksucc (off `SZ.add` off') +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_failure2 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (_: unit) + requires pts_to input #pm v ** pre ** pure (validator_failure p2 (offset `SZ.add` off) v) + returns x: t' + ensures post x +{ + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_success1 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (v2: validator p2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (off: SZ.t) + requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off) + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (offset + off) + v2 input (offset `SZ.add` off) #pm #v pre t' post + (validate_nondep_then_cont_success2 p1 p2 input offset #pm #v pre t' post off ksucc) + (validate_nondep_then_cont_failure2 p1 p2 input offset #pm #v pre t' post off kfail) +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_failure1 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (_: unit) + requires pts_to input #pm v ** pre ** pure (validator_failure p1 offset v) + returns x: t' + ensures post x +{ + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: validator p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (v2: validator p2) +: validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) +{ + v1 input offset #pm #v pre t' post + (validate_nondep_then_cont_success1 p1 v2 input offset #pm #v pre t' post ksucc kfail) + (validate_nondep_then_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) +} +``` + let split_nondep_then_post' (#t1 #t2: Type0) (#k1: parser_kind) From cbf928e47cbc4e851488e60e27c39a3024f98d71 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 01:10:52 -0700 Subject: [PATCH 09/89] common validate_dtuple2 and nondep_then --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 54 ++++- .../pulse/LowParse.Pulse.Combinators.fst | 198 ++++++++++++++---- src/lowparse/pulse/LowParse.Pulse.Util.fst | 50 +++++ 3 files changed, 259 insertions(+), 43 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 8f2f40276..0fa1dfb38 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -24,11 +24,13 @@ let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))) inline_for_extraction -let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = - (input: slice byte) -> - (offset: SZ.t) -> - (#pm: perm) -> - (#v: Ghost.erased bytes) -> +let validator_for (#t: Type0) (#k: parser_kind) (p: parser k t) + (input: slice byte) + (offset: SZ.t) + (pm: perm) + (v: Ghost.erased bytes) +: Tot Type += (pre: vprop) -> (t': Type0) -> (post: (t' -> vprop)) -> @@ -38,6 +40,14 @@ let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) (fun x -> post x) +inline_for_extraction +let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = + (input: slice byte) -> + (offset: SZ.t) -> + (#pm: perm) -> + (#v: Ghost.erased bytes) -> + validator_for p input offset pm v + inline_for_extraction let validate (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) @@ -311,6 +321,40 @@ fn peek_stick } ``` +```pulse +fn peek_stick_gen + (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased bytes) + (offset: SZ.t) + (consumed: SZ.t) + requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (consumed))) + returns input': slice byte + ensures exists* v' . pts_to_serialized s input' #pm v' ** (pts_to_serialized s input' #pm v' @==> pts_to input #pm v) ** pure ( + validator_success #k #t p offset v consumed /\ + parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v consumed) + ) +{ + let split123 = slice_split_stick false input offset; + match split123 { Mktuple2 input1 input23 -> { + unfold (slice_split_stick_post input pm v offset split123); + unfold (slice_split_stick_post' input pm v offset input1 input23); + with v23 . assert (pts_to input23 #pm v23); + stick_elim_partial_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); + let split23 = peek_stick s input23 consumed; + match split23 { Mktuple2 input2 input3 -> { + unfold (peek_stick_post s input23 pm v23 consumed split23); + unfold (peek_stick_post' s input23 pm v23 consumed input2 input3); + with v' . assert (pts_to_serialized s input2 #pm v'); + stick_elim_partial_r (pts_to_serialized s input2 #pm _) (pts_to input3 #pm _) (pts_to input23 #pm v23); + stick_trans (pts_to_serialized s input2 #pm _) (pts_to input23 #pm _) (pts_to input #pm _); + input2 + }} + }} +} +``` + inline_for_extraction let reader (#t: Type0) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index b3c230740..e12e1ece9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -399,12 +399,13 @@ let nondep_then_eq_dtuple2 inline_for_extraction ```pulse -fn validate_nondep_then_cont_success2 - (#t1 #t2: Type0) +fn validate_dtuple2_cont_success2 + (#t1: Type0) + (#t2: t1 -> Type0) (#k1: parser_kind) (p1: parser k1 t1) (#k2: parser_kind) - (p2: parser k2 t2) + (p2: ((x: t1) -> parser k2 (t2 x))) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -412,27 +413,28 @@ fn validate_nondep_then_cont_success2 (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (key: Ghost.erased t1) + (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) (off': SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success p2 (offset `SZ.add` off) v off') + requires pts_to input #pm v ** pre ** pure (validator_success (p2 key) (offset `SZ.add` off) v off') returns x: t' ensures post x { pts_to_len input; // for SZ.fits (off + off') - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); ksucc (off `SZ.add` off') } ``` inline_for_extraction ```pulse -fn validate_nondep_then_cont_failure2 - (#t1 #t2: Type0) +fn validate_dtuple2_cont_failure2 + (#t1: Type0) + (#t2: t1 -> Type0) (#k1: parser_kind) (p1: parser k1 t1) (#k2: parser_kind) - (p2: parser k2 t2) + (p2: ((x: t1) -> parser k2 (t2 x))) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -440,56 +442,60 @@ fn validate_nondep_then_cont_failure2 (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (key: Ghost.erased t1) + (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) (_: unit) - requires pts_to input #pm v ** pre ** pure (validator_failure p2 (offset `SZ.add` off) v) + requires pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (offset `SZ.add` off) v) returns x: t' ensures post x { - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); kfail () } ``` inline_for_extraction ```pulse -fn validate_nondep_then_cont_success1 - (#t1 #t2: Type0) +fn validate_dtuple2_cont_success1 + (#t1: Type0) + (#t2: t1 -> Type0) (#k1: parser_kind) (p1: parser k1 t1) (#k2: parser_kind) - (#p2: parser k2 t2) - (v2: validator p2) + (p2: ((x: t1) -> parser k2 (t2 x))) (input: slice byte) (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) + (v2: (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) -> (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) -> + validator_for (p2 key) input (SZ.add offset off) pm v) (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) (off: SZ.t) requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off) returns x: t' ensures post x { - pts_to_len input; // for SZ.fits (offset + off) - v2 input (offset `SZ.add` off) #pm #v pre t' post - (validate_nondep_then_cont_success2 p1 p2 input offset #pm #v pre t' post off ksucc) - (validate_nondep_then_cont_failure2 p1 p2 input offset #pm #v pre t' post off kfail) + pts_to_len input; // for SZ.fits (offset + off) + let key : Ghost.erased t1 = Ghost.hide (fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v))))); + v2 off key pre t' post + (validate_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post key off ksucc) + (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post key off kfail) } ``` inline_for_extraction ```pulse -fn validate_nondep_then_cont_failure1 - (#t1 #t2: Type0) +fn validate_dtuple2_cont_failure1 + (#t1: Type0) + (#t2: t1 -> Type0) (#k1: parser_kind) (p1: parser k1 t1) (#k2: parser_kind) - (p2: parser k2 t2) + (p2: ((x: t1) -> parser k2 (t2 x))) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -497,20 +503,48 @@ fn validate_nondep_then_cont_failure1 (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) (_: unit) requires pts_to input #pm v ** pre ** pure (validator_failure p1 offset v) returns x: t' ensures post x { - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); kfail () } ``` inline_for_extraction ```pulse -fn validate_nondep_then +fn validate_dtuple2_gen + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: validator p1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (input: slice byte) + (offset: SZ.t) + (pm: perm) + (v: Ghost.erased bytes) + (v2: (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) -> (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) -> + validator_for (p2 key) input (SZ.add offset off) pm v) +: validator_for #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) input offset pm v += + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) +{ + v1 input offset #pm #v pre t' post + (validate_dtuple2_cont_success1 p1 p2 input offset #pm #v v2 pre t' post ksucc kfail) + (validate_dtuple2_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) +} +``` + +inline_for_extraction +let validate_nondep_then (#t1 #t2: Type0) (#k1: parser_kind) (#p1: parser k1 t1) @@ -519,23 +553,111 @@ fn validate_nondep_then (#p2: parser k2 t2) (v2: validator p2) : validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) -= += Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); + validate_ext + (validate_synth + (fun input offset #pm #v -> + validate_dtuple2_gen + v1 + (fun _ -> p2) + input + offset + pm + v + (fun off _ -> v2 input (SZ.add offset off) #pm #v) + ) + pair_of_dtuple2 + ) + (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) + +inline_for_extraction +```pulse +fn validate_dtuple2_payload_cont + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> validator (p2 x))) (input: slice byte) (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) + (pm: perm) + (v: Ghost.erased bytes) + (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (ksucc: ((off': SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (p2 key) (SZ.add offset off) v off')) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (SZ.add offset off) v)) (fun x -> post x))) + (input_key: slice byte) + (x: t1 { x == Ghost.reveal key }) + requires (pts_to_serialized s1 input_key #pm key ** (pre ** (pts_to_serialized s1 input_key #pm key @==> pts_to input #pm v))) + returns res: t' + ensures post res { - v1 input offset #pm #v pre t' post - (validate_nondep_then_cont_success1 p1 v2 input offset #pm #v pre t' post ksucc kfail) - (validate_nondep_then_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) + elim_stick (pts_to_serialized s1 input_key #pm key) _; + v2 x input (SZ.add offset off) pre t' post ksucc kfail } ``` +inline_for_extraction +```pulse +fn validate_dtuple2_payload + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (#s1: serializer p1) + (r1: reader s1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> validator (p2 x))) + (input: slice byte) + (offset: SZ.t) + (pm: perm) + (v: Ghost.erased bytes) + (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) +: validator_for #(t2 key) #k2 (p2 key) input (SZ.add offset off) pm v += + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off': SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (p2 key) (SZ.add offset off) v off')) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (SZ.add offset off) v)) (fun x -> post x))) +{ + let input' = peek_stick_gen s1 input offset off; + assert (pts_to_serialized s1 input' #pm key); + r1 input' _ t' post (validate_dtuple2_payload_cont s1 p2 v2 input offset pm v off key pre t' post ksucc kfail input') +} +``` + +inline_for_extraction +let validate_dtuple2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: validator p1) + (#s1: serializer p1) + (r1: reader s1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> validator (p2 x))) +: Tot (validator (tot_parse_dtuple2 p1 p2)) += fun input offset #pm #v -> + validate_dtuple2_gen + v1 + p2 + input + offset + pm + v + (validate_dtuple2_payload r1 p2 v2 input offset pm v) + let split_nondep_then_post' (#t1 #t2: Type0) (#k1: parser_kind) diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index ddfdf854c..9ce8e00fe 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -228,3 +228,53 @@ fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2 } } ``` + +let slice_split_stick_post' + (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) + (s1: S.slice t) + (s2: S.slice t) +: Tot vprop += exists* v1 v2 . + S.pts_to s1 #p v1 ** + S.pts_to s2 #p v2 ** + ((S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) @==> S.pts_to s #p v) ** + pure ( + SZ.v i <= Seq.length v /\ + (v1, v2) == Seq.split v (SZ.v i) + ) + +let slice_split_stick_post + (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) + (res: (S.slice t & S.slice t)) +: Tot vprop += let (s1, s2) = res in + slice_split_stick_post' s p v i s1 s2 + +inline_for_extraction +```pulse +fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) + requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) + returns res: (S.slice t & S.slice t) + ensures slice_split_stick_post s p v i res +{ + Seq.lemma_split v (SZ.v i); + let res = S.split mutb s i; + match res { Mktuple2 s1 s2 -> { + unfold (S.split_post s p v i res); + unfold (S.split_post' s p v i s1 s2); + with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); + ghost + fn aux + (_: unit) + requires (S.is_split s p i s1 s2 ** (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2)) + ensures (S.pts_to s #p v) + { + S.join s1 s2 s + }; + intro_stick _ _ _ aux; + fold (slice_split_stick_post' s p v i s1 s2); + fold (slice_split_stick_post s p v i (s1, s2)); + (s1, s2) + }} +} +``` From 6352265d3ff240a41c42382dfa9846f4b88b9df8 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 10:07:37 -0700 Subject: [PATCH 10/89] split_dtuple2, simplified split_nondep_then --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 92 +++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 195 +++++++++++++++--- src/lowparse/pulse/LowParse.Pulse.Util.fst | 66 ++++++ 3 files changed, 319 insertions(+), 34 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 0fa1dfb38..2c2b0bfc3 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -12,6 +12,98 @@ let serializer #k = tot_serializer #k let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : vprop = pts_to input #pm (bare_serialize s v) +```pulse +ghost +fn pts_to_serialized_intro_stick + (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v0: bytes) (v: t) + requires (pts_to input #pm v0 ** pure (v0 == bare_serialize s v)) + ensures (pts_to_serialized s input #pm v ** (pts_to_serialized s input #pm v @==> pts_to input #pm v0)) +{ + rewrite_with_stick (pts_to input #pm v0) (pts_to_serialized s input #pm v) +} +``` + +```pulse +ghost +fn pts_to_serialized_elim_stick + (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: t) + requires (pts_to_serialized s input #pm v) + ensures (pts_to input #pm (bare_serialize s v) ** (pts_to input #pm (bare_serialize s v) @==> pts_to_serialized s input #pm v)) +{ + rewrite_with_stick (pts_to_serialized s input #pm v) (pts_to input #pm (bare_serialize s v)) +} +``` + +let serializer_ext_eq + (#t: Type0) + (#k1: parser_kind) + (#p1: parser k1 t) + (s1: serializer p1) + (#k2: parser_kind) + (#p2: parser k2 t) + (s2: serializer p2) + (v: t) +: Lemma + (requires forall x . parse p1 x == parse p2 x) + (ensures (bare_serialize s1 v == bare_serialize s2 v)) += let s2' = serialize_ext #k1 p1 s1 #k2 p2 in + serializer_unique p2 s2 s2' v + +```pulse +ghost +fn pts_to_serialized_ext + (#t: Type0) + (#k1: parser_kind) + (#p1: parser k1 t) + (s1: serializer p1) + (#k2: parser_kind) + (#p2: parser k2 t) + (s2: serializer p2) + (input: slice byte) + (#pm: perm) + (#v: t) + requires pts_to_serialized s1 input #pm v ** pure ( + forall x . parse p1 x == parse p2 x + ) + ensures pts_to_serialized s2 input #pm v +{ + serializer_ext_eq s1 s2 v; + unfold (pts_to_serialized s1 input #pm v); + fold (pts_to_serialized s2 input #pm v) +} +``` + +```pulse +ghost +fn pts_to_serialized_ext_stick + (#t: Type0) + (#k1: parser_kind) + (#p1: parser k1 t) + (s1: serializer p1) + (#k2: parser_kind) + (#p2: parser k2 t) + (s2: serializer p2) + (input: slice byte) + (#pm: perm) + (#v: t) + requires pts_to_serialized s1 input #pm v ** pure ( + forall x . parse p1 x == parse p2 x + ) + ensures pts_to_serialized s2 input #pm v ** (pts_to_serialized s2 input #pm v @==> pts_to_serialized s1 input #pm v) +{ + pts_to_serialized_ext s1 s2 input; + ghost + fn aux + (_: unit) + requires emp ** pts_to_serialized s2 input #pm v + ensures pts_to_serialized s1 input #pm v + { + pts_to_serialized_ext s2 s1 input + }; + intro_stick _ _ _ aux +} +``` + let validator_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : GTot bool = SZ.v offset <= Seq.length v && ( let pv = parse p (Seq.slice v (SZ.v offset) (Seq.length v)) in diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index e12e1ece9..d64743775 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -107,6 +107,34 @@ fn pts_to_serialized_synth_elim } ``` +```pulse +ghost +fn pts_to_serialized_synth_stick + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (#v: t) + requires pts_to_serialized s input #pm v + ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) ** (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) @==> pts_to_serialized s input #pm v) +{ + pts_to_serialized_synth_intro s f f' input; + ghost + fn aux + (_: unit) + requires emp ** pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) + ensures pts_to_serialized s input #pm v + { + pts_to_serialized_synth_elim s f f' input v + }; + intro_stick _ _ _ aux +} +``` + ```pulse ghost fn pts_to_serialized_synth_l2r @@ -149,6 +177,34 @@ fn pts_to_serialized_synth_r2l } ``` +```pulse +ghost +fn pts_to_serialized_synth_l2r_stick + (#t #t': Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (#v: t') + requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + ensures pts_to_serialized s input #pm (f' v) ** (pts_to_serialized s input #pm (f' v) @==> pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) +{ + pts_to_serialized_synth_l2r s f f' input; + ghost + fn aux + (_: unit) + requires emp ** pts_to_serialized s input #pm (f' v) + ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + { + pts_to_serialized_synth_r2l s f f' input v + }; + intro_stick _ _ _ aux +} +``` + inline_for_extraction ```pulse fn validate_filter_gen_cont_success @@ -658,6 +714,85 @@ let validate_dtuple2 v (validate_dtuple2_payload r1 p2 v2 input offset pm v) +let split_dtuple2_post' + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (#k2: parser_kind) + (#p2: (x: t1) -> parser k2 (t2 x)) + (s2: (x: t1) -> serializer (p2 x)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased (dtuple2 t1 t2)) + (left right: slice byte) +: Tot vprop += pts_to_serialized s1 left #pm (dfst v) ** + pts_to_serialized (s2 (dfst v)) right #pm (dsnd v) ** + ((pts_to_serialized s1 left #pm (dfst v) ** + pts_to_serialized (s2 (dfst v)) right #pm (dsnd v)) @==> + pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + +let split_dtuple2_post + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (#k2: parser_kind) + (#p2: (x: t1) -> parser k2 (t2 x)) + (s2: (x: t1) -> serializer (p2 x)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased (dtuple2 t1 t2)) + (res: (slice byte & slice byte)) +: Tot vprop += let (left, right) = res in + split_dtuple2_post' s1 s2 input pm v left right + +inline_for_extraction +```pulse +fn split_dtuple2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: ((x: t1) -> parser k2 (t2 x))) + (s2: (x: t1) -> serializer (p2 x)) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (dtuple2 t1 t2)) + requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + returns res: (slice byte & slice byte) + ensures split_dtuple2_post s1 s2 input pm v res +{ + rewrite_with_stick + (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); + parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); + let i = j1 input 0sz; + let res = slice_append_split_stick false input i; + match res { + Mktuple2 input1 input2 -> { + unfold (slice_append_split_stick_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); + unfold (slice_append_split_stick_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); + stick_trans (_ ** _) _ _; + pts_to_serialized_intro_stick s1 input1 (dfst v); + pts_to_serialized_intro_stick (s2 (dfst v)) input2 (dsnd v); + stick_prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); + stick_trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; + fold (split_dtuple2_post' s1 s2 input pm v input1 input2); + fold (split_dtuple2_post s1 s2 input pm v (input1, input2)); + (input1, input2) + } + } +} +``` + let split_nondep_then_post' (#t1 #t2: Type0) (#k1: parser_kind) @@ -711,39 +846,31 @@ fn split_nondep_then returns res: (slice byte & slice byte) ensures split_nondep_then_post s1 s2 input pm v res { - tot_serialize_nondep_then_eq s1 s2 v; - rewrite (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) - as (pts_to input #pm (bare_serialize s1 (fst v) `Seq.append` bare_serialize s2 (snd v))); - parse_serialize_strong_prefix s1 (fst v) (bare_serialize s2 (snd v)); - let i = j1 input 0sz; - let res = slice_append_split false input i; - match res { - Mktuple2 input1 input2 -> { - unfold (slice_append_split_post input pm (bare_serialize s1 (fst v)) (bare_serialize s2 (snd v)) i res); - unfold (slice_append_split_post' input pm (bare_serialize s1 (fst v)) (bare_serialize s2 (snd v)) i input1 input2); - fold (pts_to_serialized s1 input1 #pm (fst v)); - fold (pts_to_serialized s2 input2 #pm (snd v)); - ghost - fn aux - (_foo: unit) - requires is_split input pm i input1 input2 ** (pts_to_serialized s1 input1 #pm (fst v) ** pts_to_serialized s2 input2 #pm (snd v)) - ensures pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v - { - unfold (pts_to_serialized s1 input1 #pm (fst v)); - unfold (pts_to_serialized s2 input2 #pm (snd v)); - join input1 input2 input; - rewrite (pts_to input #pm (bare_serialize s1 (fst v) `Seq.append` bare_serialize s2 (snd v))) - as (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) - }; - intro_stick - (pts_to_serialized s1 input1 #pm (fst v) ** pts_to_serialized s2 input2 #pm (snd v)) - (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) - (is_split input pm i input1 input2) - aux; - fold (split_nondep_then_post' s1 s2 input pm v input1 input2); - fold (split_nondep_then_post s1 s2 input pm v (input1, input2)); - (input1, input2) - } - } + Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); + pts_to_serialized_ext_stick + (tot_serialize_nondep_then s1 s2) + (tot_serialize_synth + (tot_parse_dtuple2 p1 (fun _ -> p2)) + pair_of_dtuple2 + (tot_serialize_dtuple2 s1 (fun _ -> s2)) + dtuple2_of_pair + () + ) + input; + pts_to_serialized_synth_l2r_stick + (tot_serialize_dtuple2 s1 (fun _ -> s2)) + pair_of_dtuple2 + dtuple2_of_pair + input; + stick_trans (pts_to_serialized (tot_serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; + let res = split_dtuple2 s1 j1 (fun _ -> s2) input; + match res { Mktuple2 input1 input2 -> { + unfold (split_dtuple2_post s1 (fun _ -> s2) input pm (dtuple2_of_pair v) res); + unfold (split_dtuple2_post' s1 (fun _ -> s2) input pm (dtuple2_of_pair v) input1 input2); + stick_trans (_ ** _) _ _; + fold (split_nondep_then_post' s1 s2 input pm v input1 input2); + fold (split_nondep_then_post s1 s2 input pm v (input1, input2)); + (input1, input2) + }} } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 9ce8e00fe..f4a238bb5 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -229,6 +229,52 @@ fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2 } ``` +let slice_append_split_stick_post' + (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + (s1: S.slice t) + (s2: S.slice t) +: Tot vprop += + S.pts_to s1 #p v1 ** + S.pts_to s2 #p v2 ** + ((S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) @==> S.pts_to s #p (v1 `Seq.append` v2)) + +let slice_append_split_stick_post + (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + (res: (S.slice t & S.slice t)) +: Tot vprop += let (s1, s2) = res in + slice_append_split_stick_post' s p v1 v2 i s1 s2 + +inline_for_extraction +```pulse +fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) + requires S.pts_to input #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) + returns res: (S.slice t & S.slice t) + ensures slice_append_split_stick_post input p v1 v2 i res +{ + let res = slice_append_split mutb input i; + match res { + Mktuple2 input1 input2 -> { + unfold (slice_append_split_post input p v1 v2 i res); + unfold (slice_append_split_post' input p v1 v2 i input1 input2); + ghost + fn aux + (_foo: unit) + requires S.is_split input p i input1 input2 ** (S.pts_to input1 #p v1 ** S.pts_to input2 #p v2) + ensures S.pts_to input #p (v1 `Seq.append` v2) + { + S.join input1 input2 input; + }; + intro_stick _ _ _ aux; + fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); + fold (slice_append_split_stick_post input p v1 v2 i (input1, input2)); + (input1, input2) + } + } +} +``` + let slice_split_stick_post' (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) @@ -278,3 +324,23 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos }} } ``` + +```pulse +ghost +fn rewrite_with_stick + (p1 p2: vprop) + requires p1 ** pure (p1 == p2) + ensures p2 ** (p2 @==> p1) +{ + rewrite p1 as p2; + ghost + fn aux + (_: unit) + requires emp ** p2 + ensures p1 + { + rewrite p2 as p1 + }; + intro_stick _ _ _ aux +} +``` From 53d8e8bd65772cd94ccd828d41811b24511d18df Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 10:33:19 -0700 Subject: [PATCH 11/89] jump_nondep_then, jump_dtuple2 --- .../pulse/LowParse.Pulse.Combinators.fst | 88 +++++++++++++++++++ 1 file changed, 88 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index d64743775..ee2c5f987 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -714,6 +714,94 @@ let validate_dtuple2 v (validate_dtuple2_payload r1 p2 v2 input offset pm v) +inline_for_extraction +```pulse +fn jump_nondep_then + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: jumper p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (v2: jumper p2) +: jumper #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + nondep_then_eq #k1 p1 #k2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + pts_to_len input; + let consumed1 = v1 input offset; + let consumed2 = v2 input (offset `SZ.add` consumed1); + SZ.add consumed1 consumed2 +} +``` + +inline_for_extraction +```pulse +fn jump_dtuple2_payload_cont + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> jumper (p2 x))) + (input: slice byte) + (offset: SZ.t) + (pm: perm) + (v: Ghost.erased bytes) + (off: SZ.t { jumper_pre (tot_parse_dtuple2 p1 p2) offset v /\ validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) + (input_key: slice byte) + (x: t1 { x == Ghost.reveal key }) + requires (pts_to_serialized s1 input_key #pm key ** (pts_to_serialized s1 input_key #pm key @==> pts_to input #pm v)) + returns res: SZ.t + ensures (pts_to input #pm v ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v res)) +{ + elim_stick (pts_to_serialized s1 input_key #pm key) _; + pts_to_len input; + let off2 = v2 x input (SZ.add offset off); + SZ.add off off2 +} +``` + +inline_for_extraction +```pulse +fn jump_dtuple2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: jumper p1) + (#s1: serializer p1) + (r1: reader s1) + (#k2: parser_kind) + (#p2: (x: t1) -> parser k2 (t2 x)) + (v2: (x: t1) -> jumper (p2 x)) +: jumper #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + pts_to_len input; + let off1 = v1 input offset; + let input1 = peek_stick_gen s1 input offset off1; + with key . assert (pts_to_serialized s1 input1 #pm key); + r1 + input1 + (pts_to_serialized s1 input1 #pm key @==> pts_to input #pm v) + SZ.t + (fun res -> pts_to input #pm v ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v res)) + (jump_dtuple2_payload_cont s1 p2 v2 input offset pm v off1 key input1) +} +``` + let split_dtuple2_post' (#t1: Type0) (#t2: t1 -> Type0) From 4fd8c754ff3e913ad6c6fe54438a0e8346f3d434 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 11:27:18 -0700 Subject: [PATCH 12/89] validate_and_read --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 125 +++++++++++++++++++++ 1 file changed, 125 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 2c2b0bfc3..3df15b085 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -585,3 +585,128 @@ let reader_of_leaf_reader (r: leaf_reader s) : reader s = reader_of_leaf_reader' r + +inline_for_extraction +let validate_and_read (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = + (input: slice byte) -> + (offset: SZ.t) -> + (#pm: perm) -> + (#v: Ghost.erased bytes) -> + (pre: vprop) -> + (t': Type0) -> + (post: (t' -> vprop)) -> + (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) -> + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -> + stt t' + (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) + (fun x -> post x) + +inline_for_extraction +```pulse +fn validate_and_read_intro_cont_read + (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (off: SZ.t) + (#v': Ghost.erased t) + (input': slice byte { validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (Ghost.reveal v', SZ.v off) }) + (x: t { x == Ghost.reveal v' }) + requires pts_to_serialized s input' #pm v' ** (pre ** (pts_to_serialized s input' #pm v' @==> pts_to input #pm v)) + returns x': t' + ensures post x' +{ + elim_stick _ _; + ksucc off x +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_intro_cont_validate + (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) (r: reader s) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (off: SZ.t) + requires pts_to input #pm v ** pre ** pure (validator_success p offset v off) + returns x' : t' + ensures post x' +{ + let input' = peek_stick_gen s input offset off; + with v' . assert (pts_to_serialized s input' #pm v'); + r input' #pm #v' _ _ _ (validate_and_read_intro_cont_read s input offset #pm #v pre t' post ksucc off #_ input') +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_intro + (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) (#s: serializer p) (r: reader s) +: validate_and_read #t #k p += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) +{ + w input offset _ _ _ (validate_and_read_intro_cont_validate r input offset #pm #v pre t' post ksucc) kfail +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_elim_cont + (#t: Type0) (#k: parser_kind) (p: parser k t) + (input: slice byte) + (offset: SZ.t) + (pm: perm) + (v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) + (off: SZ.t) + (x: t) + requires pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) + returns x' : t' + ensures post x' +{ + ksucc off +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_elim + (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validate_and_read p) +: validator #t #k p += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) +{ + w input offset pre t' post (validate_and_read_elim_cont p input offset pm v pre t' post ksucc) kfail +} +``` From ad3005835453206fb2e89e29ad3a494310468c28 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 14:40:50 -0700 Subject: [PATCH 13/89] validate_filter -> validate_and_read_filter --- .../pulse/LowParse.Pulse.Combinators.fst | 117 ++++-------------- 1 file changed, 23 insertions(+), 94 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index ee2c5f987..bc14ed550 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -207,13 +207,11 @@ fn pts_to_serialized_synth_l2r_stick inline_for_extraction ```pulse -fn validate_filter_gen_cont_success +fn validate_and_read_filter_cont_failure (#t: Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: serializer p) + (p: tot_parser k t) (f: (t -> bool)) - (f': (input: slice byte) -> (pm: perm) -> (v: Ghost.erased t) -> stt bool (pts_to_serialized s input #pm v) (fun res -> pts_to_serialized s input #pm v ** pure (res == f v))) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -221,41 +219,20 @@ fn validate_filter_gen_cont_success (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off)) (fun x -> post x))) (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) - (off: SZ.t) - requires (pts_to input #pm v ** pre ** pure (validator_success (p) offset v off)) + (_: unit) + requires (pts_to input #pm v ** pre ** pure (validator_failure (p) offset v)) returns res: t' ensures post res { - Seq.lemma_split v (SZ.v offset); - let split123 = split false input offset; - match split123 { Mktuple2 input1 input23 -> { - unfold (split_post input pm v offset split123); - unfold (split_post' input pm v offset input1 input23); - with v1 v23 . assert (pts_to input1 #pm v1 ** pts_to input23 #pm v23); - tot_parse_filter_eq p f v23; - let split23 = peek_stick s input23 off; - match split23 { Mktuple2 input2 input3 -> { - unfold (peek_stick_post s input23 pm v23 off split23); - unfold (peek_stick_post' s input23 pm v23 off input2 input3); - let cond = f' input2 pm _; - elim_stick (pts_to_serialized s input2 #pm _ ** _) _; - join input1 input23 input; - rewrite (pts_to input #pm (Seq.append v1 v23)) as (pts_to input #pm v); - if cond { - ksucc off - } else { - kfail () - } - }} - }} + tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () } ``` inline_for_extraction ```pulse -fn validate_filter_gen_cont_failure +fn validate_and_read_filter_cont_success (#t: Type0) (#k: parser_kind) (p: tot_parser k t) @@ -267,28 +244,32 @@ fn validate_filter_gen_cont_failure (pre: vprop) (t': Type0) (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: parse_filter_refine f) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off /\ parse (tot_parse_filter p f) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) - (_: unit) - requires (pts_to input #pm v ** pre ** pure (validator_failure (p) offset v)) + (off: SZ.t) + (x: t) + requires pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) returns res: t' ensures post res { tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); - kfail () + if f x { + ksucc off x + } else { + kfail () + } } ``` inline_for_extraction ```pulse -fn validate_filter_gen +fn validate_and_read_filter (#t: Type0) (#k: parser_kind) (#p: tot_parser k t) - (s: serializer p) - (w: validator p) + (w: validate_and_read p) (f: (t -> bool)) - (f': (input: slice byte) -> (pm: perm) -> (v: Ghost.erased t) -> stt bool (pts_to_serialized s input #pm v) (fun res -> pts_to_serialized s input #pm v ** pure (res == f v))) -: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) +: validate_and_read #_ #(parse_filter_kind k) (tot_parse_filter p f) = (input: slice byte) (offset: SZ.t) @@ -297,67 +278,15 @@ fn validate_filter_gen (pre: vprop) (t': Type0) (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off)) (fun x -> post x))) + (ksucc: ((off: SZ.t) -> (x: parse_filter_refine f) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off /\ parse (tot_parse_filter p f) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) { - w input offset pre t' post - (validate_filter_gen_cont_success s f f' input offset #pm #v pre t' post ksucc kfail) - (validate_filter_gen_cont_failure p f input offset #pm #v pre t' post kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_filter_read_cond_cont - (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) - (s: serializer p) - (f: (t -> bool)) - (input: slice byte) - (pm: perm) - (v: Ghost.erased t) - (x: t { x == Ghost.reveal v }) - requires (pts_to_serialized s input #pm v ** emp) - returns res: bool - ensures pts_to_serialized s input #pm v ** pure (res == f v) -{ - f x -} -``` - -inline_for_extraction -```pulse -fn validate_filter_read_cond - (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) - (#s: serializer p) - (r: reader s) - (f: (t -> bool)) - (input: slice byte) - (pm: perm) - (v: Ghost.erased t) - requires (pts_to_serialized s input #pm v) - returns res: bool - ensures pts_to_serialized s input #pm v ** pure (res == f v) -{ - r input #pm #v emp bool (fun res -> pts_to_serialized s input #pm v ** pure (res == f v)) (validate_filter_read_cond_cont s f input pm v) + w input offset #pm #v pre t' post + (validate_and_read_filter_cont_success p f input offset #pm #v pre t' post ksucc kfail) + (validate_and_read_filter_cont_failure p f input offset #pm #v pre t' post kfail) } ``` -inline_for_extraction -let validate_filter - (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) - (#s: serializer p) - (r: reader s) - (w: validator p) - (f: (t -> bool)) -: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) -= validate_filter_gen s w f (validate_filter_read_cond r f) - inline_for_extraction ```pulse fn jump_filter From dc5e56b72dcefa982b9fa2552136a87635ddacc5 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 15:17:30 -0700 Subject: [PATCH 14/89] validate_dtuple2 with validate_and_read tag, remove validate_nondep_then --- .../pulse/LowParse.Pulse.Combinators.fst | 145 ++---------------- 1 file changed, 14 insertions(+), 131 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index bc14ed550..f5c461eaa 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -452,23 +452,22 @@ fn validate_dtuple2_cont_success1 (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) - (v2: (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) -> (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) -> - validator_for (p2 key) input (SZ.add offset off) pm v) + (v2: (x: t1) -> validator (p2 x)) (pre: vprop) (t': Type0) (post: (t' -> vprop)) (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) (off: SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off) + (x: t1) + requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) returns x: t' ensures post x { pts_to_len input; // for SZ.fits (offset + off) - let key : Ghost.erased t1 = Ghost.hide (fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v))))); - v2 off key pre t' post - (validate_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post key off ksucc) - (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post key off kfail) + v2 x input (offset `SZ.add` off) #pm #v pre t' post + (validate_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post x off ksucc) + (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post x off kfail) } ``` @@ -500,22 +499,21 @@ fn validate_dtuple2_cont_failure1 inline_for_extraction ```pulse -fn validate_dtuple2_gen +fn validate_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) (#k1: parser_kind) (#p1: parser k1 t1) - (v1: validator p1) + (v1: validate_and_read p1) (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) + (#p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> validator (p2 x))) +: validator #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) += (input: slice byte) (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) - (v2: (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) -> (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) -> - validator_for (p2 key) input (SZ.add offset off) pm v) -: validator_for #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) input offset pm v -= + (#pm: perm) + (#v: Ghost.erased bytes) (pre: vprop) (t': Type0) (post: (t' -> vprop)) @@ -528,121 +526,6 @@ fn validate_dtuple2_gen } ``` -inline_for_extraction -let validate_nondep_then - (#t1 #t2: Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (v1: validator p1) - (#k2: parser_kind) - (#p2: parser k2 t2) - (v2: validator p2) -: validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) -= Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); - validate_ext - (validate_synth - (fun input offset #pm #v -> - validate_dtuple2_gen - v1 - (fun _ -> p2) - input - offset - pm - v - (fun off _ -> v2 input (SZ.add offset off) #pm #v) - ) - pair_of_dtuple2 - ) - (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) - -inline_for_extraction -```pulse -fn validate_dtuple2_payload_cont - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (s1: serializer p1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (v2: ((x: t1) -> validator (p2 x))) - (input: slice byte) - (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) - (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off': SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (p2 key) (SZ.add offset off) v off')) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (SZ.add offset off) v)) (fun x -> post x))) - (input_key: slice byte) - (x: t1 { x == Ghost.reveal key }) - requires (pts_to_serialized s1 input_key #pm key ** (pre ** (pts_to_serialized s1 input_key #pm key @==> pts_to input #pm v))) - returns res: t' - ensures post res -{ - elim_stick (pts_to_serialized s1 input_key #pm key) _; - v2 x input (SZ.add offset off) pre t' post ksucc kfail -} -``` - -inline_for_extraction -```pulse -fn validate_dtuple2_payload - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (#s1: serializer p1) - (r1: reader s1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (v2: ((x: t1) -> validator (p2 x))) - (input: slice byte) - (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) - (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) -: validator_for #(t2 key) #k2 (p2 key) input (SZ.add offset off) pm v -= - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off': SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (p2 key) (SZ.add offset off) v off')) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (SZ.add offset off) v)) (fun x -> post x))) -{ - let input' = peek_stick_gen s1 input offset off; - assert (pts_to_serialized s1 input' #pm key); - r1 input' _ t' post (validate_dtuple2_payload_cont s1 p2 v2 input offset pm v off key pre t' post ksucc kfail input') -} -``` - -inline_for_extraction -let validate_dtuple2 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (v1: validator p1) - (#s1: serializer p1) - (r1: reader s1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (v2: ((x: t1) -> validator (p2 x))) -: Tot (validator (tot_parse_dtuple2 p1 p2)) -= fun input offset #pm #v -> - validate_dtuple2_gen - v1 - p2 - input - offset - pm - v - (validate_dtuple2_payload r1 p2 v2 input offset pm v) - inline_for_extraction ```pulse fn jump_nondep_then From fd887c5f3488da0fb2facbabe3ca823cfe3bd159 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 15:21:29 -0700 Subject: [PATCH 15/89] restore validate_nondep_then --- .../pulse/LowParse.Pulse.Combinators.fst | 139 ++++++++++++++++++ 1 file changed, 139 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index f5c461eaa..94ac338ac 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -382,6 +382,145 @@ let nondep_then_eq_dtuple2 = tot_parse_synth_eq (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 input +inline_for_extraction +```pulse +fn validate_nondep_then_cont_success2 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (off': SZ.t) + requires pts_to input #pm v ** pre ** pure (validator_success p2 (offset `SZ.add` off) v off') + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (off + off') + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + ksucc (off `SZ.add` off') +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_failure2 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (_: unit) + requires pts_to input #pm v ** pre ** pure (validator_failure p2 (offset `SZ.add` off) v) + returns x: t' + ensures post x +{ + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_success1 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (v2: validator p2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (off: SZ.t) + requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off) + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (offset + off) + v2 input (offset `SZ.add` off) #pm #v pre t' post + (validate_nondep_then_cont_success2 p1 p2 input offset #pm #v pre t' post off ksucc) + (validate_nondep_then_cont_failure2 p1 p2 input offset #pm #v pre t' post off kfail) +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then_cont_failure1 + (#t1 #t2: Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: parser k2 t2) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) + (_: unit) + requires pts_to input #pm v ** pre ** pure (validator_failure p1 offset v) + returns x: t' + ensures post x +{ + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + kfail () +} +``` + +inline_for_extraction +```pulse +fn validate_nondep_then + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: validator p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (v2: validator p2) +: validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) +{ + v1 input offset #pm #v pre t' post + (validate_nondep_then_cont_success1 p1 v2 input offset #pm #v pre t' post ksucc kfail) + (validate_nondep_then_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) +} +``` + inline_for_extraction ```pulse fn validate_dtuple2_cont_success2 From cb95617d57590509511378e7a5c22844a65bcc43 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 15:36:53 -0700 Subject: [PATCH 16/89] validate_and_read_dtuple2 --- .../pulse/LowParse.Pulse.Combinators.fst | 91 +++++++++++++++++++ 1 file changed, 91 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 94ac338ac..8692d91c4 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -665,6 +665,97 @@ fn validate_dtuple2 } ``` +inline_for_extraction +```pulse +fn validate_and_read_dtuple2_cont_success2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (key: t1) + (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == key }) + (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (off': SZ.t) + (x: t2 key) + requires pts_to input #pm v ** pre ** pure (validator_success (p2 key) (offset `SZ.add` off) v off' /\ parse (p2 key) (Seq.slice v (SZ.v (offset `SZ.add` off)) (Seq.length v)) == Some (x, SZ.v off')) + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (off + off') + ksucc (off `SZ.add` off') (| key, x |) +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_dtuple2_cont_success1 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (p1: parser k1 t1) + (#k2: parser_kind) + (p2: ((x: t1) -> parser k2 (t2 x))) + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (v2: (x: t1) -> validate_and_read (p2 x)) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) + (off: SZ.t) + (x: t1) + requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) + returns x: t' + ensures post x +{ + pts_to_len input; // for SZ.fits (offset + off) + v2 x input (offset `SZ.add` off) #pm #v pre t' post + (validate_and_read_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post x off ksucc) + (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post x off kfail) +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_dtuple2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (v1: validate_and_read p1) + (#k2: parser_kind) + (#p2: ((x: t1) -> parser k2 (t2 x))) + (v2: ((x: t1) -> validate_and_read (p2 x))) +: validate_and_read #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) + (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) +{ + v1 input offset #pm #v pre t' post + (validate_and_read_dtuple2_cont_success1 p1 p2 input offset #pm #v v2 pre t' post ksucc kfail) + (validate_dtuple2_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) +} +``` + inline_for_extraction ```pulse fn jump_nondep_then From 14f5bace7062178d024f025149de4d0034707039 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 9 Jul 2024 17:28:01 -0700 Subject: [PATCH 17/89] read_synth, read_filter --- .../pulse/LowParse.Pulse.Combinators.fst | 88 ++++++++++++++++++- 1 file changed, 87 insertions(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 8692d91c4..0865e7837 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -205,6 +205,76 @@ fn pts_to_serialized_synth_l2r_stick } ``` +inline_for_extraction +let stt_cps + (#t: Type) + (y: t) +: Tot Type += (pre: vprop) -> (t': Type0) -> (post: (t' -> vprop)) -> (phi: ((y': t { y' == y }) -> stt t' pre (fun x -> post x))) -> stt t' pre (fun x -> post x) + +inline_for_extraction +let intro_stt_cps + (#t: Type) + (y: t) +: Tot (stt_cps y) += fun pre t' post phi -> phi y + +inline_for_extraction +let elim_stt_cps + (#t: Type) + (y: Ghost.erased t) + (cps: stt_cps (Ghost.reveal y)) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (phi: ((y': t { y' == Ghost.reveal y }) -> stt t' pre (fun x -> post x))) +: stt t' pre (fun x -> post x) += cps pre t' post phi + +inline_for_extraction +```pulse +fn read_synth_cont + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (f2': (x: t1) -> stt_cps (f2 x)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased t2) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (cont: (x: t2 { x == Ghost.reveal v }) -> stt t' (pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v ** pre) (fun x -> post x)) + (x: t1 { x == Ghost.reveal (f1 v) }) + requires pts_to_serialized s1 input #pm (f1 v) ** (pre ** (pts_to_serialized s1 input #pm (f1 v) @==> pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v)) + returns x': t' + ensures post x' +{ + elim_stick _ _; + f2' x (pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v ** pre) t' post cont +} +``` + +inline_for_extraction +```pulse +fn read_synth + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (f2': (x: t1) -> stt_cps (f2 x)) +: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) += (input: slice byte) + (#pm: _) + (#v: _) + (pre: _) + (t': Type0) + (post: _) + (cont: _) +{ + pts_to_serialized_synth_l2r_stick s1 f2 f1 input; + r input #pm #_ (pre ** (pts_to_serialized s1 input #pm (f1 v) @==> pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v)) t' post + (read_synth_cont r f2 f1 f2' input pm v pre t' post cont) +} +``` + inline_for_extraction ```pulse fn validate_and_read_filter_cont_failure @@ -355,7 +425,23 @@ fn pts_to_serialized_filter_elim } ``` - +inline_for_extraction +```pulse +fn read_filter + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (f: (t1 -> bool)) +: reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) += (input: slice byte) + (#pm: _) + (#v: _) + (pre: _) + (t': Type0) + (post: _) + (cont: _) +{ + pts_to_serialized_filter_elim s1 f input; + r input #pm #(Ghost.hide (Ghost.reveal v)) pre t' post cont +} +``` let pair_of_dtuple2 (#t1 #t2: Type) From b44c18b350c68fc11e4dd9c2a6a8776da468d38a Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 10 Jul 2024 09:22:39 -0700 Subject: [PATCH 18/89] validate_and_read_synth --- .../pulse/LowParse.Pulse.Combinators.fst | 75 ++++++++++++++++++- 1 file changed, 73 insertions(+), 2 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 0865e7837..933944a47 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -234,7 +234,7 @@ let elim_stt_cps inline_for_extraction ```pulse fn read_synth_cont - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (s1: serializer p1) (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) (f2': (x: t1) -> stt_cps (f2 x)) (input: slice byte) @@ -271,7 +271,78 @@ fn read_synth { pts_to_serialized_synth_l2r_stick s1 f2 f1 input; r input #pm #_ (pre ** (pts_to_serialized s1 input #pm (f1 v) @==> pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v)) t' post - (read_synth_cont r f2 f1 f2' input pm v pre t' post cont) + (read_synth_cont s1 f2 f1 f2' input pm v pre t' post cont) +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_synth_cont_cont + (#k1: parser_kind) (#t1: Type0) (p1: parser k1 t1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) + (input: slice byte) + (pm: perm) + (v: Ghost.erased bytes) + (offset: SZ.t) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (cont: (off: SZ.t) -> (x: t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_synth p1 f2) offset v off /\ parse (tot_parse_synth p1 f2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x)) + (off: SZ.t) + (x: t1 { validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off) }) + (x2: t2 { x2 == f2 x }) + requires pts_to input #pm v ** pre + returns x': t' + ensures post x' +{ + tot_parse_synth_eq p1 f2 (Seq.slice v (SZ.v offset) (Seq.length v)); + cont off x2 +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_synth_cont + (#k1: parser_kind) (#t1: Type0) (p1: parser k1 t1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) + (f2': (x: t1) -> stt_cps (f2 x)) + (input: slice byte) + (pm: perm) + (v: Ghost.erased bytes) + (offset: SZ.t) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (cont: (off: SZ.t) -> (x: t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_synth p1 f2) offset v off /\ parse (tot_parse_synth p1 f2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x)) + (off: SZ.t) + (x: t1) + requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) + returns x': t' + ensures post x' +{ + f2' x (pts_to input #pm v ** pre) t' post (validate_and_read_synth_cont_cont p1 f2 input pm v offset pre t' post cont off x) +} +``` + +inline_for_extraction +```pulse +fn validate_and_read_synth + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (w: validate_and_read p1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) + (f2': (x: t1) -> stt_cps (f2 x)) +: validate_and_read #t2 #k1 (tot_parse_synth p1 f2) += (input: slice byte) + (offset: _) + (#pm: _) + (#v: _) + (pre: _) + (t': Type0) + (post: _) + (ksucc: _) + (kfail: _) +{ + tot_parse_synth_eq p1 f2 (Seq.slice v (SZ.v offset) (Seq.length v)); + w input offset #pm #v pre t' post (validate_and_read_synth_cont p1 f2 f2' input pm v offset pre t' post ksucc) kfail } ``` From 9383b8c4aa9f3862f6199887e496a3d7f1712cf0 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 10 Jul 2024 11:00:30 -0700 Subject: [PATCH 19/89] bitsum --- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 110 ++++++++++++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 16 ++- 2 files changed, 124 insertions(+), 2 deletions(-) create mode 100644 src/lowparse/pulse/LowParse.Pulse.BitSum.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst new file mode 100644 index 000000000..3bc18c894 --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -0,0 +1,110 @@ +module LowParse.Pulse.BitSum +include LowParse.Spec.BitSum +include LowParse.Pulse.Combinators +open FStar.Tactics.V2 +open LowParse.Pulse.Util +open Pulse.Lib.Stick +open Pulse.Lib.Slice + +#push-options "--print_universes" + +inline_for_extraction +```pulse +fn stt_cps_ifthenelse + (#t: Type0) + (y: Ghost.erased t) +: if_combinator_weak u#4 (stt_cps u#0 #t y) += (cond: _) + (ftrue: _) + (ffalse: _) + (pre: _) + (t': Type0) + (post: _) + (y': _) +{ + if cond { + ftrue () pre t' post y' + } else { + ffalse () pre t' post y' + } +} +``` + +inline_for_extraction +let validate_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (#b: bitsum' cl tot) + (f: filter_bitsum'_t b) + (d: destr_bitsum'_t b) + (#k: parser_kind) + (#p: parser k t) + (w: validate_and_read p) +: Pure (validate_and_read (tot_parse_bitsum' b p)) + (requires (k.parser_kind_subkind == Some ParserStrong)) + (ensures (fun _ -> True)) += [@@inline_let] + let _ = synth_bitsum'_injective b in + (validate_and_read_synth + (validate_and_read_filter + w + (filter_bitsum' b) + (fun x -> f x) + ) + (synth_bitsum' b) + (d + stt_cps + stt_cps_ifthenelse + (fun k pre t' post phi -> phi k) + ) + ) + +inline_for_extraction +let jump_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (b: bitsum' cl tot) + (#k: parser_kind) + (#p: parser k t) + (w: jumper p) +: Tot (jumper (tot_parse_bitsum' b p)) += [@@inline_let] + let _ = synth_bitsum'_injective b in + jump_synth + (jump_filter + w + (filter_bitsum' b) + ) + (synth_bitsum' b) + +inline_for_extraction +let read_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (#b: bitsum' cl tot) + (d: destr_bitsum'_t b) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (r: reader s) +: Tot (reader (tot_serialize_bitsum' b s)) += [@@inline_let] + let _ = synth_bitsum'_injective b in + [@@inline_let] + let _ = synth_bitsum'_recip_inverse b in + read_synth + (read_filter + r + (filter_bitsum' b) + ) + (synth_bitsum' b) + (synth_bitsum'_recip b) + ( + d + stt_cps + stt_cps_ifthenelse + (fun k pre t' post phi -> phi k) + ) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 933944a47..139859615 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -378,6 +378,7 @@ fn validate_and_read_filter_cont_success (#k: parser_kind) (p: tot_parser k t) (f: (t -> bool)) + (f': (x: t) -> (y: bool { y == f x })) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -394,7 +395,7 @@ fn validate_and_read_filter_cont_success ensures post res { tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); - if f x { + if f' x { ksucc off x } else { kfail () @@ -410,6 +411,7 @@ fn validate_and_read_filter (#p: tot_parser k t) (w: validate_and_read p) (f: (t -> bool)) + (f': (x: t) -> (y: bool { y == f x })) : validate_and_read #_ #(parse_filter_kind k) (tot_parse_filter p f) = (input: slice byte) @@ -423,11 +425,21 @@ fn validate_and_read_filter (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) { w input offset #pm #v pre t' post - (validate_and_read_filter_cont_success p f input offset #pm #v pre t' post ksucc kfail) + (validate_and_read_filter_cont_success p f f' input offset #pm #v pre t' post ksucc kfail) (validate_and_read_filter_cont_failure p f input offset #pm #v pre t' post kfail) } ``` +inline_for_extraction +let validate_and_read_filter' + (#t: Type0) + (#k: parser_kind) + (#p: tot_parser k t) + (w: validate_and_read p) + (f: (t -> bool)) +: validate_and_read #_ #(parse_filter_kind k) (tot_parse_filter p f) += validate_and_read_filter w f (fun x -> f x) + inline_for_extraction ```pulse fn jump_filter From 4b29d00250d6b72c2578ce1cdd8c299e00473e3b Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 10:25:37 -0700 Subject: [PATCH 20/89] validate_and_read integers; ret, empty; etc. --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 49 ++++++++++++ src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 2 +- .../pulse/LowParse.Pulse.Combinators.fst | 76 +++++++++++++++++++ src/lowparse/pulse/LowParse.Pulse.Int.fst | 16 ++++ 4 files changed, 142 insertions(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 3df15b085..fd6f110dd 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -221,6 +221,10 @@ inline_for_extraction let validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = v1 +inline_for_extraction +let validate_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validator p1 { k2 `is_weaker_than` k1 }) : validator (tot_weaken k2 p1) = + validate_ext v1 (tot_weaken k2 p1) + inline_for_extraction ```pulse fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { @@ -601,6 +605,27 @@ let validate_and_read (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) (fun x -> post x) +inline_for_extraction +```pulse +fn validate_and_read_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validate_and_read #_ #k2 p2 = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (ksucc: _) + (kfail: _) +{ + v1 input offset #pm #v pre t' post ksucc kfail +} +``` + +inline_for_extraction +let validate_and_read_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1 { k2 `is_weaker_than` k1 }) : validate_and_read (tot_weaken k2 p1) = + validate_and_read_ext v1 (tot_weaken k2 p1) + inline_for_extraction ```pulse fn validate_and_read_intro_cont_read @@ -710,3 +735,27 @@ fn validate_and_read_elim w input offset pre t' post (validate_and_read_elim_cont p input offset pm v pre t' post ksucc) kfail } ``` + +inline_for_extraction +```pulse +fn ifthenelse_validate_and_read + (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) (wtrue: squash (cond == true) -> validate_and_read p) (wfalse: squash (cond == false) -> validate_and_read p) +: validate_and_read #t #k p += + (input: _) + (offset: _) + (#pm: _) + (#v: _) + (pre: _) + (t': Type0) + (post: _) + (ksucc: _) + (kfail: _) +{ + if cond { + wtrue () input offset pre t' post ksucc kfail + } else { + wfalse () input offset pre t' post ksucc kfail + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index 3bc18c894..70c225ce9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -88,7 +88,7 @@ let read_bitsum' (d: destr_bitsum'_t b) (#k: parser_kind) (#p: parser k t) - (s: serializer p) + (#s: serializer p) (r: reader s) : Tot (reader (tot_serialize_bitsum' b s)) = [@@inline_let] diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 139859615..d09c7f31d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -8,6 +8,68 @@ open Pulse.Lib.Slice module SZ = FStar.SizeT +inline_for_extraction +```pulse +fn validate_ret + (#t: Type0) + (x: t) +: validator #t #parse_ret_kind (tot_parse_ret x) += (input: slice byte) + (offset: _) + (#pm: _) + (#v: _) + (pre: _) + (t': Type0) + (post: _) + (ksucc: _) + (kfail: _) +{ + ksucc 0sz +} +``` + +inline_for_extraction +```pulse +fn leaf_read_ret + (#t: Type0) + (x: t) + (v_unique: (v' : t) -> Lemma (x == v')) +: leaf_reader #t #parse_ret_kind #(tot_parse_ret x) (tot_serialize_ret x v_unique) += (input: slice byte) + (#pm: _) + (#v: _) +{ + v_unique v; + x +} +``` + +inline_for_extraction +let read_ret + (#t: Type0) + (x: t) + (v_unique: (v' : t) -> Lemma (x == v')) +: Tot (reader (tot_serialize_ret x v_unique)) += reader_of_leaf_reader (leaf_read_ret x v_unique) + +inline_for_extraction +let validate_and_read_ret + (#t: Type0) + (x: t) + (v_unique: (v' : t) -> Lemma (x == v')) +: Tot (validate_and_read (tot_parse_ret x)) += validate_and_read_intro (validate_ret x) (read_ret x v_unique) + +inline_for_extraction +let validate_empty : validator tot_parse_empty = validate_ret () + +inline_for_extraction +let read_empty : reader tot_serialize_empty = read_ret () (fun _ -> ()) + +inline_for_extraction +let validate_and_read_empty : validate_and_read tot_parse_empty = + validate_and_read_intro validate_empty read_empty + let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) @@ -275,6 +337,13 @@ fn read_synth } ``` +inline_for_extraction +let read_synth' + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) +: Tot (reader (tot_serialize_synth p1 f2 s1 f1 ())) += read_synth r f2 f1 (fun x -> intro_stt_cps (f2 x)) + inline_for_extraction ```pulse fn validate_and_read_synth_cont_cont @@ -346,6 +415,13 @@ fn validate_and_read_synth } ``` +inline_for_extraction +let validate_and_read_synth' + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (w: validate_and_read p1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) +: Tot (validate_and_read (tot_parse_synth p1 f2)) += validate_and_read_synth w f2 (fun x -> intro_stt_cps (f2 x)) + inline_for_extraction ```pulse fn validate_and_read_filter_cont_failure diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 30452d13f..144fac10a 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -38,6 +38,10 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 inline_for_extraction let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) +inline_for_extraction +let validate_and_read_u8 : validate_and_read tot_parse_u8 = + validate_and_read_intro validate_u8 read_u8 + inline_for_extraction let validate_u16 : validator tot_parse_u16 = validate_total_constant_size tot_parse_u16 2sz @@ -68,6 +72,10 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_ inline_for_extraction let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) +inline_for_extraction +let validate_and_read_u16 : validate_and_read tot_parse_u16 = + validate_and_read_intro validate_u16 read_u16 + inline_for_extraction let validate_u32 : validator tot_parse_u32 = validate_total_constant_size tot_parse_u32 4sz @@ -98,6 +106,10 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_ inline_for_extraction let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) +inline_for_extraction +let validate_and_read_u32 : validate_and_read tot_parse_u32 = + validate_and_read_intro validate_u32 read_u32 + inline_for_extraction let validate_u64 : validator tot_parse_u64 = validate_total_constant_size tot_parse_u64 8sz @@ -127,3 +139,7 @@ fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_ inline_for_extraction let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) + +inline_for_extraction +let validate_and_read_u64 : validate_and_read tot_parse_u64 = + validate_and_read_intro validate_u64 read_u64 From 944b32de821267b74597aaf62fcdb61d8eb43da8 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 11:48:16 -0700 Subject: [PATCH 21/89] ifthenelse_jumper, jump_ret, jump_empty --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 18 ++++++++++++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 6 ++++++ 2 files changed, 24 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index fd6f110dd..1827aa121 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -284,6 +284,24 @@ let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (pts_to input #pm v ** pure (jumper_pre' p offset v)) (fun res -> pts_to input #pm v ** pure (validator_success p offset v res)) +inline_for_extraction +```pulse +fn ifthenelse_jumper (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) (jtrue: squash (cond == true) -> jumper p) (jfalse: squash (cond == false) -> jumper p) +: jumper #t #k p += + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + if cond { + jtrue () input offset + } else { + jfalse () input offset + } +} +``` + inline_for_extraction let jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = v1 diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index d09c7f31d..ead85534b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -60,6 +60,9 @@ let validate_and_read_ret : Tot (validate_and_read (tot_parse_ret x)) = validate_and_read_intro (validate_ret x) (read_ret x v_unique) +inline_for_extraction +let jump_ret (#t: Type) (x: t) : jumper (tot_parse_ret x) = jump_constant_size (tot_parse_ret x) 0sz + inline_for_extraction let validate_empty : validator tot_parse_empty = validate_ret () @@ -70,6 +73,9 @@ inline_for_extraction let validate_and_read_empty : validate_and_read tot_parse_empty = validate_and_read_intro validate_empty read_empty +inline_for_extraction +let jump_empty : jumper tot_parse_empty = jump_ret () + let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) From 769d3b11373188702cf56f6d5cac3f0627f55cc0 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 15:53:03 -0700 Subject: [PATCH 22/89] some extraction fixes: avoid pure lets in F* assemblies, replace record field accesses, normalize with postprocess_with, etc. --- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 38 +++++++++++++------ .../pulse/LowParse.Pulse.Combinators.fst | 19 +++++++--- .../pulse/LowParse.Pulse.Endianness.fst | 8 ++-- src/lowparse/pulse/LowParse.Pulse.Int.fst | 3 +- 4 files changed, 46 insertions(+), 22 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index 70c225ce9..c57fd5104 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -70,14 +70,32 @@ let jump_bitsum' (#p: parser k t) (w: jumper p) : Tot (jumper (tot_parse_bitsum' b p)) -= [@@inline_let] - let _ = synth_bitsum'_injective b in - jump_synth += jump_synth (jump_filter w (filter_bitsum' b) ) - (synth_bitsum' b) + (synth_bitsum'_injective b; synth_bitsum' b) + +inline_for_extraction +```pulse +fn read_bitsum'_cont + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (b: bitsum' cl tot) + (k_: bitsum'_type b) + (pre: vprop) + (t': Type0) + (post: (t' -> vprop)) + (phi: (k': bitsum'_type b { k' == k_ }) -> stt t' pre (fun x -> post x)) + requires pre + returns x: t' + ensures post x +{ + phi k_ +} +``` inline_for_extraction let read_bitsum' @@ -91,20 +109,16 @@ let read_bitsum' (#s: serializer p) (r: reader s) : Tot (reader (tot_serialize_bitsum' b s)) -= [@@inline_let] - let _ = synth_bitsum'_injective b in - [@@inline_let] - let _ = synth_bitsum'_recip_inverse b in - read_synth += read_synth (read_filter r (filter_bitsum' b) ) - (synth_bitsum' b) - (synth_bitsum'_recip b) + (synth_bitsum'_injective b; synth_bitsum' b) + (synth_bitsum'_recip_inverse b; synth_bitsum'_recip b) ( d stt_cps stt_cps_ifthenelse - (fun k pre t' post phi -> phi k) + (read_bitsum'_cont b) ) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index ead85534b..a71787d92 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -115,15 +115,24 @@ let validate_synth w inline_for_extraction -let jump_synth - (#t #t': Type) +```pulse +fn jump_synth + (#t #t': Type0) (#k: parser_kind) (#p: tot_parser k t) (w: jumper p) (f: (t -> t') { synth_injective f }) -: Tot (jumper (tot_parse_synth p f)) -= Classical.forall_intro (tot_parse_synth_eq' p f); - w +: jumper #t' #k (tot_parse_synth #k #t p f) += + (input: _) + (offset: _) + (#pm: _) + (#v: _) +{ + tot_parse_synth_eq' p f (Seq.slice v (SZ.v offset) (Seq.length v)); + w input offset #pm #v +} +``` ```pulse ghost diff --git a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst index 2124abe49..abde8470a 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst @@ -44,7 +44,7 @@ fn be_to_n_0 (pos: SZ.t) { E.reveal_be_to_n (Seq.slice (v) 0 0); - u.zero + UIntType?.zero u } ``` @@ -67,7 +67,7 @@ fn be_to_n_1 E.reveal_be_to_n (Seq.slice (v) 0 1); E.reveal_be_to_n (Seq.slice (v) 0 0); let last = S.op_Array_Access x 0sz; - (u.from_byte last) + UIntType?.from_byte u last } ``` @@ -95,8 +95,8 @@ fn be_to_n_S let pos' = pos `SZ.sub` 1sz; let last = S.op_Array_Access x pos'; let n = ih x #pm #v pos'; - let blast = u.from_byte last; - u.add blast (u.mul256 n) + let blast = UIntType?.from_byte u last; + UIntType?.add u blast (u.mul256 n) } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 144fac10a..f31f2bc14 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -18,7 +18,8 @@ let jump_u8 : jumper tot_parse_u8 = inline_for_extraction noextract -let be_to_n_1 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint8 1) +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let be_to_n_1 = (E.mk_be_to_n EI.uint8 1) inline_for_extraction ```pulse From 376403df72b745a0d054d1ca89dad9ca178d51c4 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 16:57:25 -0700 Subject: [PATCH 23/89] mark more combinators inline_for_extraction --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 15 +++++++++++++-- src/lowparse/pulse/LowParse.Pulse.Combinators.fst | 1 + 2 files changed, 14 insertions(+), 2 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 1827aa121..37b066e87 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -303,8 +303,16 @@ fn ifthenelse_jumper (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) ``` inline_for_extraction -let jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = - v1 +```pulse +fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = + (input: slice byte) + (offset: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + v1 input offset #pm #v +} +``` inline_for_extraction ```pulse @@ -348,6 +356,7 @@ let peek_post = let (left, right) = res in peek_post' s input pm v consumed left right +inline_for_extraction ```pulse fn peek (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) @@ -402,6 +411,7 @@ let peek_stick_post = let (left, right) = res in peek_stick_post' s input pm v consumed left right +inline_for_extraction ```pulse fn peek_stick (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) @@ -435,6 +445,7 @@ fn peek_stick } ``` +inline_for_extraction ```pulse fn peek_stick_gen (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index a71787d92..7cfae59d5 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -551,6 +551,7 @@ fn jump_filter } ``` +inline_for_extraction let parse_filter_refine_intro (#t: Type) (f: (t -> GTot bool)) From 9d360cbc127d44fa59a06a5ded18ee1b2849838e Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 18:02:49 -0700 Subject: [PATCH 24/89] Karamel chokes on tuples. Use a custom slice_pair type instead --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 32 ++++----- .../pulse/LowParse.Pulse.Combinators.fst | 24 +++---- src/lowparse/pulse/LowParse.Pulse.Util.fst | 67 +++++++++++++------ 3 files changed, 73 insertions(+), 50 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 37b066e87..c234c1a78 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -351,9 +351,9 @@ let peek_post (pm: perm) (v: bytes) (consumed: SZ.t) - (res: (slice byte & slice byte)) + (res: slice_pair byte) : Tot vprop -= let (left, right) = res in += let SlicePair left right = res in peek_post' s input pm v consumed left right inline_for_extraction @@ -365,22 +365,22 @@ fn peek (#v: Ghost.erased bytes) (consumed: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) - returns res: (slice byte & slice byte) + returns res: slice_pair byte ensures peek_post s input pm v consumed res { - let s1s2 = split false input #pm #v consumed; + let s1s2 = slice_split false input #pm #v consumed; match s1s2 { - Mktuple2 s1 s2 -> { + SlicePair s1 s2 -> { Seq.lemma_split v (SZ.v consumed); let v1 = Ghost.hide (fst (Some?.v (parse p v))); parse_injective #k p (bare_serialize s v1) v; - unfold (split_post input pm v consumed (s1, s2)); + unfold (split_post0 input pm v consumed (SlicePair s1 s2)); unfold (split_post' input pm v consumed s1 s2); with v1' . assert (pts_to s1 #pm v1'); rewrite (pts_to s1 #pm v1') as (pts_to_serialized s s1 #pm v1); fold (peek_post' s input pm v consumed s1 s2); - fold (peek_post s input pm v consumed (s1, s2)); - (s1, s2) + fold (peek_post s input pm v consumed (SlicePair s1 s2)); + (SlicePair s1 s2) } } } @@ -406,9 +406,9 @@ let peek_stick_post (pm: perm) (v: bytes) (consumed: SZ.t) - (res: (slice byte & slice byte)) + (res: slice_pair byte) : Tot vprop -= let (left, right) = res in += let (SlicePair left right) = res in peek_stick_post' s input pm v consumed left right inline_for_extraction @@ -420,11 +420,11 @@ fn peek_stick (#v: Ghost.erased bytes) (consumed: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) - returns res: (slice byte & slice byte) + returns res: (slice_pair byte) ensures peek_stick_post s input pm v consumed res { let res = peek s input consumed; - match res { Mktuple2 left right -> { + match res { SlicePair left right -> { unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); @@ -439,8 +439,8 @@ fn peek_stick }; intro_stick (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) aux; fold (peek_stick_post' s input pm v consumed left right); - fold (peek_stick_post s input pm v consumed (left, right)); - (left, right) + fold (peek_stick_post s input pm v consumed (left `SlicePair` right)); + (left `SlicePair` right) }} } ``` @@ -462,13 +462,13 @@ fn peek_stick_gen ) { let split123 = slice_split_stick false input offset; - match split123 { Mktuple2 input1 input23 -> { + match split123 { SlicePair input1 input23 -> { unfold (slice_split_stick_post input pm v offset split123); unfold (slice_split_stick_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); stick_elim_partial_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); let split23 = peek_stick s input23 consumed; - match split23 { Mktuple2 input2 input3 -> { + match split23 { SlicePair input2 input3 -> { unfold (peek_stick_post s input23 pm v23 consumed split23); unfold (peek_stick_post' s input23 pm v23 consumed input2 input3); with v' . assert (pts_to_serialized s input2 #pm v'); diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 7cfae59d5..75dee6699 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1137,9 +1137,9 @@ let split_dtuple2_post (input: slice byte) (pm: perm) (v: Ghost.erased (dtuple2 t1 t2)) - (res: (slice byte & slice byte)) + (res: slice_pair byte) : Tot vprop -= let (left, right) = res in += let (SlicePair left right) = res in split_dtuple2_post' s1 s2 input pm v left right inline_for_extraction @@ -1158,7 +1158,7 @@ fn split_dtuple2 (#pm: perm) (#v: Ghost.erased (dtuple2 t1 t2)) requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v - returns res: (slice byte & slice byte) + returns res: slice_pair byte ensures split_dtuple2_post s1 s2 input pm v res { rewrite_with_stick @@ -1168,7 +1168,7 @@ fn split_dtuple2 let i = j1 input 0sz; let res = slice_append_split_stick false input i; match res { - Mktuple2 input1 input2 -> { + SlicePair input1 input2 -> { unfold (slice_append_split_stick_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); unfold (slice_append_split_stick_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); stick_trans (_ ** _) _ _; @@ -1177,8 +1177,8 @@ fn split_dtuple2 stick_prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); stick_trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; fold (split_dtuple2_post' s1 s2 input pm v input1 input2); - fold (split_dtuple2_post s1 s2 input pm v (input1, input2)); - (input1, input2) + fold (split_dtuple2_post s1 s2 input pm v (input1 `SlicePair` input2)); + (input1 `SlicePair` input2) } } } @@ -1214,9 +1214,9 @@ let split_nondep_then_post (input: slice byte) (pm: perm) (v: Ghost.erased (t1 & t2)) - (res: (slice byte & slice byte)) + (res: slice_pair byte) : Tot vprop -= let (left, right) = res in += let (SlicePair left right) = res in split_nondep_then_post' s1 s2 input pm v left right inline_for_extraction @@ -1234,7 +1234,7 @@ fn split_nondep_then (#pm: perm) (#v: Ghost.erased (t1 & t2)) requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v - returns res: (slice byte & slice byte) + returns res: slice_pair byte ensures split_nondep_then_post s1 s2 input pm v res { Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); @@ -1255,13 +1255,13 @@ fn split_nondep_then input; stick_trans (pts_to_serialized (tot_serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; let res = split_dtuple2 s1 j1 (fun _ -> s2) input; - match res { Mktuple2 input1 input2 -> { + match res { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 (fun _ -> s2) input pm (dtuple2_of_pair v) res); unfold (split_dtuple2_post' s1 (fun _ -> s2) input pm (dtuple2_of_pair v) input1 input2); stick_trans (_ ** _) _ _; fold (split_nondep_then_post' s1 s2 input pm v input1 input2); - fold (split_nondep_then_post s1 s2 input pm v (input1, input2)); - (input1, input2) + fold (split_nondep_then_post s1 s2 input pm v (input1 `SlicePair` input2)); + (input1 `SlicePair` input2) }} } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index f4a238bb5..aa56409cf 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -184,6 +184,26 @@ fn stick_elim_partial_r } ``` +noeq type slice_pair (t: Type) = | SlicePair: (fst: S.slice t) -> (snd: S.slice t) -> slice_pair t + +let split_post0 + (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) + (res: (slice_pair t)) +: Tot vprop += let SlicePair s1 s2 = res in + S.split_post' s p v i s1 s2 + +```pulse +fn slice_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) + requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) + returns res: slice_pair t + ensures split_post0 s p v i res +{ + admit () +} +``` + +noextract let slice_append_split_precond (#t: Type) (mutb: bool) (p: perm) (v1: Ghost.erased (Seq.seq t)) (i: SZ.t) : Tot prop @@ -201,29 +221,30 @@ let slice_append_split_post' let slice_append_split_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: (S.slice t & S.slice t)) + (res: slice_pair t) : Tot vprop -= let (s1, s2) = res in += let SlicePair s1 s2 = res in slice_append_split_post' s p v1 v2 i s1 s2 inline_for_extraction +noextract ```pulse fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to s #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: (S.slice t & S.slice t) + returns res: slice_pair t ensures slice_append_split_post s p v1 v2 i res { let vs = Ghost.hide (Seq.split (Seq.append v1 v2) (SZ.v i)); assert (pure (fst vs `Seq.equal` v1)); assert (pure (snd vs `Seq.equal` v2)); - let res = S.split mutb s i; + let res = slice_split mutb s i; match res { - Mktuple2 s1 s2 -> { - unfold (S.split_post s p (Seq.append v1 v2) i res); + SlicePair s1 s2 -> { + unfold (split_post0 s p (Seq.append v1 v2) i res); unfold (S.split_post' s p (Seq.append v1 v2) i s1 s2); fold (slice_append_split_post' s p v1 v2 i s1 s2); - fold (slice_append_split_post s p v1 v2 i (s1, s2)); - (s1, s2) + fold (slice_append_split_post s p v1 v2 i (SlicePair s1 s2)); + (SlicePair s1 s2) } } } @@ -241,21 +262,22 @@ let slice_append_split_stick_post' let slice_append_split_stick_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: (S.slice t & S.slice t)) + (res: slice_pair t) : Tot vprop -= let (s1, s2) = res in += let SlicePair s1 s2 = res in slice_append_split_stick_post' s p v1 v2 i s1 s2 inline_for_extraction +noextract ```pulse fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to input #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: (S.slice t & S.slice t) + returns res: slice_pair t ensures slice_append_split_stick_post input p v1 v2 i res { let res = slice_append_split mutb input i; match res { - Mktuple2 input1 input2 -> { + SlicePair input1 input2 -> { unfold (slice_append_split_post input p v1 v2 i res); unfold (slice_append_split_post' input p v1 v2 i input1 input2); ghost @@ -268,8 +290,8 @@ fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm }; intro_stick _ _ _ aux; fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); - fold (slice_append_split_stick_post input p v1 v2 i (input1, input2)); - (input1, input2) + fold (slice_append_split_stick_post input p v1 v2 i (SlicePair input1 input2)); + (SlicePair input1 input2) } } } @@ -291,22 +313,23 @@ let slice_split_stick_post' let slice_split_stick_post (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: (S.slice t & S.slice t)) + (res: slice_pair t) : Tot vprop -= let (s1, s2) = res in += let (SlicePair s1 s2) = res in slice_split_stick_post' s p v i s1 s2 inline_for_extraction +noextract ```pulse fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) - returns res: (S.slice t & S.slice t) + returns res: slice_pair t ensures slice_split_stick_post s p v i res { Seq.lemma_split v (SZ.v i); - let res = S.split mutb s i; - match res { Mktuple2 s1 s2 -> { - unfold (S.split_post s p v i res); + let res = slice_split mutb s i; + match res { SlicePair s1 s2 -> { + unfold (split_post0 s p v i res); unfold (S.split_post' s p v i s1 s2); with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); ghost @@ -319,8 +342,8 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos }; intro_stick _ _ _ aux; fold (slice_split_stick_post' s p v i s1 s2); - fold (slice_split_stick_post s p v i (s1, s2)); - (s1, s2) + fold (slice_split_stick_post s p v i (SlicePair s1 s2)); + (SlicePair s1 s2) }} } ``` From 92f63439d50f9aabbbd5059101e8742c84bd1c01 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 18:36:43 -0700 Subject: [PATCH 25/89] Karamel chokes on local ghost functions, hoist them instead (mainly for intro_stick) --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 34 ++++++++++++----- src/lowparse/pulse/LowParse.Pulse.Util.fst | 43 +++++++++++++--------- 2 files changed, 49 insertions(+), 28 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index c234c1a78..47c5513cb 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -411,6 +411,29 @@ let peek_stick_post = let (SlicePair left right) = res in peek_stick_post' s input pm v consumed left right +```pulse +ghost +fn peek_stick_aux + (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (input: slice byte) + (pm: perm) + (consumed: SZ.t) + (v: bytes) + (left right: slice byte) + (v1: t) + (v2: bytes) + (hyp: squash ( + bare_serialize s v1 `Seq.append` v2 == v + )) + (_: unit) + requires (is_split input pm consumed left right ** (pts_to_serialized s left #pm v1 ** pts_to right #pm v2)) + ensures pts_to input #pm v +{ + unfold (pts_to_serialized s left #pm v1); + join left right input +} +``` + inline_for_extraction ```pulse fn peek_stick @@ -428,16 +451,7 @@ fn peek_stick unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); - ghost - fn aux - (_foo: unit) - requires (is_split input pm consumed left right ** (pts_to_serialized s left #pm v1 ** pts_to right #pm v2)) - ensures pts_to input #pm v - { - unfold (pts_to_serialized s left #pm v1); - join left right input - }; - intro_stick (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) aux; + intro_stick (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_stick_aux s input pm consumed v left right v1 v2 ()); fold (peek_stick_post' s input pm v consumed left right); fold (peek_stick_post s input pm v consumed (left `SlicePair` right)); (left `SlicePair` right) diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index aa56409cf..4441a0735 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -267,6 +267,17 @@ let slice_append_split_stick_post = let SlicePair s1 s2 = res in slice_append_split_stick_post' s p v1 v2 i s1 s2 +```pulse +ghost +fn slice_append_split_stick_aux + (#t: Type) (input: S.slice t) (p: perm) (v1 v2: (Seq.seq t)) (i: SZ.t) (input1 input2: S.slice t) (_: unit) + requires S.is_split input p i input1 input2 ** (S.pts_to input1 #p v1 ** S.pts_to input2 #p v2) + ensures S.pts_to input #p (v1 `Seq.append` v2) +{ + S.join input1 input2 input +} +``` + inline_for_extraction noextract ```pulse @@ -280,15 +291,7 @@ fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm SlicePair input1 input2 -> { unfold (slice_append_split_post input p v1 v2 i res); unfold (slice_append_split_post' input p v1 v2 i input1 input2); - ghost - fn aux - (_foo: unit) - requires S.is_split input p i input1 input2 ** (S.pts_to input1 #p v1 ** S.pts_to input2 #p v2) - ensures S.pts_to input #p (v1 `Seq.append` v2) - { - S.join input1 input2 input; - }; - intro_stick _ _ _ aux; + intro_stick _ _ _ (slice_append_split_stick_aux input p v1 v2 i input1 input2); fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); fold (slice_append_split_stick_post input p v1 v2 i (SlicePair input1 input2)); (SlicePair input1 input2) @@ -318,6 +321,18 @@ let slice_split_stick_post = let (SlicePair s1 s2) = res in slice_split_stick_post' s p v i s1 s2 +```pulse +ghost +fn slice_split_stick_aux + (#t: Type) (s: S.slice t) (p: perm) (v: Seq.seq t) (i: SZ.t) + (s1 s2: S.slice t) (v1 v2: Seq.seq t) (hyp: squash (v == Seq.append v1 v2)) (_: unit) + requires (S.is_split s p i s1 s2 ** (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2)) + ensures (S.pts_to s #p v) + { + S.join s1 s2 s + } +``` + inline_for_extraction noextract ```pulse @@ -332,15 +347,7 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos unfold (split_post0 s p v i res); unfold (S.split_post' s p v i s1 s2); with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); - ghost - fn aux - (_: unit) - requires (S.is_split s p i s1 s2 ** (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2)) - ensures (S.pts_to s #p v) - { - S.join s1 s2 s - }; - intro_stick _ _ _ aux; + intro_stick _ _ _ (slice_split_stick_aux s p v i s1 s2 v1 v2 ()); fold (slice_split_stick_post' s p v i s1 s2); fold (slice_split_stick_post s p v i (SlicePair s1 s2)); (SlicePair s1 s2) From 5e76a3f4c76145c18ae9c7eb86152b9a6423a36a Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 20:30:00 -0700 Subject: [PATCH 26/89] slice_pair moved to Pulse --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 4 +- src/lowparse/pulse/LowParse.Pulse.Util.fst | 63 ++++++++-------------- 2 files changed, 24 insertions(+), 43 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 47c5513cb..349167d40 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -368,13 +368,13 @@ fn peek returns res: slice_pair byte ensures peek_post s input pm v consumed res { - let s1s2 = slice_split false input #pm #v consumed; + let s1s2 = split false input #pm #v consumed; match s1s2 { SlicePair s1 s2 -> { Seq.lemma_split v (SZ.v consumed); let v1 = Ghost.hide (fst (Some?.v (parse p v))); parse_injective #k p (bare_serialize s v1) v; - unfold (split_post0 input pm v consumed (SlicePair s1 s2)); + unfold (split_post input pm v consumed (SlicePair s1 s2)); unfold (split_post' input pm v consumed s1 s2); with v1' . assert (pts_to s1 #pm v1'); rewrite (pts_to s1 #pm v1') as (pts_to_serialized s s1 #pm v1); diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 4441a0735..e84ba5661 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -184,25 +184,6 @@ fn stick_elim_partial_r } ``` -noeq type slice_pair (t: Type) = | SlicePair: (fst: S.slice t) -> (snd: S.slice t) -> slice_pair t - -let split_post0 - (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: (slice_pair t)) -: Tot vprop -= let SlicePair s1 s2 = res in - S.split_post' s p v i s1 s2 - -```pulse -fn slice_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) - requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) - returns res: slice_pair t - ensures split_post0 s p v i res -{ - admit () -} -``` - noextract let slice_append_split_precond (#t: Type) (mutb: bool) (p: perm) (v1: Ghost.erased (Seq.seq t)) (i: SZ.t) @@ -221,9 +202,9 @@ let slice_append_split_post' let slice_append_split_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: slice_pair t) + (res: S.slice_pair t) : Tot vprop -= let SlicePair s1 s2 = res in += let S.SlicePair s1 s2 = res in slice_append_split_post' s p v1 v2 i s1 s2 inline_for_extraction @@ -231,20 +212,20 @@ noextract ```pulse fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to s #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: slice_pair t + returns res: S.slice_pair t ensures slice_append_split_post s p v1 v2 i res { let vs = Ghost.hide (Seq.split (Seq.append v1 v2) (SZ.v i)); assert (pure (fst vs `Seq.equal` v1)); assert (pure (snd vs `Seq.equal` v2)); - let res = slice_split mutb s i; + let res = S.split mutb s i; match res { - SlicePair s1 s2 -> { - unfold (split_post0 s p (Seq.append v1 v2) i res); + S.SlicePair s1 s2 -> { + unfold (S.split_post s p (Seq.append v1 v2) i res); unfold (S.split_post' s p (Seq.append v1 v2) i s1 s2); fold (slice_append_split_post' s p v1 v2 i s1 s2); - fold (slice_append_split_post s p v1 v2 i (SlicePair s1 s2)); - (SlicePair s1 s2) + fold (slice_append_split_post s p v1 v2 i (S.SlicePair s1 s2)); + (S.SlicePair s1 s2) } } } @@ -262,9 +243,9 @@ let slice_append_split_stick_post' let slice_append_split_stick_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: slice_pair t) + (res: S.slice_pair t) : Tot vprop -= let SlicePair s1 s2 = res in += let S.SlicePair s1 s2 = res in slice_append_split_stick_post' s p v1 v2 i s1 s2 ```pulse @@ -283,18 +264,18 @@ noextract ```pulse fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to input #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: slice_pair t + returns res: S.slice_pair t ensures slice_append_split_stick_post input p v1 v2 i res { let res = slice_append_split mutb input i; match res { - SlicePair input1 input2 -> { + S.SlicePair input1 input2 -> { unfold (slice_append_split_post input p v1 v2 i res); unfold (slice_append_split_post' input p v1 v2 i input1 input2); intro_stick _ _ _ (slice_append_split_stick_aux input p v1 v2 i input1 input2); fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); - fold (slice_append_split_stick_post input p v1 v2 i (SlicePair input1 input2)); - (SlicePair input1 input2) + fold (slice_append_split_stick_post input p v1 v2 i (S.SlicePair input1 input2)); + (S.SlicePair input1 input2) } } } @@ -316,9 +297,9 @@ let slice_split_stick_post' let slice_split_stick_post (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: slice_pair t) + (res: S.slice_pair t) : Tot vprop -= let (SlicePair s1 s2) = res in += let (S.SlicePair s1 s2) = res in slice_split_stick_post' s p v i s1 s2 ```pulse @@ -338,19 +319,19 @@ noextract ```pulse fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) - returns res: slice_pair t + returns res: S.slice_pair t ensures slice_split_stick_post s p v i res { Seq.lemma_split v (SZ.v i); - let res = slice_split mutb s i; - match res { SlicePair s1 s2 -> { - unfold (split_post0 s p v i res); + let res = S.split mutb s i; + match res { S.SlicePair s1 s2 -> { + unfold (S.split_post s p v i res); unfold (S.split_post' s p v i s1 s2); with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); intro_stick _ _ _ (slice_split_stick_aux s p v i s1 s2 v1 v2 ()); fold (slice_split_stick_post' s p v i s1 s2); - fold (slice_split_stick_post s p v i (SlicePair s1 s2)); - (SlicePair s1 s2) + fold (slice_split_stick_post s p v i (S.SlicePair s1 s2)); + (S.SlicePair s1 s2) }} } ``` From 39a46385a7a822cca98f3bb180299f65b907b8b3 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 11 Jul 2024 20:38:36 -0700 Subject: [PATCH 27/89] read then pure cont --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 43 ++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 349167d40..9d8ce7abc 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -568,6 +568,49 @@ fn read } ``` +inline_for_extraction +```pulse +fn read_then_cont + (#t: Type0) + (#t': Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: reader s) + (f: t -> t') + (input: slice byte) + (pm: perm) + (v: Ghost.erased t) + (v': t { Ghost.reveal v == v' }) + requires (pts_to_serialized s input #pm v ** emp) + returns v_ : t' + ensures (pts_to_serialized s input #pm v' ** pure (f (Ghost.reveal v) == v_)) +{ + f v' +} +``` + +inline_for_extraction +```pulse +fn read_then + (#t: Type0) + (#t': Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: reader s) + (f: t -> t') + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + requires (pts_to_serialized s input #pm v) + returns v' : t' + ensures (pts_to_serialized s input #pm v ** pure (f (Ghost.reveal v) == v')) +{ + call_reader r input #pm #v emp t' (fun v' -> pts_to_serialized s input #pm v ** pure (f (Ghost.reveal v) == v')) (read_then_cont r f input pm v) +} +``` + inline_for_extraction let leaf_reader (#t: Type0) From f9a36f30266c5e911dd81577075958609e2a3800 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 12 Jul 2024 10:56:25 -0700 Subject: [PATCH 28/89] read with pure continuation --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 54 +++++++++++ src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 33 +++++++ .../pulse/LowParse.Pulse.Combinators.fst | 91 +++++++++++++++++-- 3 files changed, 172 insertions(+), 6 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 9d8ce7abc..fddea88d0 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -676,6 +676,60 @@ let reader_of_leaf_reader : reader s = reader_of_leaf_reader' r +inline_for_extraction +let pure_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) +: Tot Type += (input: slice byte) -> + (#pm: perm) -> + (#v: Ghost.erased t) -> + (t': Type0) -> + (f: ((x: t { x == Ghost.reveal v }) -> Tot t')) -> + stt t' (pts_to_serialized s input #pm v) (fun x' -> pts_to_serialized s input #pm v ** pure (x' == f v)) + +inline_for_extraction +```pulse +fn pure_read + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: pure_reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + requires (pts_to_serialized s input #pm v) + returns v' : t + ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) +{ + r input #pm #v t id +} +``` + +inline_for_extraction +```pulse +fn pure_reader_of_leaf_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: leaf_reader s) +: pure_reader #t #k #p s += + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + (t': Type0) + (f: _) +{ + let x = r input #pm #v; + f x +} +``` + inline_for_extraction let validate_and_read (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (input: slice byte) -> diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index c57fd5104..7d88dc4ee 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -122,3 +122,36 @@ let read_bitsum' stt_cps_ifthenelse (read_bitsum'_cont b) ) + +inline_for_extraction +let pure_read_synth_cont_ifthenelse + (#t: Type0) + (x: Ghost.erased t) +: Tot (if_combinator_weak (pure_read_synth_cont_t #t x)) += fun cond iftrue iffalse t' g' -> if cond then iftrue () t' g' else iffalse () t' g' + +inline_for_extraction +let pure_read_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (#b: bitsum' cl tot) + (d: destr_bitsum'_t b) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: pure_reader s) +: Tot (pure_reader (tot_serialize_bitsum' b s)) += pure_read_synth + (pure_read_filter + r + (filter_bitsum' b) + ) + (synth_bitsum'_injective b; synth_bitsum' b) + (synth_bitsum'_recip_inverse b; synth_bitsum'_recip b) + ( + d + (pure_read_synth_cont_t #(bitsum'_type b)) + (pure_read_synth_cont_ifthenelse #(bitsum'_type b)) + (pure_read_synth_cont_init) + ) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 75dee6699..fd56e4ab8 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -359,6 +359,58 @@ let read_synth' : Tot (reader (tot_serialize_synth p1 f2 s1 f1 ())) = read_synth r f2 f1 (fun x -> intro_stt_cps (f2 x)) +inline_for_extraction +let pure_read_synth_cont_t + (#t: Type0) + (x: t) += (t': Type0) -> (g': ((y: t { y == x }) -> t')) -> (y': t' { y' == g' x }) + +inline_for_extraction +let pure_read_synth_cont + (#t1 #t2: Type0) + (f2: (t1 -> Tot t2)) + (f2': ((x: t1) -> pure_read_synth_cont_t (f2 x))) + (x1: Ghost.erased t1) + (t': Type0) + (g: ((x2: t2 { x2 == f2 x1 }) -> Tot t')) + (x1': t1 { x1' == Ghost.reveal x1 }) +: Tot t' += f2' x1' t' g + +inline_for_extraction +```pulse +fn pure_read_synth + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (f2': ((x: t1) -> pure_read_synth_cont_t (f2 x))) +: pure_reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) += (input: slice byte) + (#pm: _) + (#v: _) + (t': Type0) + (g: _) +{ + pts_to_serialized_synth_l2r_stick s1 f2 f1 input; + let res = r input #pm #(f1 v) t' (pure_read_synth_cont f2 f2' (f1 v) t' g); + elim_stick _ _; + res +} +``` + +inline_for_extraction +let pure_read_synth_cont_init + (#t: Type0) + (x: t) +: Tot (pure_read_synth_cont_t #t x) += fun t' g' -> g' x + +inline_for_extraction +let pure_read_synth' + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) +: pure_reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) += pure_read_synth r f2 f1 (fun x -> pure_read_synth_cont_init (f2 x)) + inline_for_extraction ```pulse fn validate_and_read_synth_cont_cont @@ -570,14 +622,12 @@ fn pts_to_serialized_filter_intro (f: (t -> bool)) (input: slice byte) (#pm: perm) - (#v: Ghost.erased t) + (#v: t) requires (pts_to_serialized s input #pm v ** pure (f v == true)) - ensures exists* (v': parse_filter_refine f) . pts_to_serialized (tot_serialize_filter s f) input #pm v' ** pure ((v' <: t) == Ghost.reveal v) + ensures exists* (v': parse_filter_refine f) . pts_to_serialized (tot_serialize_filter s f) input #pm v' ** pure ((v' <: t) == v) { unfold (pts_to_serialized s input #pm v); - let sq: squash (f v == true) = (); - let v' : Ghost.erased (parse_filter_refine f) = Ghost.hide (parse_filter_refine_intro #t f (Ghost.reveal v) sq); - fold (pts_to_serialized (tot_serialize_filter s f) input #pm v'); + fold (pts_to_serialized (tot_serialize_filter s f) input #pm v); } ``` @@ -591,7 +641,7 @@ fn pts_to_serialized_filter_elim (f: (t -> bool)) (input: slice byte) (#pm: perm) - (#v: Ghost.erased (parse_filter_refine f)) + (#v: parse_filter_refine f) requires (pts_to_serialized (tot_serialize_filter s f) input #pm v) ensures pts_to_serialized s input #pm v { @@ -618,6 +668,35 @@ fn read_filter } ``` +inline_for_extraction +let pure_read_filter_cont + (#t: Type0) + (f: t -> bool) + (v: Ghost.erased (parse_filter_refine f)) + (t': Type0) + (g: (x: parse_filter_refine f { x == Ghost.reveal v }) -> t') + (x: t { x == Ghost.reveal #t (Ghost.hide #t (Ghost.reveal v)) }) +: Tot t' += g x + +inline_for_extraction +```pulse +fn pure_read_filter + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) (f: (t1 -> bool)) +: pure_reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) += (input: slice byte) + (#pm: _) + (#v: _) + (t': Type0) + (g: _) +{ + pts_to_serialized_filter_elim s1 f input; + let res = r input #pm #(Ghost.hide (Ghost.reveal v)) t' (pure_read_filter_cont f v t' g); + pts_to_serialized_filter_intro s1 f input; + res +} +``` + let pair_of_dtuple2 (#t1 #t2: Type) (x: dtuple2 t1 (fun _ -> t2)) From b5f83dfa18cf5c99c13dee310d24d731698c05aa Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 12 Jul 2024 14:49:03 -0700 Subject: [PATCH 29/89] pure_read_ext, empty, dtuple, ifthenelse --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 52 +++++++++++++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 39 +++++++++++++- 2 files changed, 90 insertions(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index fddea88d0..e3086f5fd 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -709,6 +709,58 @@ fn pure_read } ``` +inline_for_extraction +```pulse +fn ifthenelse_pure_reader + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: bool) + (iftrue: squash (cond == true) -> pure_reader s) + (iffalse: squash (cond == false) -> pure_reader s) +:pure_reader #t #k #p s += + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + (t': Type0) + (f: _) +{ + if cond { + iftrue () input #pm #v t' f + } else { + iffalse () input #pm #v t' f + } +} +``` + +inline_for_extraction +```pulse +fn pure_reader_ext + (#t: Type0) + (#k1: parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (r1: pure_reader s1) + (#k2: parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { forall x . parse p1 x == parse p2 x }) +:pure_reader #t #k2 #p2 s2 += + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + (t': Type0) + (f: _) +{ + pts_to_serialized_ext_stick s2 s1 input; + let res = r1 input #pm #v t' f; + elim_stick _ _; + res +} +``` + inline_for_extraction ```pulse fn pure_reader_of_leaf_reader diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index fd56e4ab8..e7edb48c3 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -67,7 +67,10 @@ inline_for_extraction let validate_empty : validator tot_parse_empty = validate_ret () inline_for_extraction -let read_empty : reader tot_serialize_empty = read_ret () (fun _ -> ()) +let leaf_read_empty : leaf_reader tot_serialize_empty = leaf_read_ret () (fun _ -> ()) + +inline_for_extraction +let read_empty : reader tot_serialize_empty = reader_of_leaf_reader leaf_read_empty inline_for_extraction let validate_and_read_empty : validate_and_read tot_parse_empty = @@ -1344,3 +1347,37 @@ fn split_nondep_then }} } ``` + +inline_for_extraction +```pulse +fn pure_read_dtuple2 + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (j1: jumper p1 { k1.parser_kind_subkind == Some ParserStrong }) + (#s1: serializer p1) + (r1: pure_reader s1) + (#k2: parser_kind) + (#p2: (x: t1) -> parser k2 (t2 x)) + (#s2: (x: t1) -> serializer (p2 x)) + (r2: (x: t1) -> pure_reader (s2 x)) +: pure_reader #(dtuple2 t1 t2) #(and_then_kind k1 k2) #(tot_parse_dtuple2 p1 p2) (tot_serialize_dtuple2 s1 s2) += + (input: slice byte) + (#pm: perm) + (#v: _) + (t': Type0) + (f: _) +{ + let split12 = split_dtuple2 s1 j1 s2 input; + match split12 { SlicePair input1 input2 -> { + unfold (split_dtuple2_post s1 s2 input pm v split12); + unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); + let x1 = pure_read r1 input1; + let x2 = pure_read (r2 x1) input2; + elim_stick _ _; + f (Mkdtuple2 x1 x2) + }} +} +``` From ff928b07d3a7923eed645ab0b9bf90c75e8b2d9e Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 12 Jul 2024 16:18:24 -0700 Subject: [PATCH 30/89] properly normalize integer readers --- src/lowparse/pulse/LowParse.Pulse.Int.fst | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index f31f2bc14..879352749 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -53,7 +53,8 @@ let jump_u16 : jumper tot_parse_u16 = inline_for_extraction noextract -let be_to_n_2 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint16 2) +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let be_to_n_2 = (E.mk_be_to_n EI.uint16 2) inline_for_extraction ```pulse @@ -87,7 +88,8 @@ let jump_u32 : jumper tot_parse_u32 = inline_for_extraction noextract -let be_to_n_4 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint32 4) +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let be_to_n_4 = (E.mk_be_to_n EI.uint32 4) inline_for_extraction ```pulse @@ -121,7 +123,8 @@ let jump_u64 : jumper tot_parse_u64 = inline_for_extraction noextract -let be_to_n_8 = norm [delta_attr [`%E.must_reduce]; iota; zeta; primops] (E.mk_be_to_n EI.uint64 8) +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let be_to_n_8 = (E.mk_be_to_n EI.uint64 8) inline_for_extraction ```pulse From 5d088a0f9ca095175c3030be23ad8393ec02b534 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 12 Jul 2024 21:49:40 -0700 Subject: [PATCH 31/89] replace CPS reader with pure_reader --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 182 ++---------------- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 62 +----- .../pulse/LowParse.Pulse.Combinators.fst | 125 +++--------- src/lowparse/pulse/LowParse.Pulse.Int.fst | 8 + 4 files changed, 63 insertions(+), 314 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index e3086f5fd..fce74d3c1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -494,123 +494,6 @@ fn peek_stick_gen } ``` -inline_for_extraction -let reader - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (s: serializer p) -: Tot Type -= (input: slice byte) -> - (#pm: perm) -> - (#v: Ghost.erased t) -> - (pre: vprop) -> - (t': Type0) -> - (post: (t' -> vprop)) -> - (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x)) -> - stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x) - -inline_for_extraction -let call_reader - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: reader s) - (input: slice byte) - (#pm: perm) - (#v: Ghost.erased t) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun t' -> post t')) -: stt t' (pts_to_serialized s input #pm v ** pre) (fun t' -> post t') -= r input pre t' (fun x -> post x) (fun x -> cont x) - -#push-options "--print_implicits --query_stats" - -inline_for_extraction -```pulse -fn read_cont - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: reader s) - (input: slice byte) - (pm: perm) - (v: Ghost.erased t) - (v': t { Ghost.reveal v == v' }) - requires (pts_to_serialized s input #pm v ** emp) - returns v' : t - ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) -{ - v' -} -``` - -inline_for_extraction -```pulse -fn read - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: reader s) - (input: slice byte) - (#pm: perm) - (#v: Ghost.erased t) - requires (pts_to_serialized s input #pm v) - returns v' : t - ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) -{ - call_reader r input #pm #v emp t (fun v' -> pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) (read_cont r input pm v) -} -``` - -inline_for_extraction -```pulse -fn read_then_cont - (#t: Type0) - (#t': Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: reader s) - (f: t -> t') - (input: slice byte) - (pm: perm) - (v: Ghost.erased t) - (v': t { Ghost.reveal v == v' }) - requires (pts_to_serialized s input #pm v ** emp) - returns v_ : t' - ensures (pts_to_serialized s input #pm v' ** pure (f (Ghost.reveal v) == v_)) -{ - f v' -} -``` - -inline_for_extraction -```pulse -fn read_then - (#t: Type0) - (#t': Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: reader s) - (f: t -> t') - (input: slice byte) - (#pm: perm) - (#v: Ghost.erased t) - requires (pts_to_serialized s input #pm v) - returns v' : t' - ensures (pts_to_serialized s input #pm v ** pure (f (Ghost.reveal v) == v')) -{ - call_reader r input #pm #v emp t' (fun v' -> pts_to_serialized s input #pm v ** pure (f (Ghost.reveal v) == v')) (read_then_cont r f input pm v) -} -``` - inline_for_extraction let leaf_reader (#t: Type0) @@ -643,41 +526,7 @@ let leaf_read = r input #pm #v inline_for_extraction -```pulse -fn reader_of_leaf_reader' - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: leaf_reader s) - (input: slice byte) - (#pm: perm) - (#v: Ghost.erased t) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (cont: (x: t { x == Ghost.reveal v }) -> stt t' (pts_to_serialized s input #pm v ** pre) (fun x -> post x)) - requires (pts_to_serialized s input #pm v ** pre) - returns res: t' - ensures post res -{ - let res = leaf_read r input #pm #v; - cont res -} -``` - -inline_for_extraction -let reader_of_leaf_reader - (#t: Type0) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: leaf_reader s) -: reader s -= reader_of_leaf_reader' r - -inline_for_extraction -let pure_reader +let reader (#t: Type0) (#k: parser_kind) (#p: parser k t) @@ -692,18 +541,17 @@ let pure_reader inline_for_extraction ```pulse -fn pure_read +fn leaf_reader_of_reader (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) - (r: pure_reader s) + (r: reader s) +: leaf_reader #t #k #p s += (input: slice byte) (#pm: perm) (#v: Ghost.erased t) - requires (pts_to_serialized s input #pm v) - returns v' : t - ensures (pts_to_serialized s input #pm v' ** pure (Ghost.reveal v == v')) { r input #pm #v t id } @@ -711,15 +559,15 @@ fn pure_read inline_for_extraction ```pulse -fn ifthenelse_pure_reader +fn ifthenelse_reader (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) (cond: bool) - (iftrue: squash (cond == true) -> pure_reader s) - (iffalse: squash (cond == false) -> pure_reader s) -:pure_reader #t #k #p s + (iftrue: squash (cond == true) -> reader s) + (iffalse: squash (cond == false) -> reader s) +:reader #t #k #p s = (input: slice byte) (#pm: perm) @@ -737,16 +585,16 @@ fn ifthenelse_pure_reader inline_for_extraction ```pulse -fn pure_reader_ext +fn reader_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (#s1: serializer p1) - (r1: pure_reader s1) + (r1: reader s1) (#k2: parser_kind) (#p2: parser k2 t) (s2: serializer p2 { forall x . parse p1 x == parse p2 x }) -:pure_reader #t #k2 #p2 s2 +:reader #t #k2 #p2 s2 = (input: slice byte) (#pm: perm) @@ -763,13 +611,13 @@ fn pure_reader_ext inline_for_extraction ```pulse -fn pure_reader_of_leaf_reader +fn reader_of_leaf_reader (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) -: pure_reader #t #k #p s +: reader #t #k #p s = (input: slice byte) (#pm: perm) @@ -818,6 +666,7 @@ inline_for_extraction let validate_and_read_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1 { k2 `is_weaker_than` k1 }) : validate_and_read (tot_weaken k2 p1) = validate_and_read_ext v1 (tot_weaken k2 p1) +(* inline_for_extraction ```pulse fn validate_and_read_intro_cont_read @@ -885,6 +734,7 @@ fn validate_and_read_intro w input offset _ _ _ (validate_and_read_intro_cont_validate r input offset #pm #v pre t' post ksucc) kfail } ``` +*) inline_for_extraction ```pulse diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index 7d88dc4ee..6ed815e2e 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -78,24 +78,11 @@ let jump_bitsum' (synth_bitsum'_injective b; synth_bitsum' b) inline_for_extraction -```pulse -fn read_bitsum'_cont - (#t: eqtype) - (#tot: pos) - (#cl: uint_t tot t) - (b: bitsum' cl tot) - (k_: bitsum'_type b) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (phi: (k': bitsum'_type b { k' == k_ }) -> stt t' pre (fun x -> post x)) - requires pre - returns x: t' - ensures post x -{ - phi k_ -} -``` +let read_synth_cont_ifthenelse + (#t: Type0) + (x: Ghost.erased t) +: Tot (if_combinator_weak (read_synth_cont_t #t x)) += fun cond iftrue iffalse t' g' -> if cond then iftrue () t' g' else iffalse () t' g' inline_for_extraction let read_bitsum' @@ -118,40 +105,7 @@ let read_bitsum' (synth_bitsum'_recip_inverse b; synth_bitsum'_recip b) ( d - stt_cps - stt_cps_ifthenelse - (read_bitsum'_cont b) - ) - -inline_for_extraction -let pure_read_synth_cont_ifthenelse - (#t: Type0) - (x: Ghost.erased t) -: Tot (if_combinator_weak (pure_read_synth_cont_t #t x)) -= fun cond iftrue iffalse t' g' -> if cond then iftrue () t' g' else iffalse () t' g' - -inline_for_extraction -let pure_read_bitsum' - (#t: eqtype) - (#tot: pos) - (#cl: uint_t tot t) - (#b: bitsum' cl tot) - (d: destr_bitsum'_t b) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (r: pure_reader s) -: Tot (pure_reader (tot_serialize_bitsum' b s)) -= pure_read_synth - (pure_read_filter - r - (filter_bitsum' b) - ) - (synth_bitsum'_injective b; synth_bitsum' b) - (synth_bitsum'_recip_inverse b; synth_bitsum'_recip b) - ( - d - (pure_read_synth_cont_t #(bitsum'_type b)) - (pure_read_synth_cont_ifthenelse #(bitsum'_type b)) - (pure_read_synth_cont_init) + (read_synth_cont_t #(bitsum'_type b)) + (read_synth_cont_ifthenelse #(bitsum'_type b)) + (read_synth_cont_init) ) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index e7edb48c3..9b8b1985e 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -52,6 +52,7 @@ let read_ret : Tot (reader (tot_serialize_ret x v_unique)) = reader_of_leaf_reader (leaf_read_ret x v_unique) +(* inline_for_extraction let validate_and_read_ret (#t: Type0) @@ -59,6 +60,7 @@ let validate_and_read_ret (v_unique: (v' : t) -> Lemma (x == v')) : Tot (validate_and_read (tot_parse_ret x)) = validate_and_read_intro (validate_ret x) (read_ret x v_unique) +*) inline_for_extraction let jump_ret (#t: Type) (x: t) : jumper (tot_parse_ret x) = jump_constant_size (tot_parse_ret x) 0sz @@ -72,9 +74,11 @@ let leaf_read_empty : leaf_reader tot_serialize_empty = leaf_read_ret () (fun _ inline_for_extraction let read_empty : reader tot_serialize_empty = reader_of_leaf_reader leaf_read_empty +(* inline_for_extraction let validate_and_read_empty : validate_and_read tot_parse_empty = validate_and_read_intro validate_empty read_empty +*) inline_for_extraction let jump_empty : jumper tot_parse_empty = jump_ret () @@ -312,67 +316,16 @@ let elim_stt_cps = cps pre t' post phi inline_for_extraction -```pulse -fn read_synth_cont - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (s1: serializer p1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) - (f2': (x: t1) -> stt_cps (f2 x)) - (input: slice byte) - (pm: perm) - (v: Ghost.erased t2) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (cont: (x: t2 { x == Ghost.reveal v }) -> stt t' (pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v ** pre) (fun x -> post x)) - (x: t1 { x == Ghost.reveal (f1 v) }) - requires pts_to_serialized s1 input #pm (f1 v) ** (pre ** (pts_to_serialized s1 input #pm (f1 v) @==> pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v)) - returns x': t' - ensures post x' -{ - elim_stick _ _; - f2' x (pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v ** pre) t' post cont -} -``` - -inline_for_extraction -```pulse -fn read_synth - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) - (f2': (x: t1) -> stt_cps (f2 x)) -: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) -= (input: slice byte) - (#pm: _) - (#v: _) - (pre: _) - (t': Type0) - (post: _) - (cont: _) -{ - pts_to_serialized_synth_l2r_stick s1 f2 f1 input; - r input #pm #_ (pre ** (pts_to_serialized s1 input #pm (f1 v) @==> pts_to_serialized (tot_serialize_synth p1 f2 s1 f1 ()) input #pm v)) t' post - (read_synth_cont s1 f2 f1 f2' input pm v pre t' post cont) -} -``` - -inline_for_extraction -let read_synth' - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) -: Tot (reader (tot_serialize_synth p1 f2 s1 f1 ())) -= read_synth r f2 f1 (fun x -> intro_stt_cps (f2 x)) - -inline_for_extraction -let pure_read_synth_cont_t +let read_synth_cont_t (#t: Type0) (x: t) = (t': Type0) -> (g': ((y: t { y == x }) -> t')) -> (y': t' { y' == g' x }) inline_for_extraction -let pure_read_synth_cont +let read_synth_cont (#t1 #t2: Type0) (f2: (t1 -> Tot t2)) - (f2': ((x: t1) -> pure_read_synth_cont_t (f2 x))) + (f2': ((x: t1) -> read_synth_cont_t (f2 x))) (x1: Ghost.erased t1) (t': Type0) (g: ((x2: t2 { x2 == f2 x1 }) -> Tot t')) @@ -382,11 +335,11 @@ let pure_read_synth_cont inline_for_extraction ```pulse -fn pure_read_synth - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) +fn read_synth + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) - (f2': ((x: t1) -> pure_read_synth_cont_t (f2 x))) -: pure_reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) + (f2': ((x: t1) -> read_synth_cont_t (f2 x))) +: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) = (input: slice byte) (#pm: _) (#v: _) @@ -394,25 +347,25 @@ fn pure_read_synth (g: _) { pts_to_serialized_synth_l2r_stick s1 f2 f1 input; - let res = r input #pm #(f1 v) t' (pure_read_synth_cont f2 f2' (f1 v) t' g); + let res = r input #pm #(f1 v) t' (read_synth_cont f2 f2' (f1 v) t' g); elim_stick _ _; res } ``` inline_for_extraction -let pure_read_synth_cont_init +let read_synth_cont_init (#t: Type0) (x: t) -: Tot (pure_read_synth_cont_t #t x) +: Tot (read_synth_cont_t #t x) = fun t' g' -> g' x inline_for_extraction -let pure_read_synth' - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) +let read_synth' + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) -: pure_reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) -= pure_read_synth r f2 f1 (fun x -> pure_read_synth_cont_init (f2 x)) +: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) += read_synth r f2 f1 (fun x -> read_synth_cont_init (f2 x)) inline_for_extraction ```pulse @@ -654,25 +607,7 @@ fn pts_to_serialized_filter_elim ``` inline_for_extraction -```pulse -fn read_filter - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (f: (t1 -> bool)) -: reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) -= (input: slice byte) - (#pm: _) - (#v: _) - (pre: _) - (t': Type0) - (post: _) - (cont: _) -{ - pts_to_serialized_filter_elim s1 f input; - r input #pm #(Ghost.hide (Ghost.reveal v)) pre t' post cont -} -``` - -inline_for_extraction -let pure_read_filter_cont +let read_filter_cont (#t: Type0) (f: t -> bool) (v: Ghost.erased (parse_filter_refine f)) @@ -684,9 +619,9 @@ let pure_read_filter_cont inline_for_extraction ```pulse -fn pure_read_filter - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: pure_reader s1) (f: (t1 -> bool)) -: pure_reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) +fn read_filter + (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (f: (t1 -> bool)) +: reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) = (input: slice byte) (#pm: _) (#v: _) @@ -694,7 +629,7 @@ fn pure_read_filter (g: _) { pts_to_serialized_filter_elim s1 f input; - let res = r input #pm #(Ghost.hide (Ghost.reveal v)) t' (pure_read_filter_cont f v t' g); + let res = r input #pm #(Ghost.hide (Ghost.reveal v)) t' (read_filter_cont f v t' g); pts_to_serialized_filter_intro s1 f input; res } @@ -1124,6 +1059,7 @@ fn jump_nondep_then } ``` +(* inline_for_extraction ```pulse fn jump_dtuple2_payload_cont @@ -1186,6 +1122,7 @@ fn jump_dtuple2 (jump_dtuple2_payload_cont s1 p2 v2 input offset pm v off1 key input1) } ``` +*) let split_dtuple2_post' (#t1: Type0) @@ -1350,19 +1287,19 @@ fn split_nondep_then inline_for_extraction ```pulse -fn pure_read_dtuple2 +fn read_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) (#k1: parser_kind) (#p1: parser k1 t1) (j1: jumper p1 { k1.parser_kind_subkind == Some ParserStrong }) (#s1: serializer p1) - (r1: pure_reader s1) + (r1: reader s1) (#k2: parser_kind) (#p2: (x: t1) -> parser k2 (t2 x)) (#s2: (x: t1) -> serializer (p2 x)) - (r2: (x: t1) -> pure_reader (s2 x)) -: pure_reader #(dtuple2 t1 t2) #(and_then_kind k1 k2) #(tot_parse_dtuple2 p1 p2) (tot_serialize_dtuple2 s1 s2) + (r2: (x: t1) -> reader (s2 x)) +: reader #(dtuple2 t1 t2) #(and_then_kind k1 k2) #(tot_parse_dtuple2 p1 p2) (tot_serialize_dtuple2 s1 s2) = (input: slice byte) (#pm: perm) @@ -1374,8 +1311,8 @@ fn pure_read_dtuple2 match split12 { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v split12); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); - let x1 = pure_read r1 input1; - let x2 = pure_read (r2 x1) input2; + let x1 = leaf_reader_of_reader r1 input1; + let x2 = leaf_reader_of_reader (r2 x1) input2; elim_stick _ _; f (Mkdtuple2 x1 x2) }} diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 879352749..a9c806df4 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -39,9 +39,11 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 inline_for_extraction let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) +(* inline_for_extraction let validate_and_read_u8 : validate_and_read tot_parse_u8 = validate_and_read_intro validate_u8 read_u8 +*) inline_for_extraction let validate_u16 : validator tot_parse_u16 = @@ -74,9 +76,11 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_ inline_for_extraction let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) +(* inline_for_extraction let validate_and_read_u16 : validate_and_read tot_parse_u16 = validate_and_read_intro validate_u16 read_u16 +*) inline_for_extraction let validate_u32 : validator tot_parse_u32 = @@ -109,9 +113,11 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_ inline_for_extraction let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) +(* inline_for_extraction let validate_and_read_u32 : validate_and_read tot_parse_u32 = validate_and_read_intro validate_u32 read_u32 +*) inline_for_extraction let validate_u64 : validator tot_parse_u64 = @@ -144,6 +150,8 @@ fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_ inline_for_extraction let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) +(* inline_for_extraction let validate_and_read_u64 : validate_and_read tot_parse_u64 = validate_and_read_intro validate_u64 read_u64 +*) From cf0a5ba856ce8d57e2e97eb476368f83794be309 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 12 Jul 2024 22:46:34 -0700 Subject: [PATCH 32/89] restore validate_and_read_intro, jump_dtuple2, with leaf_reader this time --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 63 +++++++++---------- .../pulse/LowParse.Pulse.Combinators.fst | 53 ++-------------- src/lowparse/pulse/LowParse.Pulse.Int.fst | 16 ++--- 3 files changed, 41 insertions(+), 91 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index fce74d3c1..0030d123d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -525,6 +525,33 @@ let leaf_read ) = r input #pm #v +inline_for_extraction +```pulse +fn read_from_validator_success + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (r: leaf_reader s) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased bytes) + (offset: SZ.t) + (consumed: SZ.t) + requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (consumed))) + returns v' : t + ensures pts_to input #pm v ** pure ( + validator_success #k #t p offset v consumed /\ + parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v consumed) + ) +{ + let input' = peek_stick_gen s input offset consumed; + let res = r input'; + elim_stick _ _; + res +} +``` + inline_for_extraction let reader (#t: Type0) @@ -666,36 +693,10 @@ inline_for_extraction let validate_and_read_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1 { k2 `is_weaker_than` k1 }) : validate_and_read (tot_weaken k2 p1) = validate_and_read_ext v1 (tot_weaken k2 p1) -(* -inline_for_extraction -```pulse -fn validate_and_read_intro_cont_read - (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (off: SZ.t) - (#v': Ghost.erased t) - (input': slice byte { validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (Ghost.reveal v', SZ.v off) }) - (x: t { x == Ghost.reveal v' }) - requires pts_to_serialized s input' #pm v' ** (pre ** (pts_to_serialized s input' #pm v' @==> pts_to input #pm v)) - returns x': t' - ensures post x' -{ - elim_stick _ _; - ksucc off x -} -``` - inline_for_extraction ```pulse fn validate_and_read_intro_cont_validate - (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) (r: reader s) + (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -709,16 +710,15 @@ fn validate_and_read_intro_cont_validate returns x' : t' ensures post x' { - let input' = peek_stick_gen s input offset off; - with v' . assert (pts_to_serialized s input' #pm v'); - r input' #pm #v' _ _ _ (validate_and_read_intro_cont_read s input offset #pm #v pre t' post ksucc off #_ input') + let x = read_from_validator_success r input offset off; + ksucc off x } ``` inline_for_extraction ```pulse fn validate_and_read_intro - (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) (#s: serializer p) (r: reader s) + (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) (#s: serializer p) (r: leaf_reader s) : validate_and_read #t #k p = (input: slice byte) @@ -734,7 +734,6 @@ fn validate_and_read_intro w input offset _ _ _ (validate_and_read_intro_cont_validate r input offset #pm #v pre t' post ksucc) kfail } ``` -*) inline_for_extraction ```pulse diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 9b8b1985e..ae03e6a92 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -52,15 +52,13 @@ let read_ret : Tot (reader (tot_serialize_ret x v_unique)) = reader_of_leaf_reader (leaf_read_ret x v_unique) -(* inline_for_extraction let validate_and_read_ret (#t: Type0) (x: t) (v_unique: (v' : t) -> Lemma (x == v')) : Tot (validate_and_read (tot_parse_ret x)) -= validate_and_read_intro (validate_ret x) (read_ret x v_unique) -*) += validate_and_read_intro (validate_ret x) (leaf_read_ret x v_unique) inline_for_extraction let jump_ret (#t: Type) (x: t) : jumper (tot_parse_ret x) = jump_constant_size (tot_parse_ret x) 0sz @@ -74,11 +72,9 @@ let leaf_read_empty : leaf_reader tot_serialize_empty = leaf_read_ret () (fun _ inline_for_extraction let read_empty : reader tot_serialize_empty = reader_of_leaf_reader leaf_read_empty -(* inline_for_extraction let validate_and_read_empty : validate_and_read tot_parse_empty = - validate_and_read_intro validate_empty read_empty -*) + validate_and_read_intro validate_empty leaf_read_empty inline_for_extraction let jump_empty : jumper tot_parse_empty = jump_ret () @@ -1059,37 +1055,6 @@ fn jump_nondep_then } ``` -(* -inline_for_extraction -```pulse -fn jump_dtuple2_payload_cont - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (s1: serializer p1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (v2: ((x: t1) -> jumper (p2 x))) - (input: slice byte) - (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) - (off: SZ.t { jumper_pre (tot_parse_dtuple2 p1 p2) offset v /\ validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (key: Ghost.erased t1 { fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) - (input_key: slice byte) - (x: t1 { x == Ghost.reveal key }) - requires (pts_to_serialized s1 input_key #pm key ** (pts_to_serialized s1 input_key #pm key @==> pts_to input #pm v)) - returns res: SZ.t - ensures (pts_to input #pm v ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v res)) -{ - elim_stick (pts_to_serialized s1 input_key #pm key) _; - pts_to_len input; - let off2 = v2 x input (SZ.add offset off); - SZ.add off off2 -} -``` - inline_for_extraction ```pulse fn jump_dtuple2 @@ -1099,7 +1064,7 @@ fn jump_dtuple2 (#p1: parser k1 t1) (v1: jumper p1) (#s1: serializer p1) - (r1: reader s1) + (r1: leaf_reader s1) (#k2: parser_kind) (#p2: (x: t1) -> parser k2 (t2 x)) (v2: (x: t1) -> jumper (p2 x)) @@ -1112,17 +1077,11 @@ fn jump_dtuple2 { pts_to_len input; let off1 = v1 input offset; - let input1 = peek_stick_gen s1 input offset off1; - with key . assert (pts_to_serialized s1 input1 #pm key); - r1 - input1 - (pts_to_serialized s1 input1 #pm key @==> pts_to input #pm v) - SZ.t - (fun res -> pts_to input #pm v ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v res)) - (jump_dtuple2_payload_cont s1 p2 v2 input offset pm v off1 key input1) + let x = read_from_validator_success r1 input offset off1; + let off2 = v2 x input (SZ.add offset off1); + SZ.add off1 off2 } ``` -*) let split_dtuple2_post' (#t1: Type0) diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index a9c806df4..920b48aaa 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -39,11 +39,9 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 inline_for_extraction let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) -(* inline_for_extraction let validate_and_read_u8 : validate_and_read tot_parse_u8 = - validate_and_read_intro validate_u8 read_u8 -*) + validate_and_read_intro validate_u8 (read_u8' ()) inline_for_extraction let validate_u16 : validator tot_parse_u16 = @@ -76,11 +74,9 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_ inline_for_extraction let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) -(* inline_for_extraction let validate_and_read_u16 : validate_and_read tot_parse_u16 = - validate_and_read_intro validate_u16 read_u16 -*) + validate_and_read_intro validate_u16 (read_u16' ()) inline_for_extraction let validate_u32 : validator tot_parse_u32 = @@ -113,11 +109,9 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_ inline_for_extraction let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) -(* inline_for_extraction let validate_and_read_u32 : validate_and_read tot_parse_u32 = - validate_and_read_intro validate_u32 read_u32 -*) + validate_and_read_intro validate_u32 (read_u32' ()) inline_for_extraction let validate_u64 : validator tot_parse_u64 = @@ -150,8 +144,6 @@ fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_ inline_for_extraction let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) -(* inline_for_extraction let validate_and_read_u64 : validate_and_read tot_parse_u64 = - validate_and_read_intro validate_u64 read_u64 -*) + validate_and_read_intro validate_u64 (read_u64' ()) From 626d975df487417d589b4fc3f87d21258182e14d Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 15 Jul 2024 16:22:50 -0700 Subject: [PATCH 33/89] revert to non-CPS validators; jumpers with offsets instead of sizes --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 280 ++------- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 52 -- .../pulse/LowParse.Pulse.Combinators.fst | 587 ++---------------- src/lowparse/pulse/LowParse.Pulse.Int.fst | 16 - 4 files changed, 104 insertions(+), 831 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 0030d123d..401b6fa4f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -108,54 +108,47 @@ let validator_success (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t SZ.v offset <= Seq.length v && ( let pv = parse p (Seq.slice v (SZ.v offset) (Seq.length v)) in Some? pv && - snd (Some?.v pv) = SZ.v off + SZ.v offset + snd (Some?.v pv) = SZ.v off ) let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) : GTot bool = SZ.v offset <= Seq.length v && None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))) -inline_for_extraction -let validator_for (#t: Type0) (#k: parser_kind) (p: parser k t) - (input: slice byte) - (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) -: Tot Type -= - (pre: vprop) -> - (t': Type0) -> - (post: (t' -> vprop)) -> - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) -> - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -> - stt t' - (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) - (fun x -> post x) +let validator_postcond (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) (res: bool) : GTot bool = + if res + then validator_success p offset v off + else validator_failure p offset v + +module R = Pulse.Lib.Reference inline_for_extraction let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = (input: slice byte) -> - (offset: SZ.t) -> + (poffset: R.ref SZ.t) -> + (#offset: Ghost.erased SZ.t) -> (#pm: perm) -> (#v: Ghost.erased bytes) -> - validator_for p input offset pm v + stt bool + (pts_to input #pm v ** R.pts_to poffset offset ** pure (SZ.v offset <= Seq.length v)) + (fun res -> pts_to input #pm v ** (exists* off . R.pts_to poffset off ** pure (validator_postcond p offset v off res))) inline_for_extraction -let validate +```pulse +fn validate (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) (input: slice byte) - (offset: SZ.t) + (poffset: R.ref SZ.t) + (#offset: Ghost.erased SZ.t) (#pm: perm) (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -: stt t' - (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) - (fun x -> post x) -= w input offset #pm #v pre t' post ksucc kfail + requires pts_to input #pm v ** R.pts_to poffset offset ** pure (SZ.v offset <= Seq.length v) + returns res: bool + ensures pts_to input #pm v ** (exists* off . R.pts_to poffset off ** pure (validator_postcond p offset v off res)) +{ + w input poffset #offset #pm #v +} +``` let validate_nonempty_post (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) @@ -164,42 +157,6 @@ let validate_nonempty_post (off == 0sz <==> None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v)))) /\ (if off = 0sz then validator_failure p offset v else validator_success p offset v off) -inline_for_extraction -```pulse -fn validate_nonempty_success (#t: Type0) (#k: parser_kind) (p: parser k t { k.parser_kind_low > 0 }) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (sq: squash (SZ.v offset <= Seq.length v)) - (off: SZ.t) - requires pts_to input #pm v ** emp ** pure (validator_success p offset v off) - returns off: SZ.t - ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) -{ - parser_kind_prop_equiv k p; - off -} -``` - -inline_for_extraction -```pulse -fn validate_nonempty_failure (#t: Type0) (#k: parser_kind) (p: parser k t { k.parser_kind_low > 0 }) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (sq: squash (SZ.v offset <= Seq.length v)) - (_: unit) - requires pts_to input #pm v ** emp ** pure (validator_failure p offset v) - returns off: SZ.t - ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) -{ - parser_kind_prop_equiv k p; - 0sz -} -``` - inline_for_extraction ```pulse fn validate_nonempty (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p { k.parser_kind_low > 0 }) @@ -211,9 +168,14 @@ fn validate_nonempty (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validato returns off: SZ.t ensures pts_to input #pm v ** pure (validate_nonempty_post p offset v off) { - validate w input offset #pm #v emp SZ.t (fun off -> pts_to input #pm v ** pure (validate_nonempty_post p offset v off)) - (validate_nonempty_success p input offset #pm #v ()) - (validate_nonempty_failure p input offset #pm #v ()) + parser_kind_prop_equiv k p; + let mut poffset = offset; + let is_valid = validate w input poffset; + if is_valid { + !poffset; + } else { + 0sz + } } ``` @@ -234,22 +196,20 @@ fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (s }) : validator #_ #k p = (input: slice byte) - (offset: SZ.t) + (poffset: _) + (#offset: _) (#pm: perm) (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) { parser_kind_prop_equiv k p; pts_to_len input; + let offset = !poffset; if SZ.lt (SZ.sub (len input) offset) sz { - kfail () + false } else { - ksucc (sz <: SZ.t) + poffset := SZ.add offset sz; + true } } ``` @@ -327,7 +287,8 @@ fn jump_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { (#v: Ghost.erased bytes) { parser_kind_prop_equiv k p; - (sz <: SZ.t) + pts_to_len input; + SZ.add offset sz } ``` @@ -467,12 +428,12 @@ fn peek_stick_gen (#pm: perm) (#v: Ghost.erased bytes) (offset: SZ.t) - (consumed: SZ.t) - requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (consumed))) + (off: SZ.t) + requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (off))) returns input': slice byte ensures exists* v' . pts_to_serialized s input' #pm v' ** (pts_to_serialized s input' #pm v' @==> pts_to input #pm v) ** pure ( - validator_success #k #t p offset v consumed /\ - parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v consumed) + validator_success #k #t p offset v off /\ + parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { let split123 = slice_split_stick false input offset; @@ -481,6 +442,7 @@ fn peek_stick_gen unfold (slice_split_stick_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); stick_elim_partial_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); + let consumed = SZ.sub off offset; let split23 = peek_stick s input23 consumed; match split23 { SlicePair input2 input3 -> { unfold (peek_stick_post s input23 pm v23 consumed split23); @@ -537,15 +499,15 @@ fn read_from_validator_success (#pm: perm) (#v: Ghost.erased bytes) (offset: SZ.t) - (consumed: SZ.t) - requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (consumed))) + (off: SZ.t) + requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (off))) returns v' : t ensures pts_to input #pm v ** pure ( - validator_success #k #t p offset v consumed /\ - parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v consumed) + validator_success #k #t p offset v off /\ + parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { - let input' = peek_stick_gen s input offset consumed; + let input' = peek_stick_gen s input offset off; let res = r input'; elim_stick _ _; res @@ -656,147 +618,3 @@ fn reader_of_leaf_reader f x } ``` - -inline_for_extraction -let validate_and_read (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = - (input: slice byte) -> - (offset: SZ.t) -> - (#pm: perm) -> - (#v: Ghost.erased bytes) -> - (pre: vprop) -> - (t': Type0) -> - (post: (t' -> vprop)) -> - (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) -> - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -> - stt t' - (pts_to input #pm v ** pre ** pure (SZ.v offset <= Seq.length v)) - (fun x -> post x) - -inline_for_extraction -```pulse -fn validate_and_read_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validate_and_read #_ #k2 p2 = - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: _) - (kfail: _) -{ - v1 input offset #pm #v pre t' post ksucc kfail -} -``` - -inline_for_extraction -let validate_and_read_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validate_and_read p1 { k2 `is_weaker_than` k1 }) : validate_and_read (tot_weaken k2 p1) = - validate_and_read_ext v1 (tot_weaken k2 p1) - -inline_for_extraction -```pulse -fn validate_and_read_intro_cont_validate - (#t: Type0) (#k: parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (off: SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success p offset v off) - returns x' : t' - ensures post x' -{ - let x = read_from_validator_success r input offset off; - ksucc off x -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_intro - (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) (#s: serializer p) (r: leaf_reader s) -: validate_and_read #t #k p -= - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -{ - w input offset _ _ _ (validate_and_read_intro_cont_validate r input offset #pm #v pre t' post ksucc) kfail -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_elim_cont - (#t: Type0) (#k: parser_kind) (p: parser k t) - (input: slice byte) - (offset: SZ.t) - (pm: perm) - (v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) - (off: SZ.t) - (x: t) - requires pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) - returns x' : t' - ensures post x' -{ - ksucc off -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_elim - (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validate_and_read p) -: validator #t #k p -= - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success p offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure p offset v)) (fun x -> post x))) -{ - w input offset pre t' post (validate_and_read_elim_cont p input offset pm v pre t' post ksucc) kfail -} -``` - -inline_for_extraction -```pulse -fn ifthenelse_validate_and_read - (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) (wtrue: squash (cond == true) -> validate_and_read p) (wfalse: squash (cond == false) -> validate_and_read p) -: validate_and_read #t #k p -= - (input: _) - (offset: _) - (#pm: _) - (#v: _) - (pre: _) - (t': Type0) - (post: _) - (ksucc: _) - (kfail: _) -{ - if cond { - wtrue () input offset pre t' post ksucc kfail - } else { - wfalse () input offset pre t' post ksucc kfail - } -} -``` diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index 6ed815e2e..b08fbf2c1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -8,58 +8,6 @@ open Pulse.Lib.Slice #push-options "--print_universes" -inline_for_extraction -```pulse -fn stt_cps_ifthenelse - (#t: Type0) - (y: Ghost.erased t) -: if_combinator_weak u#4 (stt_cps u#0 #t y) -= (cond: _) - (ftrue: _) - (ffalse: _) - (pre: _) - (t': Type0) - (post: _) - (y': _) -{ - if cond { - ftrue () pre t' post y' - } else { - ffalse () pre t' post y' - } -} -``` - -inline_for_extraction -let validate_bitsum' - (#t: eqtype) - (#tot: pos) - (#cl: uint_t tot t) - (#b: bitsum' cl tot) - (f: filter_bitsum'_t b) - (d: destr_bitsum'_t b) - (#k: parser_kind) - (#p: parser k t) - (w: validate_and_read p) -: Pure (validate_and_read (tot_parse_bitsum' b p)) - (requires (k.parser_kind_subkind == Some ParserStrong)) - (ensures (fun _ -> True)) -= [@@inline_let] - let _ = synth_bitsum'_injective b in - (validate_and_read_synth - (validate_and_read_filter - w - (filter_bitsum' b) - (fun x -> f x) - ) - (synth_bitsum' b) - (d - stt_cps - stt_cps_ifthenelse - (fun k pre t' post phi -> phi k) - ) - ) - inline_for_extraction let jump_bitsum' (#t: eqtype) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index ae03e6a92..1e67efd65 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -15,16 +15,12 @@ fn validate_ret (x: t) : validator #t #parse_ret_kind (tot_parse_ret x) = (input: slice byte) - (offset: _) + (poffset: _) + (#offset: _) (#pm: _) (#v: _) - (pre: _) - (t': Type0) - (post: _) - (ksucc: _) - (kfail: _) { - ksucc 0sz + true } ``` @@ -52,14 +48,6 @@ let read_ret : Tot (reader (tot_serialize_ret x v_unique)) = reader_of_leaf_reader (leaf_read_ret x v_unique) -inline_for_extraction -let validate_and_read_ret - (#t: Type0) - (x: t) - (v_unique: (v' : t) -> Lemma (x == v')) -: Tot (validate_and_read (tot_parse_ret x)) -= validate_and_read_intro (validate_ret x) (leaf_read_ret x v_unique) - inline_for_extraction let jump_ret (#t: Type) (x: t) : jumper (tot_parse_ret x) = jump_constant_size (tot_parse_ret x) 0sz @@ -72,10 +60,6 @@ let leaf_read_empty : leaf_reader tot_serialize_empty = leaf_read_ret () (fun _ inline_for_extraction let read_empty : reader tot_serialize_empty = reader_of_leaf_reader leaf_read_empty -inline_for_extraction -let validate_and_read_empty : validate_and_read tot_parse_empty = - validate_and_read_intro validate_empty leaf_read_empty - inline_for_extraction let jump_empty : jumper tot_parse_empty = jump_ret () @@ -285,32 +269,6 @@ fn pts_to_serialized_synth_l2r_stick } ``` -inline_for_extraction -let stt_cps - (#t: Type) - (y: t) -: Tot Type -= (pre: vprop) -> (t': Type0) -> (post: (t' -> vprop)) -> (phi: ((y': t { y' == y }) -> stt t' pre (fun x -> post x))) -> stt t' pre (fun x -> post x) - -inline_for_extraction -let intro_stt_cps - (#t: Type) - (y: t) -: Tot (stt_cps y) -= fun pre t' post phi -> phi y - -inline_for_extraction -let elim_stt_cps - (#t: Type) - (y: Ghost.erased t) - (cps: stt_cps (Ghost.reveal y)) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (phi: ((y': t { y' == Ghost.reveal y }) -> stt t' pre (fun x -> post x))) -: stt t' pre (fun x -> post x) -= cps pre t' post phi - inline_for_extraction let read_synth_cont_t (#t: Type0) @@ -365,175 +323,55 @@ let read_synth' inline_for_extraction ```pulse -fn validate_and_read_synth_cont_cont - (#k1: parser_kind) (#t1: Type0) (p1: parser k1 t1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) - (input: slice byte) - (pm: perm) - (v: Ghost.erased bytes) - (offset: SZ.t) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (cont: (off: SZ.t) -> (x: t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_synth p1 f2) offset v off /\ parse (tot_parse_synth p1 f2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x)) - (off: SZ.t) - (x: t1 { validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off) }) - (x2: t2 { x2 == f2 x }) - requires pts_to input #pm v ** pre - returns x': t' - ensures post x' -{ - tot_parse_synth_eq p1 f2 (Seq.slice v (SZ.v offset) (Seq.length v)); - cont off x2 -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_synth_cont - (#k1: parser_kind) (#t1: Type0) (p1: parser k1 t1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) - (f2': (x: t1) -> stt_cps (f2 x)) - (input: slice byte) - (pm: perm) - (v: Ghost.erased bytes) - (offset: SZ.t) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (cont: (off: SZ.t) -> (x: t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_synth p1 f2) offset v off /\ parse (tot_parse_synth p1 f2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x)) - (off: SZ.t) - (x: t1) - requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) - returns x': t' - ensures post x' -{ - f2' x (pts_to input #pm v ** pre) t' post (validate_and_read_synth_cont_cont p1 f2 input pm v offset pre t' post cont off x) -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_synth - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (w: validate_and_read p1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) - (f2': (x: t1) -> stt_cps (f2 x)) -: validate_and_read #t2 #k1 (tot_parse_synth p1 f2) -= (input: slice byte) - (offset: _) - (#pm: _) - (#v: _) - (pre: _) - (t': Type0) - (post: _) - (ksucc: _) - (kfail: _) -{ - tot_parse_synth_eq p1 f2 (Seq.slice v (SZ.v offset) (Seq.length v)); - w input offset #pm #v pre t' post (validate_and_read_synth_cont p1 f2 f2' input pm v offset pre t' post ksucc) kfail -} -``` - -inline_for_extraction -let validate_and_read_synth' - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (w: validate_and_read p1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) -: Tot (validate_and_read (tot_parse_synth p1 f2)) -= validate_and_read_synth w f2 (fun x -> intro_stt_cps (f2 x)) - -inline_for_extraction -```pulse -fn validate_and_read_filter_cont_failure +fn validate_filter (#t: Type0) (#k: parser_kind) - (p: tot_parser k t) - (f: (t -> bool)) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) - (_: unit) - requires (pts_to input #pm v ** pre ** pure (validator_failure (p) offset v)) - returns res: t' - ensures post res -{ - tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); - kfail () -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_filter_cont_success - (#t: Type0) - (#k: parser_kind) - (p: tot_parser k t) + (#p: tot_parser k t) + (w: validator p) + (#s: serializer p) + (r: leaf_reader s) (f: (t -> bool)) (f': (x: t) -> (y: bool { y == f x })) +: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) += (input: slice byte) - (offset: SZ.t) + (poffset: _) + (#offset: _) (#pm: perm) (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: parse_filter_refine f) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off /\ parse (tot_parse_filter p f) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) - (off: SZ.t) - (x: t) - requires pts_to input #pm v ** pre ** pure (validator_success p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) - returns res: t' - ensures post res { tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); - if f' x { - ksucc off x + let offset = !poffset; + let is_valid = w input poffset; + if is_valid { + let off = !poffset; + let x = read_from_validator_success r input offset off; + f' x } else { - kfail () + false } } ``` inline_for_extraction -```pulse -fn validate_and_read_filter +let validate_filter'_test (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) - (w: validate_and_read p) (f: (t -> bool)) - (f': (x: t) -> (y: bool { y == f x })) -: validate_and_read #_ #(parse_filter_kind k) (tot_parse_filter p f) -= - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: parse_filter_refine f) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_filter p f) offset v off /\ parse (tot_parse_filter p f) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_filter p f) offset v)) (fun x -> post x))) -{ - w input offset #pm #v pre t' post - (validate_and_read_filter_cont_success p f f' input offset #pm #v pre t' post ksucc kfail) - (validate_and_read_filter_cont_failure p f input offset #pm #v pre t' post kfail) -} -``` + (x: t) +: Tot (y: bool { y == f x }) += f x inline_for_extraction -let validate_and_read_filter' +let validate_filter' (#t: Type0) (#k: parser_kind) (#p: tot_parser k t) - (w: validate_and_read p) + (w: validator p) + (#s: serializer p) + (r: leaf_reader s) (f: (t -> bool)) -: validate_and_read #_ #(parse_filter_kind k) (tot_parse_filter p f) -= validate_and_read_filter w f (fun x -> f x) +: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) += validate_filter w r f (validate_filter'_test f) inline_for_extraction ```pulse @@ -656,117 +494,6 @@ let nondep_then_eq_dtuple2 = tot_parse_synth_eq (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 input -inline_for_extraction -```pulse -fn validate_nondep_then_cont_success2 - (#t1 #t2: Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: parser k2 t2) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) - (off': SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success p2 (offset `SZ.add` off) v off') - returns x: t' - ensures post x -{ - pts_to_len input; // for SZ.fits (off + off') - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); - ksucc (off `SZ.add` off') -} -``` - -inline_for_extraction -```pulse -fn validate_nondep_then_cont_failure2 - (#t1 #t2: Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: parser k2 t2) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (off: SZ.t {validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) }) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) - (_: unit) - requires pts_to input #pm v ** pre ** pure (validator_failure p2 (offset `SZ.add` off) v) - returns x: t' - ensures post x -{ - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); - kfail () -} -``` - -inline_for_extraction -```pulse -fn validate_nondep_then_cont_success1 - (#t1 #t2: Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (#p2: parser k2 t2) - (v2: validator p2) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) - (off: SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off) - returns x: t' - ensures post x -{ - pts_to_len input; // for SZ.fits (offset + off) - v2 input (offset `SZ.add` off) #pm #v pre t' post - (validate_nondep_then_cont_success2 p1 p2 input offset #pm #v pre t' post off ksucc) - (validate_nondep_then_cont_failure2 p1 p2 input offset #pm #v pre t' post off kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_nondep_then_cont_failure1 - (#t1 #t2: Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: parser k2 t2) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) - (_: unit) - requires pts_to input #pm v ** pre ** pure (validator_failure p1 offset v) - returns x: t' - ensures post x -{ - nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); - kfail () -} -``` - inline_for_extraction ```pulse fn validate_nondep_then @@ -780,133 +507,18 @@ fn validate_nondep_then : validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_nondep_then p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) -{ - v1 input offset #pm #v pre t' post - (validate_nondep_then_cont_success1 p1 v2 input offset #pm #v pre t' post ksucc kfail) - (validate_nondep_then_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_dtuple2_cont_success2 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (key: Ghost.erased t1) - (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) - (off': SZ.t) - requires pts_to input #pm v ** pre ** pure (validator_success (p2 key) (offset `SZ.add` off) v off') - returns x: t' - ensures post x -{ - pts_to_len input; // for SZ.fits (off + off') - ksucc (off `SZ.add` off') -} -``` - -inline_for_extraction -```pulse -fn validate_dtuple2_cont_failure2 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (key: Ghost.erased t1) - (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == Ghost.reveal key }) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) - (_: unit) - requires pts_to input #pm v ** pre ** pure (validator_failure (p2 key) (offset `SZ.add` off) v) - returns x: t' - ensures post x -{ - kfail () -} -``` - -inline_for_extraction -```pulse -fn validate_dtuple2_cont_success1 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) + (poffset: _) + (#offset: _) (#pm: perm) (#v: Ghost.erased bytes) - (v2: (x: t1) -> validator (p2 x)) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) - (off: SZ.t) - (x: t1) - requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) - returns x: t' - ensures post x { - pts_to_len input; // for SZ.fits (offset + off) - v2 x input (offset `SZ.add` off) #pm #v pre t' post - (validate_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post x off ksucc) - (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post x off kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_dtuple2_cont_failure1 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) - (_: unit) - requires pts_to input #pm v ** pre ** pure (validator_failure p1 offset v) - returns x: t' - ensures post x -{ - kfail () + nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); + let is_valid1 = v1 input poffset; + if is_valid1 { + v2 input poffset + } else { + false + } } ``` @@ -917,116 +529,29 @@ fn validate_dtuple2 (#t2: t1 -> Type0) (#k1: parser_kind) (#p1: parser k1 t1) - (v1: validate_and_read p1) + (v1: validator p1) + (#s1: serializer p1) + (r1: leaf_reader s1) (#k2: parser_kind) (#p2: ((x: t1) -> parser k2 (t2 x))) (v2: ((x: t1) -> validator (p2 x))) : validator #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off)) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) -{ - v1 input offset #pm #v pre t' post - (validate_dtuple2_cont_success1 p1 p2 input offset #pm #v v2 pre t' post ksucc kfail) - (validate_dtuple2_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_dtuple2_cont_success2 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) + (poffset: _) + (#offset: _) (#pm: perm) (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (key: t1) - (off: SZ.t { validator_success p1 offset v off /\ SZ.fits (SZ.v offset + SZ.v off) /\ fst (Some?.v (parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)))) == key }) - (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (off': SZ.t) - (x: t2 key) - requires pts_to input #pm v ** pre ** pure (validator_success (p2 key) (offset `SZ.add` off) v off' /\ parse (p2 key) (Seq.slice v (SZ.v (offset `SZ.add` off)) (Seq.length v)) == Some (x, SZ.v off')) - returns x: t' - ensures post x { - pts_to_len input; // for SZ.fits (off + off') - ksucc (off `SZ.add` off') (| key, x |) -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_dtuple2_cont_success1 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (p1: parser k1 t1) - (#k2: parser_kind) - (p2: ((x: t1) -> parser k2 (t2 x))) - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (v2: (x: t1) -> validate_and_read (p2 x)) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_parse_dtuple2 p1 p2) offset v)) (fun x -> post x))) - (off: SZ.t) - (x: t1) - requires pts_to input #pm v ** pre ** pure (validator_success p1 offset v off /\ parse p1 (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off)) - returns x: t' - ensures post x -{ - pts_to_len input; // for SZ.fits (offset + off) - v2 x input (offset `SZ.add` off) #pm #v pre t' post - (validate_and_read_dtuple2_cont_success2 p1 p2 input offset #pm #v pre t' post x off ksucc) - (validate_dtuple2_cont_failure2 p1 p2 input offset #pm #v pre t' post x off kfail) -} -``` - -inline_for_extraction -```pulse -fn validate_and_read_dtuple2 - (#t1: Type0) - (#t2: t1 -> Type0) - (#k1: parser_kind) - (#p1: parser k1 t1) - (v1: validate_and_read p1) - (#k2: parser_kind) - (#p2: ((x: t1) -> parser k2 (t2 x))) - (v2: ((x: t1) -> validate_and_read (p2 x))) -: validate_and_read #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) -= - (input: slice byte) - (offset: SZ.t) - (#pm: perm) - (#v: Ghost.erased bytes) - (pre: vprop) - (t': Type0) - (post: (t' -> vprop)) - (ksucc: ((off: SZ.t) -> (x: dtuple2 t1 t2) -> stt t' (pts_to input #pm v ** pre ** pure (validator_success (tot_parse_dtuple2 p1 p2) offset v off /\ parse (tot_parse_dtuple2 p1 p2) (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (x, SZ.v off))) (fun x -> post x))) - (kfail: (unit -> stt t' (pts_to input #pm v ** pre ** pure (validator_failure (tot_nondep_then p1 p2) offset v)) (fun x -> post x))) -{ - v1 input offset #pm #v pre t' post - (validate_and_read_dtuple2_cont_success1 p1 p2 input offset #pm #v v2 pre t' post ksucc kfail) - (validate_dtuple2_cont_failure1 p1 p2 input offset #pm #v pre t' post kfail) + let offset = !poffset; + let is_valid1 = v1 input poffset; + if is_valid1 { + let off = !poffset; + let x = read_from_validator_success r1 input offset off; + v2 x input poffset + } else { + false + } } ``` @@ -1049,9 +574,8 @@ fn jump_nondep_then { nondep_then_eq #k1 p1 #k2 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); pts_to_len input; - let consumed1 = v1 input offset; - let consumed2 = v2 input (offset `SZ.add` consumed1); - SZ.add consumed1 consumed2 + let off1 = v1 input offset; + v2 input off1 } ``` @@ -1078,8 +602,7 @@ fn jump_dtuple2 pts_to_len input; let off1 = v1 input offset; let x = read_from_validator_success r1 input offset off1; - let off2 = v2 x input (SZ.add offset off1); - SZ.add off1 off2 + v2 x input off1 } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 920b48aaa..81f6702f5 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -39,10 +39,6 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 inline_for_extraction let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) -inline_for_extraction -let validate_and_read_u8 : validate_and_read tot_parse_u8 = - validate_and_read_intro validate_u8 (read_u8' ()) - inline_for_extraction let validate_u16 : validator tot_parse_u16 = validate_total_constant_size tot_parse_u16 2sz @@ -74,10 +70,6 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_ inline_for_extraction let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) -inline_for_extraction -let validate_and_read_u16 : validate_and_read tot_parse_u16 = - validate_and_read_intro validate_u16 (read_u16' ()) - inline_for_extraction let validate_u32 : validator tot_parse_u32 = validate_total_constant_size tot_parse_u32 4sz @@ -109,10 +101,6 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_ inline_for_extraction let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) -inline_for_extraction -let validate_and_read_u32 : validate_and_read tot_parse_u32 = - validate_and_read_intro validate_u32 (read_u32' ()) - inline_for_extraction let validate_u64 : validator tot_parse_u64 = validate_total_constant_size tot_parse_u64 8sz @@ -143,7 +131,3 @@ fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_ inline_for_extraction let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) - -inline_for_extraction -let validate_and_read_u64 : validate_and_read tot_parse_u64 = - validate_and_read_intro validate_u64 (read_u64' ()) From f85e1c088f28ec63a414ac84ccfd98525d06f4b3 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 15 Jul 2024 16:44:57 -0700 Subject: [PATCH 34/89] eta-expand validate_synth because of the pure bind --- .../pulse/LowParse.Pulse.Combinators.fst | 17 +++++++++++++---- 1 file changed, 13 insertions(+), 4 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 1e67efd65..3c354858f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -91,15 +91,24 @@ let tot_parse_synth_eq' = parse_synth_eq #k p1 f2 b inline_for_extraction -let validate_synth +```pulse +fn validate_synth (#t #t': Type) (#k: parser_kind) (#p: tot_parser k t) (w: validator p) (f: (t -> t') { synth_injective f }) -: Tot (validator (tot_parse_synth p f)) -= Classical.forall_intro (tot_parse_synth_eq' p f); - w +: validator #t' #k (tot_parse_synth p f) += (input: slice byte) + (poffset: _) + (#offset: _) + (#pm: _) + (#v: _) +{ + tot_parse_synth_eq' p f (Seq.slice v (SZ.v offset) (Seq.length v)); + w input poffset #offset #pm #v +} +``` inline_for_extraction ```pulse From 0238383fa415204e1c1dc57de02aa12d022f5f13 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 15 Jul 2024 19:43:05 -0700 Subject: [PATCH 35/89] ifthenelse_validator --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 23 ++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 401b6fa4f..9345b97f4 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -150,6 +150,29 @@ fn validate } ``` +inline_for_extraction +```pulse +fn ifthenelse_validator + (#t: Type0) (#k: parser_kind) (p: parser k t) + (cond: bool) + (wtrue: squash (cond == true) -> validator p) + (wfalse: squash (cond == false) -> validator p) +: validator #t #k p += + (input: slice byte) + (poffset: R.ref SZ.t) + (#offset: Ghost.erased SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + if cond { + wtrue () input poffset + } else { + wfalse () input poffset + } +} +``` + let validate_nonempty_post (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) : Tot prop From fc169ddea346266c63ecdecce66068824cf32baa Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 16 Jul 2024 14:41:56 -0700 Subject: [PATCH 36/89] bound offset on validator failure --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 9345b97f4..eaf2535b8 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -116,9 +116,11 @@ let validator_failure (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t None? (parse p (Seq.slice v (SZ.v offset) (Seq.length v))) let validator_postcond (#k: parser_kind) (#t: Type) (p: parser k t) (offset: SZ.t) (v: bytes) (off: SZ.t) (res: bool) : GTot bool = + SZ.v off <= Seq.length v && ( if res then validator_success p offset v off else validator_failure p offset v +) module R = Pulse.Lib.Reference From d75f3b63781b6213e83550edb54ec05afa0e2ae4 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 17 Jul 2024 10:56:36 -0700 Subject: [PATCH 37/89] jump_vclist; dtuple2, nondep_then, nlist accessors --- .../pulse/LowParse.Pulse.Combinators.fst | 118 +++++++++++ src/lowparse/pulse/LowParse.Pulse.Util.fst | 17 ++ src/lowparse/pulse/LowParse.Pulse.VCList.fst | 199 ++++++++++++++++++ 3 files changed, 334 insertions(+) create mode 100644 src/lowparse/pulse/LowParse.Pulse.VCList.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 3c354858f..b24b9eeff 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -694,6 +694,66 @@ fn split_dtuple2 } ``` +inline_for_extraction +```pulse +fn dtuple2_dfst + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: ((x: t1) -> parser k2 (t2 x))) + (s2: (x: t1) -> serializer (p2 x)) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (dtuple2 t1 t2)) + requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + returns res: slice byte + ensures pts_to_serialized s1 res #pm (dfst v) ** + (pts_to_serialized s1 res #pm (dfst v) @==> pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) +{ + let spl = split_dtuple2 s1 j1 s2 input; + match spl { SlicePair input1 input2 -> { + unfold (split_dtuple2_post s1 s2 input pm v spl); + unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); + stick_elim_partial_r _ _ _; + input1 + }} +} +``` + +inline_for_extraction +```pulse +fn dtuple2_dsnd + (#t1: Type0) + (#t2: t1 -> Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: ((x: t1) -> parser k2 (t2 x))) + (s2: (x: t1) -> serializer (p2 x)) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (dtuple2 t1 t2)) + requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + returns res: slice byte + ensures pts_to_serialized (s2 (dfst v)) res #pm (dsnd v) ** + (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v) @==> pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) +{ + let spl = split_dtuple2 s1 j1 s2 input; + match spl { SlicePair input1 input2 -> { + unfold (split_dtuple2_post s1 s2 input pm v spl); + unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); + stick_elim_partial_l _ _ _; + input2 + }} +} +``` + let split_nondep_then_post' (#t1 #t2: Type0) (#k1: parser_kind) @@ -776,6 +836,64 @@ fn split_nondep_then } ``` +inline_for_extraction +```pulse +fn nondep_then_fst + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (t1 & t2)) + requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + returns res: slice byte + ensures pts_to_serialized s1 res #pm (fst v) ** + (pts_to_serialized s1 res #pm (fst v) @==> pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) +{ + let spl = split_nondep_then s1 j1 s2 input; + match spl { SlicePair input1 input2 -> { + unfold (split_nondep_then_post s1 s2 input pm v spl); + unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); + stick_elim_partial_r _ _ _; + input1 + }} +} +``` + +inline_for_extraction +```pulse +fn nondep_then_snd + (#t1 #t2: Type0) + (#k1: parser_kind) + (#p1: parser k1 t1) + (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) + (j1: jumper p1) + (#k2: parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (t1 & t2)) + requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + returns res: slice byte + ensures pts_to_serialized s2 res #pm (snd v) ** + (pts_to_serialized s2 res #pm (snd v) @==> pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) +{ + let spl = split_nondep_then s1 j1 s2 input; + match spl { SlicePair input1 input2 -> { + unfold (split_nondep_then_post s1 s2 input pm v spl); + unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); + stick_elim_partial_l _ _ _; + input2 + }} +} +``` + inline_for_extraction ```pulse fn read_dtuple2 diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index e84ba5661..5828ae74b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -5,6 +5,23 @@ include Pulse.Lib.Stick module S = Pulse.Lib.Slice module SZ = FStar.SizeT +```pulse +ghost +fn stick_id + (p: vprop) +requires emp +ensures (p @==> p) +{ + ghost fn aux (_: unit) + requires emp ** p + ensures p + { + () + }; + intro_stick p p emp aux +} +``` + ```pulse ghost fn stick_trans diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst new file mode 100644 index 000000000..86224760a --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -0,0 +1,199 @@ +module LowParse.Pulse.VCList +include LowParse.Spec.VCList +include LowParse.Pulse.Combinators +open FStar.Tactics.V2 +open LowParse.Pulse.Util +open Pulse.Lib.Stick +open Pulse.Lib.Slice + +module SZ = FStar.SizeT +module R = Pulse.Lib.Reference + +inline_for_extraction +```pulse +fn jump_nlist + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (j: jumper p) + (n0: SZ.t) +: jumper #(nlist (SZ.v n0) t) #(parse_nlist_kind (SZ.v n0) k) (tot_parse_nlist (SZ.v n0) p) += + (input: slice byte) + (offset0: SZ.t) + (#pm: perm) + (#v: Ghost.erased bytes) +{ + let mut pn = n0; + let mut poffset = offset0; + while ( + let n = !pn; + (SZ.gt n 0sz) + ) invariant b . exists* n offset . ( + pts_to input #pm v ** R.pts_to pn n ** R.pts_to poffset offset ** pure ( + SZ.v offset <= Seq.length v /\ ( + let pr0 = parse_consume (tot_parse_nlist (SZ.v n0) p) (Seq.slice v (SZ.v offset0) (Seq.length v)) in + let pr = parse_consume (tot_parse_nlist (SZ.v n) p) (Seq.slice v (SZ.v offset) (Seq.length v)) in + Some? pr0 /\ Some? pr /\ + SZ.v offset0 + Some?.v pr0 == SZ.v offset + Some?.v pr /\ + b == (SZ.v n > 0) + ))) { + let n = !pn; + let offset = !poffset; + tot_parse_nlist_eq (SZ.v n) p (Seq.slice v (SZ.v offset) (Seq.length v)); + let offset' = j input offset; + pn := (SZ.sub n 1sz); + poffset := offset'; + }; + !poffset +} +``` + +```pulse +ghost +fn nlist_cons_as_nondep_then + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (n: pos) + (input: slice byte) + (#pm: perm) + (#v: nlist n t) +requires + pts_to_serialized (tot_serialize_nlist n s) input #pm v +ensures exists* v' . + pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v' ** + (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + pure ( + v == (fst v' :: snd v') + ) +{ + synth_inverse_1 t (n - 1); + synth_inverse_2 t (n - 1); + rewrite_with_stick + (pts_to_serialized (tot_serialize_nlist n s) input #pm v) + (pts_to_serialized (tot_serialize_synth _ (synth_nlist (n - 1)) (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); + pts_to_serialized_synth_l2r_stick + (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) + (synth_nlist (n - 1)) + (synth_nlist_recip (n - 1)) + input; + stick_trans (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm _) _ _ +} +``` + +inline_for_extraction +```pulse +fn nlist_hd + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (j: jumper p) + (n: pos) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (nlist n t)) +requires + pts_to_serialized (tot_serialize_nlist n s) input #pm v +returns input' : slice byte +ensures exists* v' . + pts_to_serialized s input' #pm v' ** + (pts_to_serialized s input' #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + pure ( + Cons? v /\ + v' == List.Tot.hd v + ) +{ + nlist_cons_as_nondep_then s n input; + let res = nondep_then_fst s j (tot_serialize_nlist (n - 1) s) input; + stick_trans (pts_to_serialized s res #pm _) _ _; + res +} +``` + +inline_for_extraction +```pulse +fn nlist_tl + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (j: jumper p) + (n: pos) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (nlist n t)) +requires + pts_to_serialized (tot_serialize_nlist n s) input #pm v +returns input' : slice byte +ensures exists* v' . + pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v' ** + (pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + pure ( + Cons? v /\ + v' == List.Tot.tl v + ) +{ + nlist_cons_as_nondep_then s n input; + let res = nondep_then_snd s j (tot_serialize_nlist (n - 1) s) input; + stick_trans (pts_to_serialized (tot_serialize_nlist (n - 1) s) res #pm _) _ _; + res +} +``` + +inline_for_extraction +```pulse +fn nlist_nth + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (j: jumper p) + (n0: SZ.t) + (input: slice byte) + (#pm: perm) + (#v0: Ghost.erased (nlist (SZ.v n0) t)) + (i0: SZ.t { SZ.v i0 < SZ.v n0 }) +requires + pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0 +returns input' : slice byte +ensures exists* v . + pts_to_serialized s input' #pm v ** + (pts_to_serialized s input' #pm v @==> pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + pure (v == List.Tot.index v0 (SZ.v i0)) +{ + stick_id (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); + let mut pi = 0sz; + let mut pres = input; + while ( + let i = !pi; + (SZ.lt i i0) + ) invariant b . exists* i res (n: nat) (v: nlist n t) . ( + R.pts_to pi i ** R.pts_to pres res ** + pts_to_serialized (tot_serialize_nlist n s) res #pm v ** + (pts_to_serialized (tot_serialize_nlist n s) res #pm v @==> pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + pure ( + SZ.v i <= SZ.v i0 /\ + (b == (SZ.v i < SZ.v i0)) /\ + n == SZ.v n0 - SZ.v i /\ + List.Tot.index v0 (SZ.v i0) == List.Tot.index v (SZ.v i0 - SZ.v i) + )) { + let res = !pres; + let i = !pi; + let res2 = nlist_tl s j (SZ.v n0 - SZ.v i) res; + pi := (SZ.add i 1sz); + pres := res2; + stick_trans + (pts_to_serialized (tot_serialize_nlist (SZ.v n0 - SZ.v i - 1) s) res2 #pm _) + _ + _ + }; + let res = !pres; + let res2 = nlist_hd s j (SZ.v n0 - SZ.v i0) res; + stick_trans + (pts_to_serialized s res2 #pm _) _ _; + res2 +} +``` From 4c57be367be2ed5b999d1f78323dfc115618aec5 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 17 Jul 2024 17:32:14 -0700 Subject: [PATCH 38/89] pts_to_serialized_nlist_1, pts_to_serialized_length, etc. --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 13 +++++ .../pulse/LowParse.Pulse.Combinators.fst | 28 +++++++++++ src/lowparse/pulse/LowParse.Pulse.VCList.fst | 49 +++++++++++++++++++ 3 files changed, 90 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index eaf2535b8..fb8dd3dd3 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -34,6 +34,19 @@ fn pts_to_serialized_elim_stick } ``` +```pulse +ghost +fn pts_to_serialized_length + (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: t) + requires (pts_to_serialized s input #pm v) + ensures (pts_to_serialized s input #pm v ** pure (Seq.length (bare_serialize s v) == SZ.v (len input))) +{ + unfold (pts_to_serialized s input #pm v); + pts_to_len input; + fold (pts_to_serialized s input #pm v) +} +``` + let serializer_ext_eq (#t: Type0) (#k1: parser_kind) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index b24b9eeff..43be563a9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -278,6 +278,34 @@ fn pts_to_serialized_synth_l2r_stick } ``` +```pulse +ghost +fn pts_to_serialized_synth_l2r_stick' + (#t #t': Type0) + (#k_: parser_kind) + (#p_: tot_parser k_ t') + (#s_: tot_serializer p_) + (#k: parser_kind) + (#p: tot_parser k t) + (s: tot_serializer p) + (f: (t -> t') { synth_injective f }) + (f': (t' -> t) { synth_inverse f f' }) + (input: slice byte) + (#pm: perm) + (#v: t') + requires pts_to_serialized s_ input #pm v ** pure (forall x . parse p_ x == parse (tot_parse_synth p f) x) + ensures pts_to_serialized s input #pm (f' v) ** (pts_to_serialized s input #pm (f' v) @==> pts_to_serialized s_ input #pm v) +{ + pts_to_serialized_ext_stick + s_ + (tot_serialize_synth p f s f' ()) + input; + pts_to_serialized_synth_l2r_stick + s f f' input; + stick_trans _ _ (pts_to_serialized s_ input #pm v) +} +``` + inline_for_extraction let read_synth_cont_t (#t: Type0) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 86224760a..233513b51 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -197,3 +197,52 @@ ensures exists* v . res2 } ``` + +let synth_nlist_1 + (#t: Type) + (x: t) +: Tot (nlist 1 t) += [x] + +let synth_nlist_1_recip + (#t: Type) + (x: nlist 1 t) +: Tot t += List.Tot.hd x + +let tot_parse_nlist_1_eq + (#t: Type) + (#k: parser_kind) + (p: parser k t) + (b: bytes) +: Lemma + (parse (tot_parse_nlist 1 p) b == parse (tot_parse_synth p synth_nlist_1) b) += tot_parse_nlist_eq 1 p b; + tot_parse_synth_eq p synth_nlist_1 b + +```pulse +fn pts_to_serialized_nlist_1 + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased t) + requires pts_to_serialized s input #pm v + ensures exists* v' . pts_to_serialized (tot_serialize_nlist 1 s) input #pm v' ** + (pts_to_serialized (tot_serialize_nlist 1 s) input #pm v' @==> + pts_to_serialized s input #pm v) ** + pure ((v' <: list t) == [Ghost.reveal v]) +{ + pts_to_serialized_synth_stick s synth_nlist_1 synth_nlist_1_recip input; + Classical.forall_intro (tot_parse_nlist_1_eq p); + pts_to_serialized_ext_stick + (tot_serialize_synth p synth_nlist_1 s synth_nlist_1_recip ()) + (tot_serialize_nlist 1 s) + input; + stick_trans + (pts_to_serialized (tot_serialize_nlist 1 s) input #pm _) + _ _ +} +``` From a8eb4004ac4a345a6d57f2eadd6ed2339269e0a2 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 18 Jul 2024 21:41:33 -0700 Subject: [PATCH 39/89] Util: stick -> trade --- src/lowparse/pulse/LowParse.Pulse.Util.fst | 173 +++++++++------------ 1 file changed, 75 insertions(+), 98 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 5828ae74b..c1f830311 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -1,203 +1,178 @@ module LowParse.Pulse.Util include Pulse.Lib.Pervasives -include Pulse.Lib.Stick +open Pulse.Lib.Trade module S = Pulse.Lib.Slice module SZ = FStar.SizeT +module T = FStar.Tactics ```pulse ghost -fn stick_id +fn trade_id + (#[T.exact (`emp_inames)] is:inames) (p: vprop) requires emp -ensures (p @==> p) +ensures (trade #is p p) { ghost fn aux (_: unit) requires emp ** p ensures p + opens is { () }; - intro_stick p p emp aux + intro_trade p p emp aux } ``` ```pulse ghost -fn stick_trans - (p1 p2 p3: vprop) - requires (p1 @==> p2) ** (p2 @==> p3) - ensures (p1 @==> p3) -{ - ghost - fn aux - (_foo: unit) - requires ((p1 @==> p2) ** (p2 @==> p3)) ** p1 - ensures p3 - { - elim_stick p1 p2; - elim_stick p2 p3 - }; - intro_stick p1 p3 ((p1 @==> p2) ** (p2 @==> p3)) aux -} -``` - -```pulse -ghost -fn stick_reg_l +fn trade_reg_l + (#is : inames) (p p1 p2: vprop) - requires (p1 @==> p2) - ensures ((p ** p1) @==> (p ** p2)) + requires (trade #is p1 p2) + ensures (trade #is (p ** p1) (p ** p2)) { ghost fn aux (_foo: unit) - requires ((p1 @==> p2) ** (p ** p1)) + requires ((trade #is p1 p2) ** (p ** p1)) ensures (p ** p2) + opens is { - elim_stick p1 p2 + elim_trade #is p1 p2 }; - intro_stick (p ** p1) (p ** p2) (p1 @==> p2) aux + intro_trade (p ** p1) (p ** p2) (trade #is p1 p2) aux } ``` ```pulse ghost -fn stick_eq +fn trade_eq + (#[T.exact (`emp_inames)] is:inames) (p1 p2: vprop) requires pure (p1 == p2) // ideally with `vprop_equivs ()` - ensures (p1 @==> p2) + ensures (trade #is p1 p2) { ghost fn aux (_foo: unit) requires pure (p1 == p2) ** p1 ensures p2 + opens is { rewrite p1 as p2 }; - intro_stick p1 p2 (pure (p1 == p2)) aux -} -``` - -```pulse -ghost -fn stick_rewrite_l - (l1 l2 r: vprop) - requires (l1 @==> r) ** pure (l1 == l2) - ensures l2 @==> r -{ - rewrite (l1 @==> r) as (l2 @==> r) -} -``` - -```pulse -ghost -fn stick_rewrite_r - (l r1 r2: vprop) - requires (l @==> r1) ** pure (r1 == r2) - ensures l @==> r2 -{ - rewrite (l @==> r1) as (l @==> r2) + intro_trade p1 p2 (pure (p1 == p2)) aux } ``` ```pulse ghost -fn stick_reg_r +fn trade_reg_r + (#is: inames) (p1 p2 p: vprop) - requires (p1 @==> p2) - ensures ((p1 ** p) @==> (p2 ** p)) + requires (trade #is p1 p2) + ensures (trade #is (p1 ** p) (p2 ** p)) { - stick_reg_l p p1 p2; + trade_reg_l p p1 p2; vprop_equivs (); - rewrite ((p ** p1) @==> (p ** p2)) - as ((p1 ** p) @==> (p2 ** p)) + rewrite (trade #is (p ** p1) (p ** p2)) + as (trade #is (p1 ** p) (p2 ** p)) } ``` ```pulse ghost -fn stick_weak_concl_l +fn trade_weak_concl_l + (#is: inames) (p1 p2 p: vprop) - requires (p1 @==> p2) ** p - ensures (p1 @==> (p ** p2)) + requires (trade #is p1 p2) ** p + ensures (trade #is p1 (p ** p2)) { ghost fn aux (_foo: unit) - requires ((p1 @==> p2) ** p) ** p1 + requires ((trade #is p1 p2) ** p) ** p1 ensures p ** p2 + opens is { - elim_stick p1 p2 + elim_trade #is p1 p2 }; - intro_stick p1 (p ** p2) ((p1 @==> p2) ** p) aux + intro_trade p1 (p ** p2) ((trade #is p1 p2) ** p) aux } ``` ```pulse ghost -fn stick_weak_concl_r +fn trade_weak_concl_r + (#is: inames) (p1 p2 p: vprop) - requires (p1 @==> p2) ** p - ensures (p1 @==> (p2 ** p)) + requires (trade #is p1 p2) ** p + ensures (trade #is p1 (p2 ** p)) { - stick_weak_concl_l p1 p2 p; + trade_weak_concl_l p1 p2 p; vprop_equivs (); - stick_eq (p ** p2) (p2 ** p); - stick_trans p1 _ _ + trade_eq is (p ** p2) (p2 ** p); // FIXME: why is the `is` argument explicit? + trade_compose p1 _ _ } ``` ```pulse ghost -fn stick_prod +fn trade_prod + (#is: inames) (l1 r1 l2 r2: vprop) - requires ((l1 @==> r1) ** (l2 @==> r2)) - ensures ((l1 ** l2) @==> (r1 ** r2)) + requires (trade #is l1 r1 ** trade #is l2 r2) + ensures (trade #is (l1 ** l2) (r1 ** r2)) { ghost fn aux (_foo: unit) - requires ((l1 @==> r1) ** (l2 @==> r2)) ** (l1 ** l2) + requires ((trade #is l1 r1) ** (trade #is l2 r2)) ** (l1 ** l2) ensures r1 ** r2 + opens is { - elim_stick l1 r1; - elim_stick l2 r2 + elim_trade #is l1 r1; + elim_trade #is l2 r2 }; - intro_stick (l1 ** l2) (r1 ** r2) ((l1 @==> r1) ** (l2 @==> r2)) aux + intro_trade (l1 ** l2) (r1 ** r2) ((trade #is l1 r1) ** (trade #is l2 r2)) aux } ``` ```pulse ghost -fn stick_elim_partial_l +fn trade_elim_partial_l + (#is: inames) (p p1 p2: vprop) - requires p ** ((p ** p1) @==> p2) - ensures p1 @==> p2 + requires p ** (trade #is (p ** p1) p2) + ensures trade #is p1 p2 { ghost fn aux (_foo: unit) - requires (p ** ((p ** p1) @==> p2)) ** p1 + requires (p ** (trade #is (p ** p1) p2)) ** p1 ensures p2 + opens is { - elim_stick (p ** p1) p2 + elim_trade #is (p ** p1) p2 }; - intro_stick p1 p2 (p ** ((p ** p1) @==> p2)) aux + intro_trade p1 p2 (p ** (trade #is (p ** p1) p2)) aux } ``` ```pulse ghost -fn stick_elim_partial_r +fn trade_elim_partial_r + (#is: inames) (p1 p p2: vprop) - requires ((p1 ** p) @==> p2) ** p - ensures p1 @==> p2 + requires (trade #is (p1 ** p) p2) ** p + ensures trade #is p1 p2 { vprop_equivs (); - stick_rewrite_l (p1 ** p) (p ** p1) p2; - stick_elim_partial_l p p1 p2 + rewrite (trade #is (p1 ** p) p2) + as (trade #is (p ** p1) p2); + trade_elim_partial_l p p1 p2 } ``` @@ -256,7 +231,7 @@ let slice_append_split_stick_post' = S.pts_to s1 #p v1 ** S.pts_to s2 #p v2 ** - ((S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) @==> S.pts_to s #p (v1 `Seq.append` v2)) + (trade (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) (S.pts_to s #p (v1 `Seq.append` v2))) let slice_append_split_stick_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) @@ -289,7 +264,7 @@ fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm S.SlicePair input1 input2 -> { unfold (slice_append_split_post input p v1 v2 i res); unfold (slice_append_split_post' input p v1 v2 i input1 input2); - intro_stick _ _ _ (slice_append_split_stick_aux input p v1 v2 i input1 input2); + intro_trade _ _ _ (slice_append_split_stick_aux input p v1 v2 i input1 input2); fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); fold (slice_append_split_stick_post input p v1 v2 i (S.SlicePair input1 input2)); (S.SlicePair input1 input2) @@ -306,7 +281,7 @@ let slice_split_stick_post' = exists* v1 v2 . S.pts_to s1 #p v1 ** S.pts_to s2 #p v2 ** - ((S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) @==> S.pts_to s #p v) ** + (trade (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) (S.pts_to s #p v)) ** pure ( SZ.v i <= Seq.length v /\ (v1, v2) == Seq.split v (SZ.v i) @@ -345,7 +320,7 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos unfold (S.split_post s p v i res); unfold (S.split_post' s p v i s1 s2); with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); - intro_stick _ _ _ (slice_split_stick_aux s p v i s1 s2 v1 v2 ()); + intro_trade _ _ _ (slice_split_stick_aux s p v i s1 s2 v1 v2 ()); fold (slice_split_stick_post' s p v i s1 s2); fold (slice_split_stick_post s p v i (S.SlicePair s1 s2)); (S.SlicePair s1 s2) @@ -355,10 +330,11 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos ```pulse ghost -fn rewrite_with_stick +fn rewrite_with_trade + (#[T.exact (`emp_inames)] is:inames) (p1 p2: vprop) requires p1 ** pure (p1 == p2) - ensures p2 ** (p2 @==> p1) + ensures p2 ** (trade #is p2 p1) { rewrite p1 as p2; ghost @@ -366,9 +342,10 @@ fn rewrite_with_stick (_: unit) requires emp ** p2 ensures p1 + opens is { rewrite p2 as p1 }; - intro_stick _ _ _ aux + intro_trade _ _ _ aux } ``` From 6dee4b94324c4397fc55f034a7caa41b14f118a0 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 18 Jul 2024 22:12:04 -0700 Subject: [PATCH 40/89] move all trade lemmas to Pulse.Lib.Trade.Util --- src/lowparse/pulse/LowParse.Pulse.Util.fst | 192 --------------------- 1 file changed, 192 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index c1f830311..3ecc9804d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -6,176 +6,6 @@ module S = Pulse.Lib.Slice module SZ = FStar.SizeT module T = FStar.Tactics -```pulse -ghost -fn trade_id - (#[T.exact (`emp_inames)] is:inames) - (p: vprop) -requires emp -ensures (trade #is p p) -{ - ghost fn aux (_: unit) - requires emp ** p - ensures p - opens is - { - () - }; - intro_trade p p emp aux -} -``` - -```pulse -ghost -fn trade_reg_l - (#is : inames) - (p p1 p2: vprop) - requires (trade #is p1 p2) - ensures (trade #is (p ** p1) (p ** p2)) -{ - ghost - fn aux - (_foo: unit) - requires ((trade #is p1 p2) ** (p ** p1)) - ensures (p ** p2) - opens is - { - elim_trade #is p1 p2 - }; - intro_trade (p ** p1) (p ** p2) (trade #is p1 p2) aux -} -``` - -```pulse -ghost -fn trade_eq - (#[T.exact (`emp_inames)] is:inames) - (p1 p2: vprop) - requires pure (p1 == p2) // ideally with `vprop_equivs ()` - ensures (trade #is p1 p2) -{ - ghost - fn aux - (_foo: unit) - requires pure (p1 == p2) ** p1 - ensures p2 - opens is - { - rewrite p1 as p2 - }; - intro_trade p1 p2 (pure (p1 == p2)) aux -} -``` - -```pulse -ghost -fn trade_reg_r - (#is: inames) - (p1 p2 p: vprop) - requires (trade #is p1 p2) - ensures (trade #is (p1 ** p) (p2 ** p)) -{ - trade_reg_l p p1 p2; - vprop_equivs (); - rewrite (trade #is (p ** p1) (p ** p2)) - as (trade #is (p1 ** p) (p2 ** p)) -} -``` - -```pulse -ghost -fn trade_weak_concl_l - (#is: inames) - (p1 p2 p: vprop) - requires (trade #is p1 p2) ** p - ensures (trade #is p1 (p ** p2)) -{ - ghost - fn aux - (_foo: unit) - requires ((trade #is p1 p2) ** p) ** p1 - ensures p ** p2 - opens is - { - elim_trade #is p1 p2 - }; - intro_trade p1 (p ** p2) ((trade #is p1 p2) ** p) aux -} -``` - -```pulse -ghost -fn trade_weak_concl_r - (#is: inames) - (p1 p2 p: vprop) - requires (trade #is p1 p2) ** p - ensures (trade #is p1 (p2 ** p)) -{ - trade_weak_concl_l p1 p2 p; - vprop_equivs (); - trade_eq is (p ** p2) (p2 ** p); // FIXME: why is the `is` argument explicit? - trade_compose p1 _ _ -} -``` - -```pulse -ghost -fn trade_prod - (#is: inames) - (l1 r1 l2 r2: vprop) - requires (trade #is l1 r1 ** trade #is l2 r2) - ensures (trade #is (l1 ** l2) (r1 ** r2)) -{ - ghost - fn aux - (_foo: unit) - requires ((trade #is l1 r1) ** (trade #is l2 r2)) ** (l1 ** l2) - ensures r1 ** r2 - opens is - { - elim_trade #is l1 r1; - elim_trade #is l2 r2 - }; - intro_trade (l1 ** l2) (r1 ** r2) ((trade #is l1 r1) ** (trade #is l2 r2)) aux -} -``` - -```pulse -ghost -fn trade_elim_partial_l - (#is: inames) - (p p1 p2: vprop) - requires p ** (trade #is (p ** p1) p2) - ensures trade #is p1 p2 -{ - ghost - fn aux - (_foo: unit) - requires (p ** (trade #is (p ** p1) p2)) ** p1 - ensures p2 - opens is - { - elim_trade #is (p ** p1) p2 - }; - intro_trade p1 p2 (p ** (trade #is (p ** p1) p2)) aux -} -``` - -```pulse -ghost -fn trade_elim_partial_r - (#is: inames) - (p1 p p2: vprop) - requires (trade #is (p1 ** p) p2) ** p - ensures trade #is p1 p2 -{ - vprop_equivs (); - rewrite (trade #is (p1 ** p) p2) - as (trade #is (p ** p1) p2); - trade_elim_partial_l p p1 p2 -} -``` - noextract let slice_append_split_precond (#t: Type) (mutb: bool) (p: perm) (v1: Ghost.erased (Seq.seq t)) (i: SZ.t) @@ -327,25 +157,3 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos }} } ``` - -```pulse -ghost -fn rewrite_with_trade - (#[T.exact (`emp_inames)] is:inames) - (p1 p2: vprop) - requires p1 ** pure (p1 == p2) - ensures p2 ** (trade #is p2 p1) -{ - rewrite p1 as p2; - ghost - fn aux - (_: unit) - requires emp ** p2 - ensures p1 - opens is - { - rewrite p2 as p1 - }; - intro_trade _ _ _ aux -} -``` From 65945cb2b1d93d3412c20bebc2645a13cf98da8a Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 18 Jul 2024 22:43:21 -0700 Subject: [PATCH 41/89] LowParse.Pulse: stick -> trade --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 29 +++++----- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 1 - .../pulse/LowParse.Pulse.Combinators.fst | 58 +++++++++---------- src/lowparse/pulse/LowParse.Pulse.Util.fst | 2 +- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 32 +++++----- 5 files changed, 61 insertions(+), 61 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index fb8dd3dd3..cfb543bba 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -5,6 +5,7 @@ open Pulse.Lib.Slice open LowParse.Spec.Base module SZ = FStar.SizeT +module Trade = Pulse.Lib.Trade.Util let parser = tot_parser let serializer #k = tot_serializer #k @@ -17,9 +18,9 @@ ghost fn pts_to_serialized_intro_stick (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v0: bytes) (v: t) requires (pts_to input #pm v0 ** pure (v0 == bare_serialize s v)) - ensures (pts_to_serialized s input #pm v ** (pts_to_serialized s input #pm v @==> pts_to input #pm v0)) + ensures (pts_to_serialized s input #pm v ** (trade (pts_to_serialized s input #pm v) (pts_to input #pm v0))) { - rewrite_with_stick (pts_to input #pm v0) (pts_to_serialized s input #pm v) + Trade.rewrite_with_trade emp_inames (pts_to input #pm v0) (pts_to_serialized s input #pm v) } ``` @@ -28,9 +29,9 @@ ghost fn pts_to_serialized_elim_stick (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: t) requires (pts_to_serialized s input #pm v) - ensures (pts_to input #pm (bare_serialize s v) ** (pts_to input #pm (bare_serialize s v) @==> pts_to_serialized s input #pm v)) + ensures (pts_to input #pm (bare_serialize s v) ** (trade (pts_to input #pm (bare_serialize s v)) (pts_to_serialized s input #pm v))) { - rewrite_with_stick (pts_to_serialized s input #pm v) (pts_to input #pm (bare_serialize s v)) + Trade.rewrite_with_trade emp_inames (pts_to_serialized s input #pm v) (pts_to input #pm (bare_serialize s v)) } ``` @@ -102,7 +103,7 @@ fn pts_to_serialized_ext_stick requires pts_to_serialized s1 input #pm v ** pure ( forall x . parse p1 x == parse p2 x ) - ensures pts_to_serialized s2 input #pm v ** (pts_to_serialized s2 input #pm v @==> pts_to_serialized s1 input #pm v) + ensures pts_to_serialized s2 input #pm v ** trade (pts_to_serialized s2 input #pm v) (pts_to_serialized s1 input #pm v) { pts_to_serialized_ext s1 s2 input; ghost @@ -113,7 +114,7 @@ fn pts_to_serialized_ext_stick { pts_to_serialized_ext s2 s1 input }; - intro_stick _ _ _ aux + intro_trade _ _ _ aux } ``` @@ -393,7 +394,7 @@ let peek_stick_post' (consumed: SZ.t) (left right: slice byte) : Tot vprop -= exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** ((pts_to_serialized s left #pm v1 ** pts_to right #pm v2) @==> pts_to input #pm v) ** pure ( += exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ Seq.length (bare_serialize s v1) == SZ.v consumed /\ parse p v == Some (v1, SZ.v consumed) @@ -450,7 +451,7 @@ fn peek_stick unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); - intro_stick (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_stick_aux s input pm consumed v left right v1 v2 ()); + intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_stick_aux s input pm consumed v left right v1 v2 ()); fold (peek_stick_post' s input pm v consumed left right); fold (peek_stick_post s input pm v consumed (left `SlicePair` right)); (left `SlicePair` right) @@ -469,7 +470,7 @@ fn peek_stick_gen (off: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p offset v (off))) returns input': slice byte - ensures exists* v' . pts_to_serialized s input' #pm v' ** (pts_to_serialized s input' #pm v' @==> pts_to input #pm v) ** pure ( + ensures exists* v' . pts_to_serialized s input' #pm v' ** trade (pts_to_serialized s input' #pm v') (pts_to input #pm v) ** pure ( validator_success #k #t p offset v off /\ parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) @@ -479,15 +480,15 @@ fn peek_stick_gen unfold (slice_split_stick_post input pm v offset split123); unfold (slice_split_stick_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); - stick_elim_partial_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); + Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); let consumed = SZ.sub off offset; let split23 = peek_stick s input23 consumed; match split23 { SlicePair input2 input3 -> { unfold (peek_stick_post s input23 pm v23 consumed split23); unfold (peek_stick_post' s input23 pm v23 consumed input2 input3); with v' . assert (pts_to_serialized s input2 #pm v'); - stick_elim_partial_r (pts_to_serialized s input2 #pm _) (pts_to input3 #pm _) (pts_to input23 #pm v23); - stick_trans (pts_to_serialized s input2 #pm _) (pts_to input23 #pm _) (pts_to input #pm _); + Trade.elim_hyp_r (pts_to_serialized s input2 #pm _) (pts_to input3 #pm _) (pts_to input23 #pm v23); + Trade.trans (pts_to_serialized s input2 #pm _) (pts_to input23 #pm _) (pts_to input #pm _); input2 }} }} @@ -547,7 +548,7 @@ fn read_from_validator_success { let input' = peek_stick_gen s input offset off; let res = r input'; - elim_stick _ _; + Trade.elim _ _; res } ``` @@ -631,7 +632,7 @@ fn reader_ext { pts_to_serialized_ext_stick s2 s1 input; let res = r1 input #pm #v t' f; - elim_stick _ _; + elim_trade _ _; res } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index b08fbf2c1..bab92c49f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -3,7 +3,6 @@ include LowParse.Spec.BitSum include LowParse.Pulse.Combinators open FStar.Tactics.V2 open LowParse.Pulse.Util -open Pulse.Lib.Stick open Pulse.Lib.Slice #push-options "--print_universes" diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 43be563a9..ee25ce1c8 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -3,10 +3,10 @@ include LowParse.Spec.Combinators include LowParse.Pulse.Base open FStar.Tactics.V2 open LowParse.Pulse.Util -open Pulse.Lib.Stick open Pulse.Lib.Slice module SZ = FStar.SizeT +module Trade = Pulse.Lib.Trade.Util inline_for_extraction ```pulse @@ -193,7 +193,7 @@ fn pts_to_serialized_synth_stick (#pm: perm) (#v: t) requires pts_to_serialized s input #pm v - ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) ** (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) @==> pts_to_serialized s input #pm v) + ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) ** trade (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)) (pts_to_serialized s input #pm v) { pts_to_serialized_synth_intro s f f' input; ghost @@ -204,7 +204,7 @@ fn pts_to_serialized_synth_stick { pts_to_serialized_synth_elim s f f' input v }; - intro_stick _ _ _ aux + intro_trade _ _ _ aux } ``` @@ -263,7 +263,7 @@ fn pts_to_serialized_synth_l2r_stick (#pm: perm) (#v: t') requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v - ensures pts_to_serialized s input #pm (f' v) ** (pts_to_serialized s input #pm (f' v) @==> pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) + ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) { pts_to_serialized_synth_l2r s f f' input; ghost @@ -274,7 +274,7 @@ fn pts_to_serialized_synth_l2r_stick { pts_to_serialized_synth_r2l s f f' input v }; - intro_stick _ _ _ aux + intro_trade _ _ _ aux } ``` @@ -294,7 +294,7 @@ fn pts_to_serialized_synth_l2r_stick' (#pm: perm) (#v: t') requires pts_to_serialized s_ input #pm v ** pure (forall x . parse p_ x == parse (tot_parse_synth p f) x) - ensures pts_to_serialized s input #pm (f' v) ** (pts_to_serialized s input #pm (f' v) @==> pts_to_serialized s_ input #pm v) + ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized s_ input #pm v) { pts_to_serialized_ext_stick s_ @@ -302,7 +302,7 @@ fn pts_to_serialized_synth_l2r_stick' input; pts_to_serialized_synth_l2r_stick s f f' input; - stick_trans _ _ (pts_to_serialized s_ input #pm v) + Trade.trans _ _ (pts_to_serialized s_ input #pm v) } ``` @@ -339,7 +339,7 @@ fn read_synth { pts_to_serialized_synth_l2r_stick s1 f2 f1 input; let res = r input #pm #(f1 v) t' (read_synth_cont f2 f2' (f1 v) t' g); - elim_stick _ _; + elim_trade _ _; res } ``` @@ -659,9 +659,9 @@ let split_dtuple2_post' : Tot vprop = pts_to_serialized s1 left #pm (dfst v) ** pts_to_serialized (s2 (dfst v)) right #pm (dsnd v) ** - ((pts_to_serialized s1 left #pm (dfst v) ** - pts_to_serialized (s2 (dfst v)) right #pm (dsnd v)) @==> - pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + trade (pts_to_serialized s1 left #pm (dfst v) ** + pts_to_serialized (s2 (dfst v)) right #pm (dsnd v)) + (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) let split_dtuple2_post (#t1: Type0) @@ -699,7 +699,7 @@ fn split_dtuple2 returns res: slice_pair byte ensures split_dtuple2_post s1 s2 input pm v res { - rewrite_with_stick + Trade.rewrite_with_trade emp_inames (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); @@ -709,11 +709,11 @@ fn split_dtuple2 SlicePair input1 input2 -> { unfold (slice_append_split_stick_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); unfold (slice_append_split_stick_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); - stick_trans (_ ** _) _ _; + Trade.trans (_ ** _) _ _; pts_to_serialized_intro_stick s1 input1 (dfst v); pts_to_serialized_intro_stick (s2 (dfst v)) input2 (dsnd v); - stick_prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); - stick_trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; + Trade.prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); + Trade.trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; fold (split_dtuple2_post' s1 s2 input pm v input1 input2); fold (split_dtuple2_post s1 s2 input pm v (input1 `SlicePair` input2)); (input1 `SlicePair` input2) @@ -740,13 +740,13 @@ fn dtuple2_dfst requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s1 res #pm (dfst v) ** - (pts_to_serialized s1 res #pm (dfst v) @==> pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + trade (pts_to_serialized s1 res #pm (dfst v)) (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; match spl { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v spl); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); - stick_elim_partial_r _ _ _; + Trade.elim_hyp_r _ _ _; input1 }} } @@ -770,13 +770,13 @@ fn dtuple2_dsnd requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized (s2 (dfst v)) res #pm (dsnd v) ** - (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v) @==> pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + trade (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v)) (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; match spl { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v spl); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); - stick_elim_partial_l _ _ _; + Trade.elim_hyp_l _ _ _; input2 }} } @@ -797,9 +797,9 @@ let split_nondep_then_post' : Tot vprop = pts_to_serialized s1 left #pm (fst v) ** pts_to_serialized s2 right #pm (snd v) ** - ((pts_to_serialized s1 left #pm (fst v) ** - pts_to_serialized s2 right #pm (snd v)) @==> - pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + trade (pts_to_serialized s1 left #pm (fst v) ** + pts_to_serialized s2 right #pm (snd v)) + (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) let split_nondep_then_post (#t1 #t2: Type0) @@ -851,12 +851,12 @@ fn split_nondep_then pair_of_dtuple2 dtuple2_of_pair input; - stick_trans (pts_to_serialized (tot_serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; + Trade.trans (pts_to_serialized (tot_serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; let res = split_dtuple2 s1 j1 (fun _ -> s2) input; match res { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 (fun _ -> s2) input pm (dtuple2_of_pair v) res); unfold (split_dtuple2_post' s1 (fun _ -> s2) input pm (dtuple2_of_pair v) input1 input2); - stick_trans (_ ** _) _ _; + Trade.trans (_ ** _) _ _; fold (split_nondep_then_post' s1 s2 input pm v input1 input2); fold (split_nondep_then_post s1 s2 input pm v (input1 `SlicePair` input2)); (input1 `SlicePair` input2) @@ -881,13 +881,13 @@ fn nondep_then_fst requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s1 res #pm (fst v) ** - (pts_to_serialized s1 res #pm (fst v) @==> pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + trade (pts_to_serialized s1 res #pm (fst v)) (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; match spl { SlicePair input1 input2 -> { unfold (split_nondep_then_post s1 s2 input pm v spl); unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); - stick_elim_partial_r _ _ _; + Trade.elim_hyp_r _ _ _; input1 }} } @@ -910,13 +910,13 @@ fn nondep_then_snd requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s2 res #pm (snd v) ** - (pts_to_serialized s2 res #pm (snd v) @==> pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + trade (pts_to_serialized s2 res #pm (snd v)) (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; match spl { SlicePair input1 input2 -> { unfold (split_nondep_then_post s1 s2 input pm v spl); unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); - stick_elim_partial_l _ _ _; + Trade.elim_hyp_l _ _ _; input2 }} } @@ -950,7 +950,7 @@ fn read_dtuple2 unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); let x1 = leaf_reader_of_reader r1 input1; let x2 = leaf_reader_of_reader (r2 x1) input2; - elim_stick _ _; + elim_trade _ _; f (Mkdtuple2 x1 x2) }} } diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 3ecc9804d..2749dc8cc 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -1,6 +1,6 @@ module LowParse.Pulse.Util include Pulse.Lib.Pervasives -open Pulse.Lib.Trade +include Pulse.Lib.Trade module S = Pulse.Lib.Slice module SZ = FStar.SizeT diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 233513b51..ffa6ea303 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -3,11 +3,11 @@ include LowParse.Spec.VCList include LowParse.Pulse.Combinators open FStar.Tactics.V2 open LowParse.Pulse.Util -open Pulse.Lib.Stick open Pulse.Lib.Slice module SZ = FStar.SizeT module R = Pulse.Lib.Reference +module Trade = Pulse.Lib.Trade.Util inline_for_extraction ```pulse @@ -64,14 +64,14 @@ requires pts_to_serialized (tot_serialize_nlist n s) input #pm v ensures exists* v' . pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v' ** - (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + trade (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** pure ( v == (fst v' :: snd v') ) { synth_inverse_1 t (n - 1); synth_inverse_2 t (n - 1); - rewrite_with_stick + Trade.rewrite_with_trade emp_inames (pts_to_serialized (tot_serialize_nlist n s) input #pm v) (pts_to_serialized (tot_serialize_synth _ (synth_nlist (n - 1)) (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); pts_to_serialized_synth_l2r_stick @@ -79,7 +79,7 @@ ensures exists* v' . (synth_nlist (n - 1)) (synth_nlist_recip (n - 1)) input; - stick_trans (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm _) _ _ + Trade.trans (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm _) _ _ } ``` @@ -100,7 +100,7 @@ requires returns input' : slice byte ensures exists* v' . pts_to_serialized s input' #pm v' ** - (pts_to_serialized s input' #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + trade (pts_to_serialized s input' #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** pure ( Cons? v /\ v' == List.Tot.hd v @@ -108,7 +108,7 @@ ensures exists* v' . { nlist_cons_as_nondep_then s n input; let res = nondep_then_fst s j (tot_serialize_nlist (n - 1) s) input; - stick_trans (pts_to_serialized s res #pm _) _ _; + Trade.trans (pts_to_serialized s res #pm _) _ _; res } ``` @@ -130,7 +130,7 @@ requires returns input' : slice byte ensures exists* v' . pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v' ** - (pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v' @==> pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + trade (pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** pure ( Cons? v /\ v' == List.Tot.tl v @@ -138,7 +138,7 @@ ensures exists* v' . { nlist_cons_as_nondep_then s n input; let res = nondep_then_snd s j (tot_serialize_nlist (n - 1) s) input; - stick_trans (pts_to_serialized (tot_serialize_nlist (n - 1) s) res #pm _) _ _; + Trade.trans (pts_to_serialized (tot_serialize_nlist (n - 1) s) res #pm _) _ _; res } ``` @@ -161,10 +161,10 @@ requires returns input' : slice byte ensures exists* v . pts_to_serialized s input' #pm v ** - (pts_to_serialized s input' #pm v @==> pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + trade (pts_to_serialized s input' #pm v) (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** pure (v == List.Tot.index v0 (SZ.v i0)) { - stick_id (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); + Trade.refl emp_inames (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); let mut pi = 0sz; let mut pres = input; while ( @@ -173,7 +173,7 @@ ensures exists* v . ) invariant b . exists* i res (n: nat) (v: nlist n t) . ( R.pts_to pi i ** R.pts_to pres res ** pts_to_serialized (tot_serialize_nlist n s) res #pm v ** - (pts_to_serialized (tot_serialize_nlist n s) res #pm v @==> pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + trade (pts_to_serialized (tot_serialize_nlist n s) res #pm v) (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** pure ( SZ.v i <= SZ.v i0 /\ (b == (SZ.v i < SZ.v i0)) /\ @@ -185,14 +185,14 @@ ensures exists* v . let res2 = nlist_tl s j (SZ.v n0 - SZ.v i) res; pi := (SZ.add i 1sz); pres := res2; - stick_trans + Trade.trans (pts_to_serialized (tot_serialize_nlist (SZ.v n0 - SZ.v i - 1) s) res2 #pm _) _ _ }; let res = !pres; let res2 = nlist_hd s j (SZ.v n0 - SZ.v i0) res; - stick_trans + Trade.trans (pts_to_serialized s res2 #pm _) _ _; res2 } @@ -231,8 +231,8 @@ fn pts_to_serialized_nlist_1 (#v: Ghost.erased t) requires pts_to_serialized s input #pm v ensures exists* v' . pts_to_serialized (tot_serialize_nlist 1 s) input #pm v' ** - (pts_to_serialized (tot_serialize_nlist 1 s) input #pm v' @==> - pts_to_serialized s input #pm v) ** + trade (pts_to_serialized (tot_serialize_nlist 1 s) input #pm v') + (pts_to_serialized s input #pm v) ** pure ((v' <: list t) == [Ghost.reveal v]) { pts_to_serialized_synth_stick s synth_nlist_1 synth_nlist_1_recip input; @@ -241,7 +241,7 @@ fn pts_to_serialized_nlist_1 (tot_serialize_synth p synth_nlist_1 s synth_nlist_1_recip ()) (tot_serialize_nlist 1 s) input; - stick_trans + Trade.trans (pts_to_serialized (tot_serialize_nlist 1 s) input #pm _) _ _ } From e93bb1252986c7d1a9b9286a769384a8cf68ce53 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 18 Jul 2024 22:58:25 -0700 Subject: [PATCH 42/89] LowParse.Pulse: remove now implicit arguments to trade operations --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 4 ++-- src/lowparse/pulse/LowParse.Pulse.Combinators.fst | 2 +- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 4 ++-- 3 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index cfb543bba..99f3f9604 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -20,7 +20,7 @@ fn pts_to_serialized_intro_stick requires (pts_to input #pm v0 ** pure (v0 == bare_serialize s v)) ensures (pts_to_serialized s input #pm v ** (trade (pts_to_serialized s input #pm v) (pts_to input #pm v0))) { - Trade.rewrite_with_trade emp_inames (pts_to input #pm v0) (pts_to_serialized s input #pm v) + Trade.rewrite_with_trade (pts_to input #pm v0) (pts_to_serialized s input #pm v) } ``` @@ -31,7 +31,7 @@ fn pts_to_serialized_elim_stick requires (pts_to_serialized s input #pm v) ensures (pts_to input #pm (bare_serialize s v) ** (trade (pts_to input #pm (bare_serialize s v)) (pts_to_serialized s input #pm v))) { - Trade.rewrite_with_trade emp_inames (pts_to_serialized s input #pm v) (pts_to input #pm (bare_serialize s v)) + Trade.rewrite_with_trade (pts_to_serialized s input #pm v) (pts_to input #pm (bare_serialize s v)) } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index ee25ce1c8..772dd6f7c 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -699,7 +699,7 @@ fn split_dtuple2 returns res: slice_pair byte ensures split_dtuple2_post s1 s2 input pm v res { - Trade.rewrite_with_trade emp_inames + Trade.rewrite_with_trade (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index ffa6ea303..645fb4e70 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -71,7 +71,7 @@ ensures exists* v' . { synth_inverse_1 t (n - 1); synth_inverse_2 t (n - 1); - Trade.rewrite_with_trade emp_inames + Trade.rewrite_with_trade (pts_to_serialized (tot_serialize_nlist n s) input #pm v) (pts_to_serialized (tot_serialize_synth _ (synth_nlist (n - 1)) (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); pts_to_serialized_synth_l2r_stick @@ -164,7 +164,7 @@ ensures exists* v . trade (pts_to_serialized s input' #pm v) (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** pure (v == List.Tot.index v0 (SZ.v i0)) { - Trade.refl emp_inames (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); + Trade.refl (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); let mut pi = 0sz; let mut pres = input; while ( From 709bf3bbc304931c4df13e8f509ee886942a0dfe Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 18 Jul 2024 23:19:23 -0700 Subject: [PATCH 43/89] LowParse.Pulse: stick -> trade in names --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 42 +++++++++---------- .../pulse/LowParse.Pulse.Combinators.fst | 26 ++++++------ src/lowparse/pulse/LowParse.Pulse.Util.fst | 36 ++++++++-------- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 6 +-- 4 files changed, 55 insertions(+), 55 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 99f3f9604..6045f2283 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -15,7 +15,7 @@ let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializ ```pulse ghost -fn pts_to_serialized_intro_stick +fn pts_to_serialized_intro_trade (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v0: bytes) (v: t) requires (pts_to input #pm v0 ** pure (v0 == bare_serialize s v)) ensures (pts_to_serialized s input #pm v ** (trade (pts_to_serialized s input #pm v) (pts_to input #pm v0))) @@ -26,7 +26,7 @@ fn pts_to_serialized_intro_stick ```pulse ghost -fn pts_to_serialized_elim_stick +fn pts_to_serialized_elim_trade (#k: parser_kind) (#t: Type0) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: t) requires (pts_to_serialized s input #pm v) ensures (pts_to input #pm (bare_serialize s v) ** (trade (pts_to input #pm (bare_serialize s v)) (pts_to_serialized s input #pm v))) @@ -89,7 +89,7 @@ fn pts_to_serialized_ext ```pulse ghost -fn pts_to_serialized_ext_stick +fn pts_to_serialized_ext_trade (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) @@ -386,7 +386,7 @@ fn peek } ``` -let peek_stick_post' +let peek_trade_post' (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (pm: perm) @@ -400,7 +400,7 @@ let peek_stick_post' parse p v == Some (v1, SZ.v consumed) ) -let peek_stick_post +let peek_trade_post (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (pm: perm) @@ -409,11 +409,11 @@ let peek_stick_post (res: slice_pair byte) : Tot vprop = let (SlicePair left right) = res in - peek_stick_post' s input pm v consumed left right + peek_trade_post' s input pm v consumed left right ```pulse ghost -fn peek_stick_aux +fn peek_trade_aux (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (pm: perm) @@ -436,7 +436,7 @@ fn peek_stick_aux inline_for_extraction ```pulse -fn peek_stick +fn peek_trade (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) @@ -444,16 +444,16 @@ fn peek_stick (consumed: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) returns res: (slice_pair byte) - ensures peek_stick_post s input pm v consumed res + ensures peek_trade_post s input pm v consumed res { let res = peek s input consumed; match res { SlicePair left right -> { unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); - intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_stick_aux s input pm consumed v left right v1 v2 ()); - fold (peek_stick_post' s input pm v consumed left right); - fold (peek_stick_post s input pm v consumed (left `SlicePair` right)); + intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_trade_aux s input pm consumed v left right v1 v2 ()); + fold (peek_trade_post' s input pm v consumed left right); + fold (peek_trade_post s input pm v consumed (left `SlicePair` right)); (left `SlicePair` right) }} } @@ -461,7 +461,7 @@ fn peek_stick inline_for_extraction ```pulse -fn peek_stick_gen +fn peek_trade_gen (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) @@ -475,17 +475,17 @@ fn peek_stick_gen parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { - let split123 = slice_split_stick false input offset; + let split123 = slice_split_trade false input offset; match split123 { SlicePair input1 input23 -> { - unfold (slice_split_stick_post input pm v offset split123); - unfold (slice_split_stick_post' input pm v offset input1 input23); + unfold (slice_split_trade_post input pm v offset split123); + unfold (slice_split_trade_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); let consumed = SZ.sub off offset; - let split23 = peek_stick s input23 consumed; + let split23 = peek_trade s input23 consumed; match split23 { SlicePair input2 input3 -> { - unfold (peek_stick_post s input23 pm v23 consumed split23); - unfold (peek_stick_post' s input23 pm v23 consumed input2 input3); + unfold (peek_trade_post s input23 pm v23 consumed split23); + unfold (peek_trade_post' s input23 pm v23 consumed input2 input3); with v' . assert (pts_to_serialized s input2 #pm v'); Trade.elim_hyp_r (pts_to_serialized s input2 #pm _) (pts_to input3 #pm _) (pts_to input23 #pm v23); Trade.trans (pts_to_serialized s input2 #pm _) (pts_to input23 #pm _) (pts_to input #pm _); @@ -546,7 +546,7 @@ fn read_from_validator_success parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { - let input' = peek_stick_gen s input offset off; + let input' = peek_trade_gen s input offset off; let res = r input'; Trade.elim _ _; res @@ -630,7 +630,7 @@ fn reader_ext (t': Type0) (f: _) { - pts_to_serialized_ext_stick s2 s1 input; + pts_to_serialized_ext_trade s2 s1 input; let res = r1 input #pm #v t' f; elim_trade _ _; res diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 772dd6f7c..a0cd9ce6f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -182,7 +182,7 @@ fn pts_to_serialized_synth_elim ```pulse ghost -fn pts_to_serialized_synth_stick +fn pts_to_serialized_synth_trade (#t #t': Type0) (#k: parser_kind) (#p: tot_parser k t) @@ -252,7 +252,7 @@ fn pts_to_serialized_synth_r2l ```pulse ghost -fn pts_to_serialized_synth_l2r_stick +fn pts_to_serialized_synth_l2r_trade (#t #t': Type0) (#k: parser_kind) (#p: tot_parser k t) @@ -280,7 +280,7 @@ fn pts_to_serialized_synth_l2r_stick ```pulse ghost -fn pts_to_serialized_synth_l2r_stick' +fn pts_to_serialized_synth_l2r_trade' (#t #t': Type0) (#k_: parser_kind) (#p_: tot_parser k_ t') @@ -296,11 +296,11 @@ fn pts_to_serialized_synth_l2r_stick' requires pts_to_serialized s_ input #pm v ** pure (forall x . parse p_ x == parse (tot_parse_synth p f) x) ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized s_ input #pm v) { - pts_to_serialized_ext_stick + pts_to_serialized_ext_trade s_ (tot_serialize_synth p f s f' ()) input; - pts_to_serialized_synth_l2r_stick + pts_to_serialized_synth_l2r_trade s f f' input; Trade.trans _ _ (pts_to_serialized s_ input #pm v) } @@ -337,7 +337,7 @@ fn read_synth (t': Type0) (g: _) { - pts_to_serialized_synth_l2r_stick s1 f2 f1 input; + pts_to_serialized_synth_l2r_trade s1 f2 f1 input; let res = r input #pm #(f1 v) t' (read_synth_cont f2 f2' (f1 v) t' g); elim_trade _ _; res @@ -704,14 +704,14 @@ fn split_dtuple2 (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); let i = j1 input 0sz; - let res = slice_append_split_stick false input i; + let res = slice_append_split_trade false input i; match res { SlicePair input1 input2 -> { - unfold (slice_append_split_stick_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); - unfold (slice_append_split_stick_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); + unfold (slice_append_split_trade_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); + unfold (slice_append_split_trade_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); Trade.trans (_ ** _) _ _; - pts_to_serialized_intro_stick s1 input1 (dfst v); - pts_to_serialized_intro_stick (s2 (dfst v)) input2 (dsnd v); + pts_to_serialized_intro_trade s1 input1 (dfst v); + pts_to_serialized_intro_trade (s2 (dfst v)) input2 (dsnd v); Trade.prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); Trade.trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; fold (split_dtuple2_post' s1 s2 input pm v input1 input2); @@ -836,7 +836,7 @@ fn split_nondep_then ensures split_nondep_then_post s1 s2 input pm v res { Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); - pts_to_serialized_ext_stick + pts_to_serialized_ext_trade (tot_serialize_nondep_then s1 s2) (tot_serialize_synth (tot_parse_dtuple2 p1 (fun _ -> p2)) @@ -846,7 +846,7 @@ fn split_nondep_then () ) input; - pts_to_serialized_synth_l2r_stick + pts_to_serialized_synth_l2r_trade (tot_serialize_dtuple2 s1 (fun _ -> s2)) pair_of_dtuple2 dtuple2_of_pair diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 2749dc8cc..1f6dc1a2e 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -53,7 +53,7 @@ fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2 } ``` -let slice_append_split_stick_post' +let slice_append_split_trade_post' (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) (s2: S.slice t) @@ -63,16 +63,16 @@ let slice_append_split_stick_post' S.pts_to s2 #p v2 ** (trade (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) (S.pts_to s #p (v1 `Seq.append` v2))) -let slice_append_split_stick_post +let slice_append_split_trade_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (res: S.slice_pair t) : Tot vprop = let S.SlicePair s1 s2 = res in - slice_append_split_stick_post' s p v1 v2 i s1 s2 + slice_append_split_trade_post' s p v1 v2 i s1 s2 ```pulse ghost -fn slice_append_split_stick_aux +fn slice_append_split_trade_aux (#t: Type) (input: S.slice t) (p: perm) (v1 v2: (Seq.seq t)) (i: SZ.t) (input1 input2: S.slice t) (_: unit) requires S.is_split input p i input1 input2 ** (S.pts_to input1 #p v1 ** S.pts_to input2 #p v2) ensures S.pts_to input #p (v1 `Seq.append` v2) @@ -84,26 +84,26 @@ fn slice_append_split_stick_aux inline_for_extraction noextract ```pulse -fn slice_append_split_stick (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) +fn slice_append_split_trade (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to input #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) returns res: S.slice_pair t - ensures slice_append_split_stick_post input p v1 v2 i res + ensures slice_append_split_trade_post input p v1 v2 i res { let res = slice_append_split mutb input i; match res { S.SlicePair input1 input2 -> { unfold (slice_append_split_post input p v1 v2 i res); unfold (slice_append_split_post' input p v1 v2 i input1 input2); - intro_trade _ _ _ (slice_append_split_stick_aux input p v1 v2 i input1 input2); - fold (slice_append_split_stick_post' input p v1 v2 i input1 input2); - fold (slice_append_split_stick_post input p v1 v2 i (S.SlicePair input1 input2)); + intro_trade _ _ _ (slice_append_split_trade_aux input p v1 v2 i input1 input2); + fold (slice_append_split_trade_post' input p v1 v2 i input1 input2); + fold (slice_append_split_trade_post input p v1 v2 i (S.SlicePair input1 input2)); (S.SlicePair input1 input2) } } } ``` -let slice_split_stick_post' +let slice_split_trade_post' (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) (s2: S.slice t) @@ -117,16 +117,16 @@ let slice_split_stick_post' (v1, v2) == Seq.split v (SZ.v i) ) -let slice_split_stick_post +let slice_split_trade_post (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) (res: S.slice_pair t) : Tot vprop = let (S.SlicePair s1 s2) = res in - slice_split_stick_post' s p v i s1 s2 + slice_split_trade_post' s p v i s1 s2 ```pulse ghost -fn slice_split_stick_aux +fn slice_split_trade_aux (#t: Type) (s: S.slice t) (p: perm) (v: Seq.seq t) (i: SZ.t) (s1 s2: S.slice t) (v1 v2: Seq.seq t) (hyp: squash (v == Seq.append v1 v2)) (_: unit) requires (S.is_split s p i s1 s2 ** (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2)) @@ -139,10 +139,10 @@ fn slice_split_stick_aux inline_for_extraction noextract ```pulse -fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) +fn slice_split_trade (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) returns res: S.slice_pair t - ensures slice_split_stick_post s p v i res + ensures slice_split_trade_post s p v i res { Seq.lemma_split v (SZ.v i); let res = S.split mutb s i; @@ -150,9 +150,9 @@ fn slice_split_stick (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghos unfold (S.split_post s p v i res); unfold (S.split_post' s p v i s1 s2); with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); - intro_trade _ _ _ (slice_split_stick_aux s p v i s1 s2 v1 v2 ()); - fold (slice_split_stick_post' s p v i s1 s2); - fold (slice_split_stick_post s p v i (S.SlicePair s1 s2)); + intro_trade _ _ _ (slice_split_trade_aux s p v i s1 s2 v1 v2 ()); + fold (slice_split_trade_post' s p v i s1 s2); + fold (slice_split_trade_post s p v i (S.SlicePair s1 s2)); (S.SlicePair s1 s2) }} } diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 645fb4e70..30c065a5b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -74,7 +74,7 @@ ensures exists* v' . Trade.rewrite_with_trade (pts_to_serialized (tot_serialize_nlist n s) input #pm v) (pts_to_serialized (tot_serialize_synth _ (synth_nlist (n - 1)) (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); - pts_to_serialized_synth_l2r_stick + pts_to_serialized_synth_l2r_trade (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist (n - 1)) (synth_nlist_recip (n - 1)) @@ -235,9 +235,9 @@ fn pts_to_serialized_nlist_1 (pts_to_serialized s input #pm v) ** pure ((v' <: list t) == [Ghost.reveal v]) { - pts_to_serialized_synth_stick s synth_nlist_1 synth_nlist_1_recip input; + pts_to_serialized_synth_trade s synth_nlist_1 synth_nlist_1_recip input; Classical.forall_intro (tot_parse_nlist_1_eq p); - pts_to_serialized_ext_stick + pts_to_serialized_ext_trade (tot_serialize_synth p synth_nlist_1 s synth_nlist_1_recip ()) (tot_serialize_nlist 1 s) input; From e1ff2c4951c45f2ceca72975aa10bb7f0e94a5c5 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 19 Jul 2024 09:42:46 -0700 Subject: [PATCH 44/89] LowParse.Pulse: vprop -> slprop --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 10 +++++----- src/lowparse/pulse/LowParse.Pulse.Combinators.fst | 8 ++++---- src/lowparse/pulse/LowParse.Pulse.Util.fst | 12 ++++++------ 3 files changed, 15 insertions(+), 15 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 6045f2283..24385bdd1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -10,7 +10,7 @@ module Trade = Pulse.Lib.Trade.Util let parser = tot_parser let serializer #k = tot_serializer #k -let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : vprop = +let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : slprop = pts_to input #pm (bare_serialize s v) ```pulse @@ -338,7 +338,7 @@ let peek_post' (v: bytes) (consumed: SZ.t) (left right: slice byte) -: Tot vprop +: Tot slprop = exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** is_split input pm consumed left right ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ Seq.length (bare_serialize s v1) == SZ.v consumed /\ @@ -352,7 +352,7 @@ let peek_post (v: bytes) (consumed: SZ.t) (res: slice_pair byte) -: Tot vprop +: Tot slprop = let SlicePair left right = res in peek_post' s input pm v consumed left right @@ -393,7 +393,7 @@ let peek_trade_post' (v: bytes) (consumed: SZ.t) (left right: slice byte) -: Tot vprop +: Tot slprop = exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ Seq.length (bare_serialize s v1) == SZ.v consumed /\ @@ -407,7 +407,7 @@ let peek_trade_post (v: bytes) (consumed: SZ.t) (res: slice_pair byte) -: Tot vprop +: Tot slprop = let (SlicePair left right) = res in peek_trade_post' s input pm v consumed left right diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index a0cd9ce6f..b70d39abe 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -656,7 +656,7 @@ let split_dtuple2_post' (pm: perm) (v: Ghost.erased (dtuple2 t1 t2)) (left right: slice byte) -: Tot vprop +: Tot slprop = pts_to_serialized s1 left #pm (dfst v) ** pts_to_serialized (s2 (dfst v)) right #pm (dsnd v) ** trade (pts_to_serialized s1 left #pm (dfst v) ** @@ -676,7 +676,7 @@ let split_dtuple2_post (pm: perm) (v: Ghost.erased (dtuple2 t1 t2)) (res: slice_pair byte) -: Tot vprop +: Tot slprop = let (SlicePair left right) = res in split_dtuple2_post' s1 s2 input pm v left right @@ -794,7 +794,7 @@ let split_nondep_then_post' (pm: perm) (v: Ghost.erased (t1 & t2)) (left right: slice byte) -: Tot vprop +: Tot slprop = pts_to_serialized s1 left #pm (fst v) ** pts_to_serialized s2 right #pm (snd v) ** trade (pts_to_serialized s1 left #pm (fst v) ** @@ -813,7 +813,7 @@ let split_nondep_then_post (pm: perm) (v: Ghost.erased (t1 & t2)) (res: slice_pair byte) -: Tot vprop +: Tot slprop = let (SlicePair left right) = res in split_nondep_then_post' s1 s2 input pm v left right diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst index 1f6dc1a2e..9fdf97b15 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Util.fst @@ -16,7 +16,7 @@ let slice_append_split_post' (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) (s2: S.slice t) -: Tot vprop +: Tot slprop = S.pts_to s1 #p v1 ** S.pts_to s2 #p v2 ** @@ -25,7 +25,7 @@ let slice_append_split_post' let slice_append_split_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (res: S.slice_pair t) -: Tot vprop +: Tot slprop = let S.SlicePair s1 s2 = res in slice_append_split_post' s p v1 v2 i s1 s2 @@ -57,7 +57,7 @@ let slice_append_split_trade_post' (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) (s2: S.slice t) -: Tot vprop +: Tot slprop = S.pts_to s1 #p v1 ** S.pts_to s2 #p v2 ** @@ -66,7 +66,7 @@ let slice_append_split_trade_post' let slice_append_split_trade_post (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) (res: S.slice_pair t) -: Tot vprop +: Tot slprop = let S.SlicePair s1 s2 = res in slice_append_split_trade_post' s p v1 v2 i s1 s2 @@ -107,7 +107,7 @@ let slice_split_trade_post' (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) (s1: S.slice t) (s2: S.slice t) -: Tot vprop +: Tot slprop = exists* v1 v2 . S.pts_to s1 #p v1 ** S.pts_to s2 #p v2 ** @@ -120,7 +120,7 @@ let slice_split_trade_post' let slice_split_trade_post (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) (res: S.slice_pair t) -: Tot vprop +: Tot slprop = let (S.SlicePair s1 s2) = res in slice_split_trade_post' s p v i s1 s2 From 1dc81c2a23424a6fa2320fdde31b5c82ffc7569a Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 22 Jul 2024 16:48:19 -0700 Subject: [PATCH 45/89] WIP write --- src/lowparse/pulse/LowParse.Pulse.Write.fst | 34 +++++++++++++++++++++ 1 file changed, 34 insertions(+) create mode 100644 src/lowparse/pulse/LowParse.Pulse.Write.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Write.fst b/src/lowparse/pulse/LowParse.Pulse.Write.fst new file mode 100644 index 000000000..802631eb8 --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.Write.fst @@ -0,0 +1,34 @@ +module LowParse.Pulse.Write +include LowParse.Pulse.Base +open LowParse.Spec.Base +open LowParse.Pulse.Util +open Pulse.Lib.Slice + +module S = Pulse.Lib.Slice +module R = Pulse.Lib.Reference +module SZ = FStar.SizeT + +inline_for_extraction +let l2r_writer + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x': t') -> + (#x: Ghost.erased t) -> + (out: slice byte) -> + (offset: SZ.t) -> + (#v: Ghost.erased bytes) -> + stt SZ.t + (S.pts_to out v ** vmatch x' x ** pure ( + SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v + )) + (fun res -> exists* v' . + S.pts_to out v' ** vmatch x' x ** pure ( + let bs = bare_serialize s x in + SZ.v offset + Seq.length bs <= Seq.length v /\ + Seq.length v' == Seq.length v /\ + Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ + Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + )) From f922d50ad4a10fa8222b3ead86f8277e200e4fde Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 22 Jul 2024 17:36:21 -0700 Subject: [PATCH 46/89] WIP ghost parsers --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 25 ++++++++++------------ 1 file changed, 11 insertions(+), 14 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 24385bdd1..fb9dd78cd 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -7,9 +7,6 @@ open LowParse.Spec.Base module SZ = FStar.SizeT module Trade = Pulse.Lib.Trade.Util -let parser = tot_parser -let serializer #k = tot_serializer #k - let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : slprop = pts_to input #pm (bare_serialize s v) @@ -152,7 +149,7 @@ let validator (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = inline_for_extraction ```pulse fn validate - (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p) + (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (w: validator p) (input: slice byte) (poffset: R.ref SZ.t) (#offset: Ghost.erased SZ.t) @@ -169,7 +166,7 @@ fn validate inline_for_extraction ```pulse fn ifthenelse_validator - (#t: Type0) (#k: parser_kind) (p: parser k t) + (#t: Type0) (#k: Ghost.erased parser_kind) (p: parser k t) (cond: bool) (wtrue: squash (cond == true) -> validator p) (wfalse: squash (cond == false) -> validator p) @@ -198,7 +195,7 @@ let validate_nonempty_post inline_for_extraction ```pulse -fn validate_nonempty (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validator p { k.parser_kind_low > 0 }) +fn validate_nonempty (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (w: validator p { k.parser_kind_low > 0 }) (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -219,16 +216,16 @@ fn validate_nonempty (#t: Type0) (#k: parser_kind) (#p: parser k t) (w: validato ``` inline_for_extraction -let validate_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = +let validate_ext (#t: Type0) (#k1: Ghost.erased parser_kind) (#p1: parser k1 t) (v1: validator p1) (#k2: Ghost.erased parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : validator #_ #k2 p2 = v1 inline_for_extraction -let validate_weaken (#t: Type0) (#k1: parser_kind) (k2: parser_kind) (#p1: parser k1 t) (v1: validator p1 { k2 `is_weaker_than` k1 }) : validator (tot_weaken k2 p1) = - validate_ext v1 (tot_weaken k2 p1) +let validate_weaken (#t: Type0) (#k1: Ghost.erased parser_kind) (k2: Ghost.erased parser_kind) (#p1: parser k1 t) (v1: validator p1 { k2 `is_weaker_than` k1 }) : validator (weaken k2 p1) = + validate_ext v1 (weaken k2 p1) inline_for_extraction ```pulse -fn validate_total_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { +fn validate_total_constant_size (#t: Type0) (#k: Ghost.erased parser_kind) (p: parser k t) (sz: SZ.t { k.parser_kind_high == Some k.parser_kind_low /\ k.parser_kind_low == SZ.v sz /\ k.parser_kind_metadata == Some ParserKindMetadataTotal @@ -285,7 +282,7 @@ let jumper (#t: Type0) (#k: parser_kind) (p: parser k t) : Tot Type = inline_for_extraction ```pulse -fn ifthenelse_jumper (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) (jtrue: squash (cond == true) -> jumper p) (jfalse: squash (cond == false) -> jumper p) +fn ifthenelse_jumper (#t: Type0) (#k: Ghost.erased parser_kind) (p: parser k t) (cond: bool) (jtrue: squash (cond == true) -> jumper p) (jfalse: squash (cond == false) -> jumper p) : jumper #t #k p = (input: slice byte) @@ -303,7 +300,7 @@ fn ifthenelse_jumper (#t: Type0) (#k: parser_kind) (p: parser k t) (cond: bool) inline_for_extraction ```pulse -fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = +fn jump_ext (#t: Type0) (#k1: Ghost.erased parser_kind) (#p1: parser k1 t) (v1: jumper p1) (#k2: Ghost.erased parser_kind) (p2: parser k2 t { forall x . parse p1 x == parse p2 x }) : jumper #_ #k2 p2 = (input: slice byte) (offset: SZ.t) (#pm: perm) @@ -315,7 +312,7 @@ fn jump_ext (#t: Type0) (#k1: parser_kind) (#p1: parser k1 t) (v1: jumper p1) (# inline_for_extraction ```pulse -fn jump_constant_size (#t: Type0) (#k: parser_kind) (p: parser k t) (sz: SZ.t { +fn jump_constant_size (#t: Type0) (#k: Ghost.erased parser_kind) (p: parser k t) (sz: SZ.t { k.parser_kind_high == Some k.parser_kind_low /\ k.parser_kind_low == SZ.v sz }) @@ -359,7 +356,7 @@ let peek_post inline_for_extraction ```pulse fn peek - (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: Ghost.erased bytes) From 50ebfebbcab6c6dcaee981f07105d14602434413 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 22 Jul 2024 21:20:14 -0700 Subject: [PATCH 47/89] convert LowParse.Pulse back to ghost parsers, serializers (WIP impl_pred_recursive) --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 20 +- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 8 +- .../pulse/LowParse.Pulse.Combinators.fst | 327 +++++++++--------- src/lowparse/pulse/LowParse.Pulse.Int.fst | 64 ++-- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 99 +++--- 5 files changed, 256 insertions(+), 262 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index fb9dd78cd..81fa33d70 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -411,7 +411,7 @@ let peek_trade_post ```pulse ghost fn peek_trade_aux - (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (pm: perm) (consumed: SZ.t) @@ -434,7 +434,7 @@ fn peek_trade_aux inline_for_extraction ```pulse fn peek_trade - (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: Ghost.erased bytes) @@ -459,7 +459,7 @@ fn peek_trade inline_for_extraction ```pulse fn peek_trade_gen - (#t: Type0) (#k: parser_kind) (#p: parser k t) (s: serializer p) + (#t: Type0) (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (input: slice byte) (#pm: perm) (#v: Ghost.erased bytes) @@ -510,7 +510,7 @@ let leaf_reader inline_for_extraction let leaf_read (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) @@ -527,7 +527,7 @@ inline_for_extraction ```pulse fn read_from_validator_success (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) @@ -568,7 +568,7 @@ inline_for_extraction ```pulse fn leaf_reader_of_reader (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (r: reader s) @@ -586,7 +586,7 @@ inline_for_extraction ```pulse fn ifthenelse_reader (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (cond: bool) @@ -612,11 +612,11 @@ inline_for_extraction ```pulse fn reader_ext (#t: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t) (#s1: serializer p1) (r1: reader s1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t) (s2: serializer p2 { forall x . parse p1 x == parse p2 x }) :reader #t #k2 #p2 s2 @@ -638,7 +638,7 @@ inline_for_extraction ```pulse fn reader_of_leaf_reader (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (r: leaf_reader s) diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index bab92c49f..bb53e70b6 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -13,10 +13,10 @@ let jump_bitsum' (#tot: pos) (#cl: uint_t tot t) (b: bitsum' cl tot) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (w: jumper p) -: Tot (jumper (tot_parse_bitsum' b p)) +: Tot (jumper (parse_bitsum' b p)) = jump_synth (jump_filter w @@ -38,11 +38,11 @@ let read_bitsum' (#cl: uint_t tot t) (#b: bitsum' cl tot) (d: destr_bitsum'_t b) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (r: reader s) -: Tot (reader (tot_serialize_bitsum' b s)) +: Tot (reader (serialize_bitsum' b s)) = read_synth (read_filter r diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index b70d39abe..85d8afc83 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -13,7 +13,7 @@ inline_for_extraction fn validate_ret (#t: Type0) (x: t) -: validator #t #parse_ret_kind (tot_parse_ret x) +: validator #t #parse_ret_kind (parse_ret x) = (input: slice byte) (poffset: _) (#offset: _) @@ -30,7 +30,7 @@ fn leaf_read_ret (#t: Type0) (x: t) (v_unique: (v' : t) -> Lemma (x == v')) -: leaf_reader #t #parse_ret_kind #(tot_parse_ret x) (tot_serialize_ret x v_unique) +: leaf_reader #t #parse_ret_kind #(parse_ret x) (serialize_ret x v_unique) = (input: slice byte) (#pm: _) (#v: _) @@ -45,29 +45,29 @@ let read_ret (#t: Type0) (x: t) (v_unique: (v' : t) -> Lemma (x == v')) -: Tot (reader (tot_serialize_ret x v_unique)) +: Tot (reader (serialize_ret x v_unique)) = reader_of_leaf_reader (leaf_read_ret x v_unique) inline_for_extraction -let jump_ret (#t: Type) (x: t) : jumper (tot_parse_ret x) = jump_constant_size (tot_parse_ret x) 0sz +let jump_ret (#t: Type) (x: t) : jumper (parse_ret x) = jump_constant_size (parse_ret x) 0sz inline_for_extraction -let validate_empty : validator tot_parse_empty = validate_ret () +let validate_empty : validator parse_empty = validate_ret () inline_for_extraction -let leaf_read_empty : leaf_reader tot_serialize_empty = leaf_read_ret () (fun _ -> ()) +let leaf_read_empty : leaf_reader serialize_empty = leaf_read_ret () (fun _ -> ()) inline_for_extraction -let read_empty : reader tot_serialize_empty = reader_of_leaf_reader leaf_read_empty +let read_empty : reader serialize_empty = reader_of_leaf_reader leaf_read_empty inline_for_extraction -let jump_empty : jumper tot_parse_empty = jump_ret () +let jump_empty : jumper parse_empty = jump_ret () let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) + (#p: parser k t) + (s: serializer p) (v: t) (suff: bytes) : Lemma @@ -79,33 +79,22 @@ let parse_serialize_strong_prefix = let sv = bare_serialize s v in parse_strong_prefix #k p sv (sv `Seq.append` suff) -let tot_parse_synth_eq' - (#k: parser_kind) - (#t1: Type) - (#t2: Type) - (p1: tot_parser k t1) - (f2: (t1 -> Tot t2) {synth_injective f2}) - (b: bytes) -: Lemma - (ensures (parse (tot_parse_synth p1 f2) b == parse_synth' #k p1 f2 b)) -= parse_synth_eq #k p1 f2 b - inline_for_extraction ```pulse fn validate_synth (#t #t': Type) - (#k: parser_kind) - (#p: tot_parser k t) + (#k: Ghost.erased parser_kind) + (#p: parser k t) (w: validator p) - (f: (t -> t') { synth_injective f }) -: validator #t' #k (tot_parse_synth p f) + (f: (t -> GTot t') { synth_injective f }) +: validator #t' #k (parse_synth p f) = (input: slice byte) (poffset: _) (#offset: _) (#pm: _) (#v: _) { - tot_parse_synth_eq' p f (Seq.slice v (SZ.v offset) (Seq.length v)); + parse_synth_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); w input poffset #offset #pm #v } ``` @@ -114,18 +103,18 @@ inline_for_extraction ```pulse fn jump_synth (#t #t': Type0) - (#k: parser_kind) - (#p: tot_parser k t) + (#k: Ghost.erased parser_kind) + (#p: parser k t) (w: jumper p) - (f: (t -> t') { synth_injective f }) -: jumper #t' #k (tot_parse_synth #k #t p f) + (f: (t -> GTot t') { synth_injective f }) +: jumper #t' #k (parse_synth #k #t p f) = (input: _) (offset: _) (#pm: _) (#v: _) { - tot_parse_synth_eq' p f (Seq.slice v (SZ.v offset) (Seq.length v)); + parse_synth_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); w input offset #pm #v } ``` @@ -135,23 +124,23 @@ ghost fn pts_to_serialized_synth_intro (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (#v: t) requires pts_to_serialized s input #pm v - ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) + ensures pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v) { - tot_parse_synth_eq p f (bare_serialize s v); - parse_serialize #k #t' #(tot_parse_synth p f) (tot_serialize_synth p f s f' ()) (f v); - parse_injective #k #t' (tot_parse_synth p f) (bare_serialize s v) (bare_serialize (tot_serialize_synth p f s f' ()) (f v)); + parse_synth_eq p f (bare_serialize s v); + parse_serialize #k #t' #(parse_synth p f) (serialize_synth p f s f' ()) (f v); + parse_injective #k #t' (parse_synth p f) (bare_serialize s v) (bare_serialize (serialize_synth p f s f' ()) (f v)); unfold (pts_to_serialized s input #pm v); rewrite (pts_to input #pm (bare_serialize s v)) - as (pts_to input #pm (bare_serialize (tot_serialize_synth p f s f' ()) (f v))); - fold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)) + as (pts_to input #pm (bare_serialize (serialize_synth p f s f' ()) (f v))); + fold (pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v)) } ``` @@ -160,21 +149,21 @@ ghost fn pts_to_serialized_synth_elim (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (v: t) - requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) + requires pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v) ensures pts_to_serialized s input #pm v { - tot_parse_synth_eq p f (bare_serialize s v); - parse_serialize #k #t' #(tot_parse_synth p f) (tot_serialize_synth p f s f' ()) (f v); - parse_injective #k #t' (tot_parse_synth p f) (bare_serialize s v) (bare_serialize (tot_serialize_synth p f s f' ()) (f v)); - unfold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)); - rewrite (pts_to input #pm (bare_serialize (tot_serialize_synth p f s f' ()) (f v))) + parse_synth_eq p f (bare_serialize s v); + parse_serialize #k #t' #(parse_synth p f) (serialize_synth p f s f' ()) (f v); + parse_injective #k #t' (parse_synth p f) (bare_serialize s v) (bare_serialize (serialize_synth p f s f' ()) (f v)); + unfold (pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v)); + rewrite (pts_to input #pm (bare_serialize (serialize_synth p f s f' ()) (f v))) as (pts_to input #pm (bare_serialize s v)); fold (pts_to_serialized s input #pm v) } @@ -185,21 +174,21 @@ ghost fn pts_to_serialized_synth_trade (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (#v: t) requires pts_to_serialized s input #pm v - ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) ** trade (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v)) (pts_to_serialized s input #pm v) + ensures pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v) ** trade (pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v)) (pts_to_serialized s input #pm v) { pts_to_serialized_synth_intro s f f' input; ghost fn aux (_: unit) - requires emp ** pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm (f v) + requires emp ** pts_to_serialized (serialize_synth p f s f' ()) input #pm (f v) ensures pts_to_serialized s input #pm v { pts_to_serialized_synth_elim s f f' input v @@ -213,18 +202,18 @@ ghost fn pts_to_serialized_synth_l2r (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (#v: t') - requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + requires pts_to_serialized (serialize_synth p f s f' ()) input #pm v ensures pts_to_serialized s input #pm (f' v) { - tot_serialize_synth_eq p f s f' () v; - unfold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v); + serialize_synth_eq p f s f' () v; + unfold (pts_to_serialized (serialize_synth p f s f' ()) input #pm v); fold (pts_to_serialized s input #pm (f' v)) } ``` @@ -234,19 +223,19 @@ ghost fn pts_to_serialized_synth_r2l (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (v: t') requires pts_to_serialized s input #pm (f' v) - ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + ensures pts_to_serialized (serialize_synth p f s f' ()) input #pm v { - tot_serialize_synth_eq p f s f' () v; + serialize_synth_eq p f s f' () v; unfold (pts_to_serialized s input #pm (f' v)); - fold (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) + fold (pts_to_serialized (serialize_synth p f s f' ()) input #pm v) } ``` @@ -255,22 +244,22 @@ ghost fn pts_to_serialized_synth_l2r_trade (#t #t': Type0) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (#v: t') - requires pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v - ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v) + requires pts_to_serialized (serialize_synth p f s f' ()) input #pm v + ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized (serialize_synth p f s f' ()) input #pm v) { pts_to_serialized_synth_l2r s f f' input; ghost fn aux (_: unit) requires emp ** pts_to_serialized s input #pm (f' v) - ensures pts_to_serialized (tot_serialize_synth p f s f' ()) input #pm v + ensures pts_to_serialized (serialize_synth p f s f' ()) input #pm v { pts_to_serialized_synth_r2l s f f' input v }; @@ -283,22 +272,22 @@ ghost fn pts_to_serialized_synth_l2r_trade' (#t #t': Type0) (#k_: parser_kind) - (#p_: tot_parser k_ t') - (#s_: tot_serializer p_) + (#p_: parser k_ t') + (#s_: serializer p_) (#k: parser_kind) - (#p: tot_parser k t) - (s: tot_serializer p) - (f: (t -> t') { synth_injective f }) - (f': (t' -> t) { synth_inverse f f' }) + (#p: parser k t) + (s: serializer p) + (f: (t -> GTot t') { synth_injective f }) + (f': (t' -> GTot t) { synth_inverse f f' }) (input: slice byte) (#pm: perm) (#v: t') - requires pts_to_serialized s_ input #pm v ** pure (forall x . parse p_ x == parse (tot_parse_synth p f) x) + requires pts_to_serialized s_ input #pm v ** pure (forall x . parse p_ x == parse (parse_synth p f) x) ensures pts_to_serialized s input #pm (f' v) ** trade (pts_to_serialized s input #pm (f' v)) (pts_to_serialized s_ input #pm v) { pts_to_serialized_ext_trade s_ - (tot_serialize_synth p f s f' ()) + (serialize_synth p f s f' ()) input; pts_to_serialized_synth_l2r_trade s f f' input; @@ -315,7 +304,7 @@ let read_synth_cont_t inline_for_extraction let read_synth_cont (#t1 #t2: Type0) - (f2: (t1 -> Tot t2)) + (f2: (t1 -> GTot t2)) (f2': ((x: t1) -> read_synth_cont_t (f2 x))) (x1: Ghost.erased t1) (t': Type0) @@ -327,10 +316,10 @@ let read_synth_cont inline_for_extraction ```pulse fn read_synth - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#t2: Type0) (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) (f2': ((x: t1) -> read_synth_cont_t (f2 x))) -: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) +: reader #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) = (input: slice byte) (#pm: _) (#v: _) @@ -353,23 +342,23 @@ let read_synth_cont_init inline_for_extraction let read_synth' - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) - (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) -: reader #t2 #k1 #(tot_parse_synth p1 f2) (tot_serialize_synth p1 f2 s1 f1 ()) + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) + (#t2: Type0) (f2: (t1 -> Tot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) +: reader #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) = read_synth r f2 f1 (fun x -> read_synth_cont_init (f2 x)) inline_for_extraction ```pulse fn validate_filter (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) + (#k: Ghost.erased parser_kind) + (#p: parser k t) (w: validator p) (#s: serializer p) (r: leaf_reader s) - (f: (t -> bool)) + (f: (t -> GTot bool)) (f': (x: t) -> (y: bool { y == f x })) -: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) +: validator #_ #(parse_filter_kind k) (parse_filter p f) = (input: slice byte) (poffset: _) @@ -377,7 +366,7 @@ fn validate_filter (#pm: perm) (#v: Ghost.erased bytes) { - tot_parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); + parse_filter_eq p f (Seq.slice v (SZ.v offset) (Seq.length v)); let offset = !poffset; let is_valid = w input poffset; if is_valid { @@ -401,31 +390,31 @@ let validate_filter'_test inline_for_extraction let validate_filter' (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) + (#k: Ghost.erased parser_kind) + (#p: parser k t) (w: validator p) (#s: serializer p) (r: leaf_reader s) (f: (t -> bool)) -: validator #_ #(parse_filter_kind k) (tot_parse_filter p f) +: validator #_ #(parse_filter_kind k) (parse_filter p f) = validate_filter w r f (validate_filter'_test f) inline_for_extraction ```pulse fn jump_filter (#t: Type0) - (#k: parser_kind) - (#p: tot_parser k t) + (#k: Ghost.erased parser_kind) + (#p: parser k t) (w: jumper p) - (f: (t -> bool)) -: jumper #_ #(parse_filter_kind k) (tot_parse_filter p f) + (f: (t -> GTot bool)) +: jumper #_ #(parse_filter_kind k) (parse_filter p f) = (input: slice byte) (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) { - Classical.forall_intro (tot_parse_filter_eq p f); + Classical.forall_intro (parse_filter_eq p f); w input offset #pm #v } ``` @@ -444,17 +433,17 @@ ghost fn pts_to_serialized_filter_intro (#t: Type0) (#k: parser_kind) - (#p: tot_parser k t) + (#p: parser k t) (s: serializer p) - (f: (t -> bool)) + (f: (t -> GTot bool)) (input: slice byte) (#pm: perm) (#v: t) requires (pts_to_serialized s input #pm v ** pure (f v == true)) - ensures exists* (v': parse_filter_refine f) . pts_to_serialized (tot_serialize_filter s f) input #pm v' ** pure ((v' <: t) == v) + ensures exists* (v': parse_filter_refine f) . pts_to_serialized (serialize_filter s f) input #pm v' ** pure ((v' <: t) == v) { unfold (pts_to_serialized s input #pm v); - fold (pts_to_serialized (tot_serialize_filter s f) input #pm v); + fold (pts_to_serialized (serialize_filter s f) input #pm v); } ``` @@ -463,16 +452,16 @@ ghost fn pts_to_serialized_filter_elim (#t: Type0) (#k: parser_kind) - (#p: tot_parser k t) + (#p: parser k t) (s: serializer p) - (f: (t -> bool)) + (f: (t -> GTot bool)) (input: slice byte) (#pm: perm) (#v: parse_filter_refine f) - requires (pts_to_serialized (tot_serialize_filter s f) input #pm v) + requires (pts_to_serialized (serialize_filter s f) input #pm v) ensures pts_to_serialized s input #pm v { - unfold (pts_to_serialized (tot_serialize_filter s f) input #pm v); + unfold (pts_to_serialized (serialize_filter s f) input #pm v); fold (pts_to_serialized s input #pm v); } ``` @@ -480,7 +469,7 @@ fn pts_to_serialized_filter_elim inline_for_extraction let read_filter_cont (#t: Type0) - (f: t -> bool) + (f: t -> GTot bool) (v: Ghost.erased (parse_filter_refine f)) (t': Type0) (g: (x: parse_filter_refine f { x == Ghost.reveal v }) -> t') @@ -491,8 +480,8 @@ let read_filter_cont inline_for_extraction ```pulse fn read_filter - (#k1: parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (f: (t1 -> bool)) -: reader #(parse_filter_refine f) #(parse_filter_kind k1) #(tot_parse_filter p1 f) (tot_serialize_filter s1 f) + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (r: reader s1) (f: (t1 -> GTot bool)) +: reader #(parse_filter_refine f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) = (input: slice byte) (#pm: _) (#v: _) @@ -523,25 +512,26 @@ let dtuple2_of_pair let nondep_then_eq_dtuple2 (#t1 #t2: Type) (#k1 #k2: parser_kind) - (p1: tot_parser k1 t1) - (p2: tot_parser k2 t2) + (p1: parser k1 t1) + (p2: parser k2 t2) (input: bytes) : Lemma - (parse (tot_nondep_then p1 p2) input == parse (tot_parse_synth (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2) input) -= tot_parse_synth_eq (tot_parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; + (parse (nondep_then p1 p2) input == parse (parse_synth (parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2) input) += parse_synth_eq (parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; + parse_dtuple2_eq p1 (fun _ -> p2) input; nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 input inline_for_extraction ```pulse fn validate_nondep_then (#t1 #t2: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (v1: validator p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t2) (v2: validator p2) -: validator #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) +: validator #(t1 & t2) #(and_then_kind k1 k2) (nondep_then #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) (poffset: _) @@ -564,15 +554,15 @@ inline_for_extraction fn validate_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (v1: validator p1) (#s1: serializer p1) (r1: leaf_reader s1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: ((x: t1) -> parser k2 (t2 x))) (v2: ((x: t1) -> validator (p2 x))) -: validator #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) +: validator #(dtuple2 t1 t2) #(and_then_kind k1 k2) (parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) (poffset: _) @@ -580,6 +570,7 @@ fn validate_dtuple2 (#pm: perm) (#v: Ghost.erased bytes) { + parse_dtuple2_eq p1 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); let offset = !poffset; let is_valid1 = v1 input poffset; if is_valid1 { @@ -596,13 +587,13 @@ inline_for_extraction ```pulse fn jump_nondep_then (#t1 #t2: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (v1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t2) (v2: jumper p2) -: jumper #(t1 & t2) #(and_then_kind k1 k2) (tot_nondep_then #k1 #t1 p1 #k2 #t2 p2) +: jumper #(t1 & t2) #(and_then_kind k1 k2) (nondep_then #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) (offset: SZ.t) @@ -621,21 +612,22 @@ inline_for_extraction fn jump_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (v1: jumper p1) (#s1: serializer p1) (r1: leaf_reader s1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: (x: t1) -> parser k2 (t2 x)) (v2: (x: t1) -> jumper (p2 x)) -: jumper #(dtuple2 t1 t2) #(and_then_kind k1 k2) (tot_parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) +: jumper #(dtuple2 t1 t2) #(and_then_kind k1 k2) (parse_dtuple2 #k1 #t1 p1 #k2 #t2 p2) = (input: slice byte) (offset: SZ.t) (#pm: perm) (#v: Ghost.erased bytes) { + parse_dtuple2_eq p1 p2 (Seq.slice v (SZ.v offset) (Seq.length v)); pts_to_len input; let off1 = v1 input offset; let x = read_from_validator_success r1 input offset off1; @@ -661,7 +653,7 @@ let split_dtuple2_post' pts_to_serialized (s2 (dfst v)) right #pm (dsnd v) ** trade (pts_to_serialized s1 left #pm (dfst v) ** pts_to_serialized (s2 (dfst v)) right #pm (dsnd v)) - (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) let split_dtuple2_post (#t1: Type0) @@ -685,22 +677,23 @@ inline_for_extraction fn split_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: ((x: t1) -> parser k2 (t2 x))) (s2: (x: t1) -> serializer (p2 x)) (input: slice byte) (#pm: perm) (#v: Ghost.erased (dtuple2 t1 t2)) - requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + requires pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v returns res: slice_pair byte ensures split_dtuple2_post s1 s2 input pm v res { + serialize_dtuple2_eq s1 s2 v; Trade.rewrite_with_trade - (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); let i = j1 input 0sz; @@ -727,20 +720,20 @@ inline_for_extraction fn dtuple2_dfst (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: ((x: t1) -> parser k2 (t2 x))) (s2: (x: t1) -> serializer (p2 x)) (input: slice byte) (#pm: perm) (#v: Ghost.erased (dtuple2 t1 t2)) - requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + requires pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s1 res #pm (dfst v) ** - trade (pts_to_serialized s1 res #pm (dfst v)) (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + trade (pts_to_serialized s1 res #pm (dfst v)) (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; match spl { SlicePair input1 input2 -> { @@ -757,20 +750,20 @@ inline_for_extraction fn dtuple2_dsnd (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: ((x: t1) -> parser k2 (t2 x))) (s2: (x: t1) -> serializer (p2 x)) (input: slice byte) (#pm: perm) (#v: Ghost.erased (dtuple2 t1 t2)) - requires pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v + requires pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized (s2 (dfst v)) res #pm (dsnd v) ** - trade (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v)) (pts_to_serialized (tot_serialize_dtuple2 s1 s2) input #pm v) + trade (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v)) (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; match spl { SlicePair input1 input2 -> { @@ -799,7 +792,7 @@ let split_nondep_then_post' pts_to_serialized s2 right #pm (snd v) ** trade (pts_to_serialized s1 left #pm (fst v) ** pts_to_serialized s2 right #pm (snd v)) - (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + (pts_to_serialized (serialize_nondep_then s1 s2) input #pm v) let split_nondep_then_post (#t1 #t2: Type0) @@ -821,37 +814,37 @@ inline_for_extraction ```pulse fn split_nondep_then (#t1 #t2: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t2) (s2: serializer p2) (input: slice byte) (#pm: perm) (#v: Ghost.erased (t1 & t2)) - requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + requires pts_to_serialized (serialize_nondep_then s1 s2) input #pm v returns res: slice_pair byte ensures split_nondep_then_post s1 s2 input pm v res { Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); pts_to_serialized_ext_trade - (tot_serialize_nondep_then s1 s2) - (tot_serialize_synth - (tot_parse_dtuple2 p1 (fun _ -> p2)) + (serialize_nondep_then s1 s2) + (serialize_synth + (parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 - (tot_serialize_dtuple2 s1 (fun _ -> s2)) + (serialize_dtuple2 s1 (fun _ -> s2)) dtuple2_of_pair () ) input; pts_to_serialized_synth_l2r_trade - (tot_serialize_dtuple2 s1 (fun _ -> s2)) + (serialize_dtuple2 s1 (fun _ -> s2)) pair_of_dtuple2 dtuple2_of_pair input; - Trade.trans (pts_to_serialized (tot_serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; + Trade.trans (pts_to_serialized (serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; let res = split_dtuple2 s1 j1 (fun _ -> s2) input; match res { SlicePair input1 input2 -> { unfold (split_dtuple2_post s1 (fun _ -> s2) input pm (dtuple2_of_pair v) res); @@ -868,20 +861,20 @@ inline_for_extraction ```pulse fn nondep_then_fst (#t1 #t2: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t2) (s2: serializer p2) (input: slice byte) (#pm: perm) (#v: Ghost.erased (t1 & t2)) - requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + requires pts_to_serialized (serialize_nondep_then s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s1 res #pm (fst v) ** - trade (pts_to_serialized s1 res #pm (fst v)) (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + trade (pts_to_serialized s1 res #pm (fst v)) (pts_to_serialized (serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; match spl { SlicePair input1 input2 -> { @@ -897,20 +890,20 @@ inline_for_extraction ```pulse fn nondep_then_snd (#t1 #t2: Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (s1: serializer p1 { k1.parser_kind_subkind == Some ParserStrong }) (j1: jumper p1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: parser k2 t2) (s2: serializer p2) (input: slice byte) (#pm: perm) (#v: Ghost.erased (t1 & t2)) - requires pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v + requires pts_to_serialized (serialize_nondep_then s1 s2) input #pm v returns res: slice byte ensures pts_to_serialized s2 res #pm (snd v) ** - trade (pts_to_serialized s2 res #pm (snd v)) (pts_to_serialized (tot_serialize_nondep_then s1 s2) input #pm v) + trade (pts_to_serialized s2 res #pm (snd v)) (pts_to_serialized (serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; match spl { SlicePair input1 input2 -> { @@ -927,16 +920,16 @@ inline_for_extraction fn read_dtuple2 (#t1: Type0) (#t2: t1 -> Type0) - (#k1: parser_kind) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (j1: jumper p1 { k1.parser_kind_subkind == Some ParserStrong }) (#s1: serializer p1) (r1: reader s1) - (#k2: parser_kind) + (#k2: Ghost.erased parser_kind) (#p2: (x: t1) -> parser k2 (t2 x)) (#s2: (x: t1) -> serializer (p2 x)) (r2: (x: t1) -> reader (s2 x)) -: reader #(dtuple2 t1 t2) #(and_then_kind k1 k2) #(tot_parse_dtuple2 p1 p2) (tot_serialize_dtuple2 s1 s2) +: reader #(dtuple2 t1 t2) #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) = (input: slice byte) (#pm: perm) diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 81f6702f5..684143219 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -9,12 +9,12 @@ module SZ = FStar.SizeT module S = Pulse.Lib.Slice inline_for_extraction -let validate_u8 : validator tot_parse_u8 = - validate_total_constant_size tot_parse_u8 1sz +let validate_u8 : validator parse_u8 = + validate_total_constant_size parse_u8 1sz inline_for_extraction -let jump_u8 : jumper tot_parse_u8 = - jump_constant_size tot_parse_u8 1sz +let jump_u8 : jumper parse_u8 = + jump_constant_size parse_u8 1sz inline_for_extraction noextract @@ -23,29 +23,29 @@ let be_to_n_1 = (E.mk_be_to_n EI.uint8 1) inline_for_extraction ```pulse -fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #tot_parse_u8 tot_serialize_u8 +fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #parse_u8 serialize_u8 = (input: S.slice byte) (#pm: perm) (#v: Ghost.erased FStar.UInt8.t) { - unfold (pts_to_serialized tot_serialize_u8 input #pm v); + unfold (pts_to_serialized serialize_u8 input #pm v); serialize_u8_spec_be v; let res = be_to_n_1 input 1sz; - fold (pts_to_serialized tot_serialize_u8 input #pm v); + fold (pts_to_serialized serialize_u8 input #pm v); res } ``` inline_for_extraction -let read_u8 : reader tot_serialize_u8 = reader_of_leaf_reader (read_u8' ()) +let read_u8 : reader serialize_u8 = reader_of_leaf_reader (read_u8' ()) inline_for_extraction -let validate_u16 : validator tot_parse_u16 = - validate_total_constant_size tot_parse_u16 2sz +let validate_u16 : validator parse_u16 = + validate_total_constant_size parse_u16 2sz inline_for_extraction -let jump_u16 : jumper tot_parse_u16 = - jump_constant_size tot_parse_u16 2sz +let jump_u16 : jumper parse_u16 = + jump_constant_size parse_u16 2sz inline_for_extraction noextract @@ -54,29 +54,29 @@ let be_to_n_2 = (E.mk_be_to_n EI.uint16 2) inline_for_extraction ```pulse -fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #tot_parse_u16 tot_serialize_u16 +fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #parse_u16 serialize_u16 = (input: S.slice byte) (#pm: perm) (#v: Ghost.erased FStar.UInt16.t) { - unfold (pts_to_serialized tot_serialize_u16 input #pm v); + unfold (pts_to_serialized serialize_u16 input #pm v); serialize_u16_spec_be v; let res = be_to_n_2 input 2sz; - fold (pts_to_serialized tot_serialize_u16 input #pm v); + fold (pts_to_serialized serialize_u16 input #pm v); res } ``` inline_for_extraction -let read_u16 : reader tot_serialize_u16 = reader_of_leaf_reader (read_u16' ()) +let read_u16 : reader serialize_u16 = reader_of_leaf_reader (read_u16' ()) inline_for_extraction -let validate_u32 : validator tot_parse_u32 = - validate_total_constant_size tot_parse_u32 4sz +let validate_u32 : validator parse_u32 = + validate_total_constant_size parse_u32 4sz inline_for_extraction -let jump_u32 : jumper tot_parse_u32 = - jump_constant_size tot_parse_u32 4sz +let jump_u32 : jumper parse_u32 = + jump_constant_size parse_u32 4sz inline_for_extraction noextract @@ -85,29 +85,29 @@ let be_to_n_4 = (E.mk_be_to_n EI.uint32 4) inline_for_extraction ```pulse -fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #tot_parse_u32 tot_serialize_u32 +fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #parse_u32 serialize_u32 = (input: S.slice byte) (#pm: perm) (#v: Ghost.erased FStar.UInt32.t) { - unfold (pts_to_serialized tot_serialize_u32 input #pm v); + unfold (pts_to_serialized serialize_u32 input #pm v); serialize_u32_spec_be v; let res = be_to_n_4 input 4sz; - fold (pts_to_serialized tot_serialize_u32 input #pm v); + fold (pts_to_serialized serialize_u32 input #pm v); res } ``` inline_for_extraction -let read_u32 : reader tot_serialize_u32 = reader_of_leaf_reader (read_u32' ()) +let read_u32 : reader serialize_u32 = reader_of_leaf_reader (read_u32' ()) inline_for_extraction -let validate_u64 : validator tot_parse_u64 = - validate_total_constant_size tot_parse_u64 8sz +let validate_u64 : validator parse_u64 = + validate_total_constant_size parse_u64 8sz inline_for_extraction -let jump_u64 : jumper tot_parse_u64 = - jump_constant_size tot_parse_u64 8sz +let jump_u64 : jumper parse_u64 = + jump_constant_size parse_u64 8sz inline_for_extraction noextract @@ -116,18 +116,18 @@ let be_to_n_8 = (E.mk_be_to_n EI.uint64 8) inline_for_extraction ```pulse -fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #tot_parse_u64 tot_serialize_u64 +fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #parse_u64 serialize_u64 = (input: S.slice byte) (#pm: perm) (#v: Ghost.erased FStar.UInt64.t) { - unfold (pts_to_serialized tot_serialize_u64 input #pm v); + unfold (pts_to_serialized serialize_u64 input #pm v); serialize_u64_spec_be v; let res = be_to_n_8 input 8sz; - fold (pts_to_serialized tot_serialize_u64 input #pm v); + fold (pts_to_serialized serialize_u64 input #pm v); res } ``` inline_for_extraction -let read_u64 : reader tot_serialize_u64 = reader_of_leaf_reader (read_u64' ()) +let read_u64 : reader serialize_u64 = reader_of_leaf_reader (read_u64' ()) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 30c065a5b..008c96f26 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -13,11 +13,11 @@ inline_for_extraction ```pulse fn jump_nlist (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (j: jumper p) (n0: SZ.t) -: jumper #(nlist (SZ.v n0) t) #(parse_nlist_kind (SZ.v n0) k) (tot_parse_nlist (SZ.v n0) p) +: jumper #(nlist (SZ.v n0) t) #(parse_nlist_kind (SZ.v n0) k) (parse_nlist (SZ.v n0) p) = (input: slice byte) (offset0: SZ.t) @@ -32,15 +32,15 @@ fn jump_nlist ) invariant b . exists* n offset . ( pts_to input #pm v ** R.pts_to pn n ** R.pts_to poffset offset ** pure ( SZ.v offset <= Seq.length v /\ ( - let pr0 = parse_consume (tot_parse_nlist (SZ.v n0) p) (Seq.slice v (SZ.v offset0) (Seq.length v)) in - let pr = parse_consume (tot_parse_nlist (SZ.v n) p) (Seq.slice v (SZ.v offset) (Seq.length v)) in + let pr0 = parse_consume (parse_nlist (SZ.v n0) p) (Seq.slice v (SZ.v offset0) (Seq.length v)) in + let pr = parse_consume (parse_nlist (SZ.v n) p) (Seq.slice v (SZ.v offset) (Seq.length v)) in Some? pr0 /\ Some? pr /\ SZ.v offset0 + Some?.v pr0 == SZ.v offset + Some?.v pr /\ b == (SZ.v n > 0) ))) { let n = !pn; let offset = !poffset; - tot_parse_nlist_eq (SZ.v n) p (Seq.slice v (SZ.v offset) (Seq.length v)); + parse_nlist_eq (SZ.v n) p (Seq.slice v (SZ.v offset) (Seq.length v)); let offset' = j input offset; pn := (SZ.sub n 1sz); poffset := offset'; @@ -61,10 +61,10 @@ fn nlist_cons_as_nondep_then (#pm: perm) (#v: nlist n t) requires - pts_to_serialized (tot_serialize_nlist n s) input #pm v + pts_to_serialized (serialize_nlist n s) input #pm v ensures exists* v' . - pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v' ** - trade (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + pts_to_serialized (serialize_nondep_then s (serialize_nlist (n - 1) s)) input #pm v' ** + trade (pts_to_serialized (serialize_nondep_then s (serialize_nlist (n - 1) s)) input #pm v') (pts_to_serialized (serialize_nlist n s) input #pm v) ** pure ( v == (fst v' :: snd v') ) @@ -72,14 +72,14 @@ ensures exists* v' . synth_inverse_1 t (n - 1); synth_inverse_2 t (n - 1); Trade.rewrite_with_trade - (pts_to_serialized (tot_serialize_nlist n s) input #pm v) - (pts_to_serialized (tot_serialize_synth _ (synth_nlist (n - 1)) (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); + (pts_to_serialized (serialize_nlist n s) input #pm v) + (pts_to_serialized (serialize_synth _ (synth_nlist (n - 1)) (serialize_nondep_then s (serialize_nlist' (n - 1) s)) (synth_nlist_recip (n - 1)) ()) input #pm v); pts_to_serialized_synth_l2r_trade - (tot_serialize_nondep_then s (tot_serialize_nlist' (n - 1) s)) + (serialize_nondep_then s (serialize_nlist' (n - 1) s)) (synth_nlist (n - 1)) (synth_nlist_recip (n - 1)) input; - Trade.trans (pts_to_serialized (tot_serialize_nondep_then s (tot_serialize_nlist (n - 1) s)) input #pm _) _ _ + Trade.trans (pts_to_serialized (serialize_nondep_then s (serialize_nlist (n - 1) s)) input #pm _) _ _ } ``` @@ -87,27 +87,27 @@ inline_for_extraction ```pulse fn nlist_hd (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) (j: jumper p) - (n: pos) + (n: Ghost.erased pos) (input: slice byte) (#pm: perm) (#v: Ghost.erased (nlist n t)) requires - pts_to_serialized (tot_serialize_nlist n s) input #pm v + pts_to_serialized (serialize_nlist n s) input #pm v returns input' : slice byte ensures exists* v' . pts_to_serialized s input' #pm v' ** - trade (pts_to_serialized s input' #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + trade (pts_to_serialized s input' #pm v') (pts_to_serialized (serialize_nlist n s) input #pm v) ** pure ( Cons? v /\ v' == List.Tot.hd v ) { nlist_cons_as_nondep_then s n input; - let res = nondep_then_fst s j (tot_serialize_nlist (n - 1) s) input; + let res = nondep_then_fst #_ #(nlist (n - 1) t) s j #(parse_nlist_kind (n - 1) k) #(coerce_eq () (parse_nlist (n - 1) p)) (coerce_eq () (serialize_nlist (n - 1) s <: serializer #(parse_nlist_kind (n - 1) k) (parse_nlist (n - 1) p))) input; // FIXME: WHY WHY WHY are those reveal (hide (...)) NOT reduced? Trade.trans (pts_to_serialized s res #pm _) _ _; res } @@ -117,28 +117,28 @@ inline_for_extraction ```pulse fn nlist_tl (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) (j: jumper p) - (n: pos) + (n: Ghost.erased pos) (input: slice byte) (#pm: perm) (#v: Ghost.erased (nlist n t)) requires - pts_to_serialized (tot_serialize_nlist n s) input #pm v + pts_to_serialized (serialize_nlist n s) input #pm v returns input' : slice byte ensures exists* v' . - pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v' ** - trade (pts_to_serialized (tot_serialize_nlist (n - 1) s) input' #pm v') (pts_to_serialized (tot_serialize_nlist n s) input #pm v) ** + pts_to_serialized (serialize_nlist (n - 1) s) input' #pm v' ** + trade (pts_to_serialized (serialize_nlist (n - 1) s) input' #pm v') (pts_to_serialized (serialize_nlist n s) input #pm v) ** pure ( Cons? v /\ v' == List.Tot.tl v ) { nlist_cons_as_nondep_then s n input; - let res = nondep_then_snd s j (tot_serialize_nlist (n - 1) s) input; - Trade.trans (pts_to_serialized (tot_serialize_nlist (n - 1) s) res #pm _) _ _; + let res = nondep_then_snd #_ #(nlist (n - 1) t) s j #(parse_nlist_kind (n - 1) k) #(coerce_eq () (parse_nlist (n - 1) p)) (coerce_eq () (serialize_nlist (n - 1) s <: serializer (parse_nlist (n - 1) p))) input; // FIXME: same as above + Trade.trans (pts_to_serialized (serialize_nlist (n - 1) s) res #pm _) _ _; res } ``` @@ -147,24 +147,24 @@ inline_for_extraction ```pulse fn nlist_nth (#t: Type0) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) (j: jumper p) - (n0: SZ.t) + (n0: Ghost.erased nat) (input: slice byte) (#pm: perm) - (#v0: Ghost.erased (nlist (SZ.v n0) t)) - (i0: SZ.t { SZ.v i0 < SZ.v n0 }) + (#v0: Ghost.erased (nlist n0 t)) + (i0: SZ.t { SZ.v i0 < n0 }) requires - pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0 + pts_to_serialized (serialize_nlist n0 s) input #pm v0 returns input' : slice byte ensures exists* v . pts_to_serialized s input' #pm v ** - trade (pts_to_serialized s input' #pm v) (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + trade (pts_to_serialized s input' #pm v) (pts_to_serialized (serialize_nlist n0 s) input #pm v0) ** pure (v == List.Tot.index v0 (SZ.v i0)) { - Trade.refl (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0); + Trade.refl (pts_to_serialized (serialize_nlist n0 s) input #pm v0); let mut pi = 0sz; let mut pres = input; while ( @@ -172,26 +172,26 @@ ensures exists* v . (SZ.lt i i0) ) invariant b . exists* i res (n: nat) (v: nlist n t) . ( R.pts_to pi i ** R.pts_to pres res ** - pts_to_serialized (tot_serialize_nlist n s) res #pm v ** - trade (pts_to_serialized (tot_serialize_nlist n s) res #pm v) (pts_to_serialized (tot_serialize_nlist (SZ.v n0) s) input #pm v0) ** + pts_to_serialized (serialize_nlist n s) res #pm v ** + trade (pts_to_serialized (serialize_nlist n s) res #pm v) (pts_to_serialized (serialize_nlist n0 s) input #pm v0) ** pure ( SZ.v i <= SZ.v i0 /\ (b == (SZ.v i < SZ.v i0)) /\ - n == SZ.v n0 - SZ.v i /\ + n == n0 - SZ.v i /\ List.Tot.index v0 (SZ.v i0) == List.Tot.index v (SZ.v i0 - SZ.v i) )) { let res = !pres; let i = !pi; - let res2 = nlist_tl s j (SZ.v n0 - SZ.v i) res; + let res2 = nlist_tl s j (n0 - SZ.v i) res; pi := (SZ.add i 1sz); pres := res2; Trade.trans - (pts_to_serialized (tot_serialize_nlist (SZ.v n0 - SZ.v i - 1) s) res2 #pm _) + (pts_to_serialized (serialize_nlist (n0 - SZ.v i - 1) s) res2 #pm _) _ _ }; let res = !pres; - let res2 = nlist_hd s j (SZ.v n0 - SZ.v i0) res; + let res2 = nlist_hd s j (n0 - SZ.v i0) res; Trade.trans (pts_to_serialized s res2 #pm _) _ _; res2 @@ -210,17 +210,18 @@ let synth_nlist_1_recip : Tot t = List.Tot.hd x -let tot_parse_nlist_1_eq +let parse_nlist_1_eq (#t: Type) (#k: parser_kind) (p: parser k t) (b: bytes) : Lemma - (parse (tot_parse_nlist 1 p) b == parse (tot_parse_synth p synth_nlist_1) b) -= tot_parse_nlist_eq 1 p b; - tot_parse_synth_eq p synth_nlist_1 b + (parse (parse_nlist 1 p) b == parse (parse_synth p synth_nlist_1) b) += parse_nlist_eq 1 p b; + parse_synth_eq p synth_nlist_1 b ```pulse +ghost fn pts_to_serialized_nlist_1 (#t: Type0) (#k: parser_kind) @@ -228,21 +229,21 @@ fn pts_to_serialized_nlist_1 (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) (input: slice byte) (#pm: perm) - (#v: Ghost.erased t) + (#v: t) requires pts_to_serialized s input #pm v - ensures exists* v' . pts_to_serialized (tot_serialize_nlist 1 s) input #pm v' ** - trade (pts_to_serialized (tot_serialize_nlist 1 s) input #pm v') + ensures exists* v' . pts_to_serialized (serialize_nlist 1 s) input #pm v' ** + trade (pts_to_serialized (serialize_nlist 1 s) input #pm v') (pts_to_serialized s input #pm v) ** - pure ((v' <: list t) == [Ghost.reveal v]) + pure ((v' <: list t) == [v]) { pts_to_serialized_synth_trade s synth_nlist_1 synth_nlist_1_recip input; - Classical.forall_intro (tot_parse_nlist_1_eq p); + Classical.forall_intro (parse_nlist_1_eq p); pts_to_serialized_ext_trade - (tot_serialize_synth p synth_nlist_1 s synth_nlist_1_recip ()) - (tot_serialize_nlist 1 s) + (serialize_synth p synth_nlist_1 s synth_nlist_1_recip ()) + (serialize_nlist 1 s) input; Trade.trans - (pts_to_serialized (tot_serialize_nlist 1 s) input #pm _) + (pts_to_serialized (serialize_nlist 1 s) input #pm _) _ _ } ``` From bd2955d15044b994050b6b4b673bd4a5db7278c6 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 24 Jul 2024 18:50:33 -0700 Subject: [PATCH 48/89] WIP serialize --- src/lowparse/pulse/LowParse.Pulse.Write.fst | 139 ++++++++++++++++++++ 1 file changed, 139 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Write.fst b/src/lowparse/pulse/LowParse.Pulse.Write.fst index 802631eb8..ca0c1cd8d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Write.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Write.fst @@ -32,3 +32,142 @@ let l2r_writer Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs )) + +```pulse +fn l2r_writer_ext + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (w: l2r_writer vmatch s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { forall b . parse p1 b == parse p2 b }) +: l2r_writer #t' #t vmatch #k2 #p2 s2 += (x': t') + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + serializer_unique_strong s1 s2 x; + w x' out offset +} +``` + +(* +inline_for_extraction +let l2r_writer_lens + (#t1' #t2' #t: Type0) + (vmatch1: t1' -> t -> slprop) + (vmatch2: t2' -> t -> slprop) + (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' + (vmatch2 x2' x) + (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) + ) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: l2r_writer vmatch1 s) +: l2r_writer #t2' #t vmatch2 #k #p s += (x2': t2') -> + (#x: Ghost.erased t) -> + (out: slice byte) -> + (offset: SZ.t) -> + (#v: Ghost.erased bytes) -> + stt SZ.t + (S.pts_to out v ** vmatch x2' x ** pure ( + SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v + )) + (fun res -> exists* v' . + S.pts_to out v' ** vmatch2 x2' x ** pure ( + let bs = bare_serialize s x in + SZ.v offset + Seq.length bs <= Seq.length v /\ + Seq.length v' == Seq.length v /\ + Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ + Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + )) +*) + +inline_for_extraction +let compute_remaining_size + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x': t') -> + (#x: Ghost.erased t) -> + (out: R.ref SZ.t) -> + (#v: Ghost.erased SZ.t) -> + stt bool + (R.pts_to out v ** vmatch x' x) + (fun res -> exists* v' . + R.pts_to out v' ** vmatch x' x ** pure ( + let bs = Seq.length (bare_serialize s x) in + (res == true <==> bs <= SZ.v v) /\ + (res == true ==> bs + SZ.v v' == SZ.v v) + )) + + +inline_for_extraction +```pulse +fn compute_size + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (cr: compute_remaining_size vmatch s) + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) + requires + (R.pts_to out v ** vmatch x' x) + returns res: bool + ensures exists* v' . + R.pts_to out v' ** vmatch x' x ** pure ( + let bs = Seq.length (bare_serialize s x) in + (res == true <==> bs <= SZ.v v) /\ + (res == true ==> bs == SZ.v v') + ) +{ + let capacity = !out; + let res = cr x' out; + if res { + let remaining = !out; + out := SZ.sub capacity remaining; + true + } else { + false + } +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_ext + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (cr1: compute_remaining_size vmatch s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { + forall b . parse p1 b == parse p2 b + }) +: compute_remaining_size #t' #t vmatch #k2 #p2 s2 += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + serializer_unique_strong s1 s2 x; + cr1 x' out +} +``` From a723657fa4ebbd65fa51f7a7389d6bf69441c4b3 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 25 Jul 2024 02:20:50 -0700 Subject: [PATCH 49/89] writer, size lenses --- src/lowparse/pulse/LowParse.Pulse.Write.fst | 62 ++++++++++++++------- 1 file changed, 42 insertions(+), 20 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Write.fst b/src/lowparse/pulse/LowParse.Pulse.Write.fst index ca0c1cd8d..9e5966f6b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Write.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Write.fst @@ -56,9 +56,9 @@ fn l2r_writer_ext } ``` -(* inline_for_extraction -let l2r_writer_lens +```pulse +fn l2r_writer_lens (#t1' #t2' #t: Type0) (vmatch1: t1' -> t -> slprop) (vmatch2: t2' -> t -> slprop) @@ -71,24 +71,18 @@ let l2r_writer_lens (#s: serializer p) (w: l2r_writer vmatch1 s) : l2r_writer #t2' #t vmatch2 #k #p s -= (x2': t2') -> - (#x: Ghost.erased t) -> - (out: slice byte) -> - (offset: SZ.t) -> - (#v: Ghost.erased bytes) -> - stt SZ.t - (S.pts_to out v ** vmatch x2' x ** pure ( - SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v - )) - (fun res -> exists* v' . - S.pts_to out v' ** vmatch2 x2' x ** pure ( - let bs = bare_serialize s x in - SZ.v offset + Seq.length bs <= Seq.length v /\ - Seq.length v' == Seq.length v /\ - Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ - Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs - )) -*) += (x2': t2') + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + let x1' = lens x2' _; + let res = w x1' out offset; + elim_trade _ _; + res +} +``` inline_for_extraction let compute_remaining_size @@ -171,3 +165,31 @@ fn compute_remaining_size_ext cr1 x' out } ``` + +inline_for_extraction +```pulse +fn compute_remaining_size_lens + (#t1' #t2' #t: Type0) + (vmatch1: t1' -> t -> slprop) + (vmatch2: t2' -> t -> slprop) + (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' + (vmatch2 x2' x) + (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) + ) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: compute_remaining_size vmatch1 s) +: compute_remaining_size #t2' #t vmatch2 #k #p s += + (x2': t2') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + let x1' = lens x2' _; + let res = w x1' out; + elim_trade _ _; + res +} +``` From 0e7e174d9516cd1716e3c5add19b64c4b5b4b12d Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 25 Jul 2024 17:32:27 -0700 Subject: [PATCH 50/89] compute_remaining_size_constant_size --- src/lowparse/pulse/LowParse.Pulse.Write.fst | 31 ++++++++++++++++++++- 1 file changed, 30 insertions(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Write.fst b/src/lowparse/pulse/LowParse.Pulse.Write.fst index 9e5966f6b..80b21be64 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Write.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Write.fst @@ -104,7 +104,6 @@ let compute_remaining_size (res == true ==> bs + SZ.v v' == SZ.v v) )) - inline_for_extraction ```pulse fn compute_size @@ -193,3 +192,33 @@ fn compute_remaining_size_lens res } ``` + +inline_for_extraction +```pulse +fn compute_remaining_size_constant_size + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sz: SZ.t { + k.parser_kind_high == Some k.parser_kind_low /\ + k.parser_kind_low == SZ.v sz + }) +: compute_remaining_size #t' #t vmatch #k #p s += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + serialize_length s x; + let capacity = !out; + if (SZ.lt capacity sz) { + false + } else { + out := SZ.sub capacity sz; + true + } +} +``` From 69860446bf7dc3aaaff8b60e89f7b4e580f43a13 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 26 Jul 2024 11:00:16 -0700 Subject: [PATCH 51/89] fuse Write into Base --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 216 +++++++++++++++++++ src/lowparse/pulse/LowParse.Pulse.Write.fst | 224 -------------------- 2 files changed, 216 insertions(+), 224 deletions(-) delete mode 100644 src/lowparse/pulse/LowParse.Pulse.Write.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 81fa33d70..a175b12f9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -6,6 +6,7 @@ open LowParse.Spec.Base module SZ = FStar.SizeT module Trade = Pulse.Lib.Trade.Util +module S = Pulse.Lib.Slice let pts_to_serialized (#k: parser_kind) (#t: Type) (#p: parser k t) (s: serializer p) (input: slice byte) (#[exact (`1.0R)] pm: perm) (v: t) : slprop = pts_to input #pm (bare_serialize s v) @@ -654,3 +655,218 @@ fn reader_of_leaf_reader f x } ``` + +inline_for_extraction +let l2r_writer + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x': t') -> + (#x: Ghost.erased t) -> + (out: slice byte) -> + (offset: SZ.t) -> + (#v: Ghost.erased bytes) -> + stt SZ.t + (S.pts_to out v ** vmatch x' x ** pure ( + SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v + )) + (fun res -> exists* v' . + S.pts_to out v' ** vmatch x' x ** pure ( + let bs = bare_serialize s x in + SZ.v offset + Seq.length bs <= Seq.length v /\ + Seq.length v' == Seq.length v /\ + Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ + Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + )) + +```pulse +fn l2r_writer_ext + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (w: l2r_writer vmatch s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { forall b . parse p1 b == parse p2 b }) +: l2r_writer #t' #t vmatch #k2 #p2 s2 += (x': t') + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + serializer_unique_strong s1 s2 x; + w x' out offset +} +``` + +inline_for_extraction +```pulse +fn l2r_writer_lens + (#t1' #t2' #t: Type0) + (vmatch1: t1' -> t -> slprop) + (vmatch2: t2' -> t -> slprop) + (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' + (vmatch2 x2' x) + (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) + ) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: l2r_writer vmatch1 s) +: l2r_writer #t2' #t vmatch2 #k #p s += (x2': t2') + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + let x1' = lens x2' _; + let res = w x1' out offset; + elim_trade _ _; + res +} +``` + +inline_for_extraction +let compute_remaining_size + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x': t') -> + (#x: Ghost.erased t) -> + (out: R.ref SZ.t) -> + (#v: Ghost.erased SZ.t) -> + stt bool + (R.pts_to out v ** vmatch x' x) + (fun res -> exists* v' . + R.pts_to out v' ** vmatch x' x ** pure ( + let bs = Seq.length (bare_serialize s x) in + (res == true <==> bs <= SZ.v v) /\ + (res == true ==> bs + SZ.v v' == SZ.v v) + )) + +inline_for_extraction +```pulse +fn compute_size + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (cr: compute_remaining_size vmatch s) + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) + requires + (R.pts_to out v ** vmatch x' x) + returns res: bool + ensures exists* v' . + R.pts_to out v' ** vmatch x' x ** pure ( + let bs = Seq.length (bare_serialize s x) in + (res == true <==> bs <= SZ.v v) /\ + (res == true ==> bs == SZ.v v') + ) +{ + let capacity = !out; + let res = cr x' out; + if res { + let remaining = !out; + out := SZ.sub capacity remaining; + true + } else { + false + } +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_ext + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (cr1: compute_remaining_size vmatch s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { + forall b . parse p1 b == parse p2 b + }) +: compute_remaining_size #t' #t vmatch #k2 #p2 s2 += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + serializer_unique_strong s1 s2 x; + cr1 x' out +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_lens + (#t1' #t2' #t: Type0) + (vmatch1: t1' -> t -> slprop) + (vmatch2: t2' -> t -> slprop) + (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' + (vmatch2 x2' x) + (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) + ) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: compute_remaining_size vmatch1 s) +: compute_remaining_size #t2' #t vmatch2 #k #p s += + (x2': t2') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + let x1' = lens x2' _; + let res = w x1' out; + elim_trade _ _; + res +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_constant_size + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sz: SZ.t { + k.parser_kind_high == Some k.parser_kind_low /\ + k.parser_kind_low == SZ.v sz + }) +: compute_remaining_size #t' #t vmatch #k #p s += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + serialize_length s x; + let capacity = !out; + if (SZ.lt capacity sz) { + false + } else { + out := SZ.sub capacity sz; + true + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Write.fst b/src/lowparse/pulse/LowParse.Pulse.Write.fst deleted file mode 100644 index 80b21be64..000000000 --- a/src/lowparse/pulse/LowParse.Pulse.Write.fst +++ /dev/null @@ -1,224 +0,0 @@ -module LowParse.Pulse.Write -include LowParse.Pulse.Base -open LowParse.Spec.Base -open LowParse.Pulse.Util -open Pulse.Lib.Slice - -module S = Pulse.Lib.Slice -module R = Pulse.Lib.Reference -module SZ = FStar.SizeT - -inline_for_extraction -let l2r_writer - (#t' #t: Type0) - (vmatch: t' -> t -> slprop) - (#k: parser_kind) - (#p: parser k t) - (s: serializer p) -= (x': t') -> - (#x: Ghost.erased t) -> - (out: slice byte) -> - (offset: SZ.t) -> - (#v: Ghost.erased bytes) -> - stt SZ.t - (S.pts_to out v ** vmatch x' x ** pure ( - SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v - )) - (fun res -> exists* v' . - S.pts_to out v' ** vmatch x' x ** pure ( - let bs = bare_serialize s x in - SZ.v offset + Seq.length bs <= Seq.length v /\ - Seq.length v' == Seq.length v /\ - Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ - Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs - )) - -```pulse -fn l2r_writer_ext - (#t' #t: Type0) - (#vmatch: t' -> t -> slprop) - (#k1: Ghost.erased parser_kind) - (#p1: parser k1 t) - (#s1: serializer p1) - (w: l2r_writer vmatch s1) - (#k2: Ghost.erased parser_kind) - (#p2: parser k2 t) - (s2: serializer p2 { forall b . parse p1 b == parse p2 b }) -: l2r_writer #t' #t vmatch #k2 #p2 s2 -= (x': t') - (#x: Ghost.erased t) - (out: slice byte) - (offset: SZ.t) - (#v: Ghost.erased bytes) -{ - serializer_unique_strong s1 s2 x; - w x' out offset -} -``` - -inline_for_extraction -```pulse -fn l2r_writer_lens - (#t1' #t2' #t: Type0) - (vmatch1: t1' -> t -> slprop) - (vmatch2: t2' -> t -> slprop) - (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' - (vmatch2 x2' x) - (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) - ) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (w: l2r_writer vmatch1 s) -: l2r_writer #t2' #t vmatch2 #k #p s -= (x2': t2') - (#x: Ghost.erased t) - (out: slice byte) - (offset: SZ.t) - (#v: Ghost.erased bytes) -{ - let x1' = lens x2' _; - let res = w x1' out offset; - elim_trade _ _; - res -} -``` - -inline_for_extraction -let compute_remaining_size - (#t' #t: Type0) - (vmatch: t' -> t -> slprop) - (#k: parser_kind) - (#p: parser k t) - (s: serializer p) -= (x': t') -> - (#x: Ghost.erased t) -> - (out: R.ref SZ.t) -> - (#v: Ghost.erased SZ.t) -> - stt bool - (R.pts_to out v ** vmatch x' x) - (fun res -> exists* v' . - R.pts_to out v' ** vmatch x' x ** pure ( - let bs = Seq.length (bare_serialize s x) in - (res == true <==> bs <= SZ.v v) /\ - (res == true ==> bs + SZ.v v' == SZ.v v) - )) - -inline_for_extraction -```pulse -fn compute_size - (#t' #t: Type0) - (#vmatch: t' -> t -> slprop) - (#k: Ghost.erased parser_kind) - (#p: parser k t) - (#s: serializer p) - (cr: compute_remaining_size vmatch s) - (x': t') - (#x: Ghost.erased t) - (out: R.ref SZ.t) - (#v: Ghost.erased SZ.t) - requires - (R.pts_to out v ** vmatch x' x) - returns res: bool - ensures exists* v' . - R.pts_to out v' ** vmatch x' x ** pure ( - let bs = Seq.length (bare_serialize s x) in - (res == true <==> bs <= SZ.v v) /\ - (res == true ==> bs == SZ.v v') - ) -{ - let capacity = !out; - let res = cr x' out; - if res { - let remaining = !out; - out := SZ.sub capacity remaining; - true - } else { - false - } -} -``` - -inline_for_extraction -```pulse -fn compute_remaining_size_ext - (#t' #t: Type0) - (#vmatch: t' -> t -> slprop) - (#k1: Ghost.erased parser_kind) - (#p1: parser k1 t) - (#s1: serializer p1) - (cr1: compute_remaining_size vmatch s1) - (#k2: Ghost.erased parser_kind) - (#p2: parser k2 t) - (s2: serializer p2 { - forall b . parse p1 b == parse p2 b - }) -: compute_remaining_size #t' #t vmatch #k2 #p2 s2 -= - (x': t') - (#x: Ghost.erased t) - (out: R.ref SZ.t) - (#v: Ghost.erased SZ.t) -{ - serializer_unique_strong s1 s2 x; - cr1 x' out -} -``` - -inline_for_extraction -```pulse -fn compute_remaining_size_lens - (#t1' #t2' #t: Type0) - (vmatch1: t1' -> t -> slprop) - (vmatch2: t2' -> t -> slprop) - (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' - (vmatch2 x2' x) - (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) - ) - (#k: parser_kind) - (#p: parser k t) - (#s: serializer p) - (w: compute_remaining_size vmatch1 s) -: compute_remaining_size #t2' #t vmatch2 #k #p s -= - (x2': t2') - (#x: Ghost.erased t) - (out: R.ref SZ.t) - (#v: Ghost.erased SZ.t) -{ - let x1' = lens x2' _; - let res = w x1' out; - elim_trade _ _; - res -} -``` - -inline_for_extraction -```pulse -fn compute_remaining_size_constant_size - (#t' #t: Type0) - (#vmatch: t' -> t -> slprop) - (#k: Ghost.erased parser_kind) - (#p: parser k t) - (s: serializer p) - (sz: SZ.t { - k.parser_kind_high == Some k.parser_kind_low /\ - k.parser_kind_low == SZ.v sz - }) -: compute_remaining_size #t' #t vmatch #k #p s -= - (x': t') - (#x: Ghost.erased t) - (out: R.ref SZ.t) - (#v: Ghost.erased SZ.t) -{ - serialize_length s x; - let capacity = !out; - if (SZ.lt capacity sz) { - false - } else { - out := SZ.sub capacity sz; - true - } -} -``` From 26891a1cb22ab8848eecb9f87b5214a6457332f3 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 26 Jul 2024 13:00:40 -0700 Subject: [PATCH 52/89] compose writing lenses, write ints --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 155 +++++++++++++++--- .../pulse/LowParse.Pulse.Endianness.fst | 25 ++- src/lowparse/pulse/LowParse.Pulse.Int.fst | 100 +++++++++++ 3 files changed, 248 insertions(+), 32 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index a175b12f9..0e3087db5 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -704,34 +704,146 @@ fn l2r_writer_ext } ``` +inline_for_extraction +let vmatch_lens + (#t1' #t2' #t: Type0) + (vmatch1: t1' -> t -> slprop) + (vmatch2: t2' -> t -> slprop) +: Type += (x1': t1') -> (x: Ghost.erased t) -> stt t2' + (vmatch1 x1' x) + (fun x2' -> vmatch2 x2' x ** trade (vmatch2 x2' x) (vmatch1 x1' x)) + +inline_for_extraction +```pulse +fn vmatch_lens_compose + (#t1' #t2' #t3' #t: Type0) + (#vmatch1: t1' -> t -> slprop) + (#vmatch2: t2' -> t -> slprop) + (#vmatch3: t3' -> t -> slprop) + (l12: vmatch_lens vmatch1 vmatch2) + (l23: vmatch_lens vmatch2 vmatch3) +: vmatch_lens #t1' #t3' #t vmatch1 vmatch3 += (x1': t1') + (x: Ghost.erased t) +{ + let x2' = l12 x1' x; + let x3' = l23 x2' x; + Trade.trans _ _ (vmatch1 x1' x); + x3' +} +``` + inline_for_extraction ```pulse fn l2r_writer_lens (#t1' #t2' #t: Type0) - (vmatch1: t1' -> t -> slprop) - (vmatch2: t2' -> t -> slprop) - (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' - (vmatch2 x2' x) - (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) - ) + (#vmatch1: t1' -> t -> slprop) + (#vmatch2: t2' -> t -> slprop) + (lens: vmatch_lens vmatch1 vmatch2) (#k: parser_kind) (#p: parser k t) (#s: serializer p) - (w: l2r_writer vmatch1 s) -: l2r_writer #t2' #t vmatch2 #k #p s -= (x2': t2') + (w: l2r_writer vmatch2 s) +: l2r_writer #t1' #t vmatch1 #k #p s += (x1': t1') (#x: Ghost.erased t) (out: slice byte) (offset: SZ.t) (#v: Ghost.erased bytes) { - let x1' = lens x2' _; - let res = w x1' out offset; + let x2' = lens x1' _; + let res = w x2' out offset; elim_trade _ _; res } ``` +let eq_as_slprop (t: Type) (x x': t) : slprop = pure (x == x') + +let ref_pts_to (t: Type0) (p: perm) (r: ref t) (v: t) : slprop = + R.pts_to r #p v + +```pulse +ghost +fn ref_pts_to_lens_aux + (#t: Type) + (p: perm) + (r: R.ref t) + (v: t) + (x: t) + (_: unit) + requires ref_pts_to t p r v ** eq_as_slprop t x v + ensures ref_pts_to t p r v +{ + unfold (eq_as_slprop t x v) +} +``` + +inline_for_extraction +```pulse +fn ref_pts_to_lens + (t: Type0) + (p: perm) +: vmatch_lens #(ref t) #t #t (ref_pts_to t p) (eq_as_slprop t) += + (r: ref t) + (v: Ghost.erased t) +{ + unfold (ref_pts_to t p r (Ghost.reveal v)); + let x = !r; + fold (ref_pts_to t p r v); + fold (eq_as_slprop t x v); + Trade.intro _ _ _ (ref_pts_to_lens_aux p r v x); + x +} +``` + +inline_for_extraction +let l2r_leaf_writer + (#t: Type) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x: t) -> + (out: slice byte) -> + (offset: SZ.t) -> + (#v: Ghost.erased bytes) -> + stt SZ.t + (S.pts_to out v ** pure ( + SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v + )) + (fun res -> exists* v' . + S.pts_to out v' ** pure ( + let bs = bare_serialize s x in + SZ.v offset + Seq.length bs <= Seq.length v /\ + Seq.length v' == Seq.length v /\ + Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ + Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + )) + +inline_for_extraction +```pulse +fn l2r_writer_of_leaf_writer + (#t: Type) + (#k: parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: l2r_leaf_writer s) +: l2r_writer #t #t (eq_as_slprop t) #k #p s += (x': t) + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + unfold (eq_as_slprop t x' x); + let res = w x' out offset; + fold (eq_as_slprop t x' x); + res +} +``` + inline_for_extraction let compute_remaining_size (#t' #t: Type0) @@ -817,25 +929,22 @@ inline_for_extraction ```pulse fn compute_remaining_size_lens (#t1' #t2' #t: Type0) - (vmatch1: t1' -> t -> slprop) - (vmatch2: t2' -> t -> slprop) - (lens: (x2': t2') -> (x: Ghost.erased t) -> stt t1' - (vmatch2 x2' x) - (fun x1' -> vmatch1 x1' x ** trade (vmatch1 x1' x) (vmatch2 x2' x)) - ) + (#vmatch1: t1' -> t -> slprop) + (#vmatch2: t2' -> t -> slprop) + (lens: vmatch_lens vmatch1 vmatch2) (#k: parser_kind) (#p: parser k t) (#s: serializer p) - (w: compute_remaining_size vmatch1 s) -: compute_remaining_size #t2' #t vmatch2 #k #p s + (w: compute_remaining_size vmatch2 s) +: compute_remaining_size #t1' #t vmatch1 #k #p s = - (x2': t2') + (x1': t1') (#x: Ghost.erased t) (out: R.ref SZ.t) (#v: Ghost.erased SZ.t) { - let x1' = lens x2' _; - let res = w x1' out; + let x2' = lens x1' _; + let res = w x2' out; elim_trade _ _; res } @@ -845,7 +954,7 @@ inline_for_extraction ```pulse fn compute_remaining_size_constant_size (#t' #t: Type0) - (#vmatch: t' -> t -> slprop) + (vmatch: t' -> t -> slprop) (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) diff --git a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst index abde8470a..645dcb6b7 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst @@ -120,7 +120,7 @@ let rec mk_be_to_n then be_to_n_1 u else be_to_n_S (mk_be_to_n u (len - 1)) -// Disclaimer: a function of type n_to_be_t is ultimately meant to be called with a buffer of size len, so we do not care about any output bytes beyond position len. +// Writing from right to left: pos should be the position one past the end of the writing zone inline_for_extraction noextract @@ -136,16 +136,18 @@ let n_to_be_t (pos: SZ.t) -> stt unit (S.pts_to x v ** pure ( - SZ.v pos == len /\ - len <= Seq.length v /\ + len <= SZ.v pos /\ + SZ.v pos <= Seq.length v /\ u.v n < pow2 (8 * len) )) (fun _ -> exists* v' . S.pts_to x v' ** pure ( - SZ.v pos == len /\ - len <= Seq.length v /\ - Seq.length v' == Seq.length v /\ + len <= SZ.v pos /\ + SZ.v pos <= Seq.length v /\ u.v n < pow2 (8 * len) /\ - Seq.slice (v') 0 len `Seq.equal` n_to_be len (u.v n) + Seq.length v' == Seq.length v /\ + Seq.slice v' 0 (SZ.v pos - len) `Seq.equal` Seq.slice (v) 0 (SZ.v pos - len) /\ + Seq.slice v' (SZ.v pos - len) (SZ.v pos) `Seq.equal` n_to_be len (u.v n) /\ + Seq.slice v' (SZ.v pos) (Seq.length v') `Seq.equal` Seq.slice v (SZ.v pos) (Seq.length v) )) inline_for_extraction @@ -182,7 +184,7 @@ fn n_to_be_1 E.reveal_n_to_be 1 (u.v n); E.reveal_n_to_be 0 (u.v n / pow2 8); let n' = u.to_byte n; - S.op_Array_Assignment x 0sz n' + S.op_Array_Assignment x (pos `SZ.sub` 1sz) n' } ``` @@ -201,12 +203,17 @@ fn n_to_be_S (#v: Ghost.erased (Seq.seq U8.t)) (pos: SZ.t) { + Seq.lemma_split (Seq.slice v (SZ.v pos - 1) (Seq.length v)) 1; reveal_n_to_be (len + 1) (u.v n); let lo = u.to_byte n; let hi = u.div256 n; let pos' = pos `SZ.sub` 1sz; + with v1 . assert (S.pts_to x v1); + Seq.lemma_split (Seq.slice v1 (SZ.v pos - 1) (Seq.length v1)) 1; let _ = ih hi x pos'; - S.op_Array_Assignment x pos' lo + S.op_Array_Assignment x pos' lo; + with v2 . assert (S.pts_to x v2); + Seq.lemma_split (Seq.slice v2 (SZ.v pos - 1) (Seq.length v2)) 1; } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 684143219..0f690ab46 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -39,6 +39,31 @@ fn read_u8' (_: unit) : leaf_reader #FStar.UInt8.t #parse_u8_kind #parse_u8 seri inline_for_extraction let read_u8 : reader serialize_u8 = reader_of_leaf_reader (read_u8' ()) +inline_for_extraction +let size_u8 (#t: Type) (vmatch: t -> FStar.UInt8.t -> slprop) : compute_remaining_size vmatch serialize_u8 = + compute_remaining_size_constant_size vmatch serialize_u8 1sz + +inline_for_extraction +noextract +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let n_to_be_1 = (E.mk_n_to_be EI.uint8 1) + +inline_for_extraction +```pulse +fn l2r_leaf_write_u8 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt8.t #parse_u8_kind #parse_u8 serialize_u8 += (n: _) + (x: _) + (pos: SZ.t) + (#v: Ghost.erased (Seq.seq byte)) +{ + S.pts_to_len x; + serialize_u8_spec_be n; + let pos' = SZ.add pos 1sz; + n_to_be_1 n x pos'; + pos' +} +``` + inline_for_extraction let validate_u16 : validator parse_u16 = validate_total_constant_size parse_u16 2sz @@ -70,6 +95,31 @@ fn read_u16' (_: unit) : leaf_reader #FStar.UInt16.t #parse_u16_kind #parse_u16 inline_for_extraction let read_u16 : reader serialize_u16 = reader_of_leaf_reader (read_u16' ()) +inline_for_extraction +let size_u16 (#t: Type) (vmatch: t -> FStar.UInt16.t -> slprop) : compute_remaining_size vmatch serialize_u16 = + compute_remaining_size_constant_size vmatch serialize_u16 2sz + +inline_for_extraction +noextract +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let n_to_be_2 = (E.mk_n_to_be EI.uint16 2) + +inline_for_extraction +```pulse +fn l2r_leaf_write_u16 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt16.t #parse_u16_kind #parse_u16 serialize_u16 += (n: _) + (x: _) + (pos: SZ.t) + (#v: Ghost.erased (Seq.seq byte)) +{ + S.pts_to_len x; + serialize_u16_spec_be n; + let pos' = SZ.add pos 2sz; + n_to_be_2 n x pos'; + pos' +} +``` + inline_for_extraction let validate_u32 : validator parse_u32 = validate_total_constant_size parse_u32 4sz @@ -101,6 +151,31 @@ fn read_u32' (_: unit) : leaf_reader #FStar.UInt32.t #parse_u32_kind #parse_u32 inline_for_extraction let read_u32 : reader serialize_u32 = reader_of_leaf_reader (read_u32' ()) +inline_for_extraction +let size_u32 (#t: Type) (vmatch: t -> FStar.UInt32.t -> slprop) : compute_remaining_size vmatch serialize_u32 = + compute_remaining_size_constant_size vmatch serialize_u32 4sz + +inline_for_extraction +noextract +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let n_to_be_4 = (E.mk_n_to_be EI.uint32 4) + +inline_for_extraction +```pulse +fn l2r_leaf_write_u32 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt32.t #parse_u32_kind #parse_u32 serialize_u32 += (n: _) + (x: _) + (pos: SZ.t) + (#v: Ghost.erased (Seq.seq byte)) +{ + S.pts_to_len x; + serialize_u32_spec_be n; + let pos' = SZ.add pos 4sz; + n_to_be_4 n x pos'; + pos' +} +``` + inline_for_extraction let validate_u64 : validator parse_u64 = validate_total_constant_size parse_u64 8sz @@ -131,3 +206,28 @@ fn read_u64' (_: unit) : leaf_reader #FStar.UInt64.t #parse_u64_kind #parse_u64 inline_for_extraction let read_u64 : reader serialize_u64 = reader_of_leaf_reader (read_u64' ()) + +inline_for_extraction +let size_u64 (#t: Type) (vmatch: t -> FStar.UInt64.t -> slprop) : compute_remaining_size vmatch serialize_u64 = + compute_remaining_size_constant_size vmatch serialize_u64 8sz + +inline_for_extraction +noextract +[@@FStar.Tactics.postprocess_with (fun _ -> FStar.Tactics.norm [delta_attr [`%E.must_reduce]; iota; zeta; primops]; FStar.Tactics.trefl ())] +let n_to_be_8 = (E.mk_n_to_be EI.uint64 8) + +inline_for_extraction +```pulse +fn l2r_leaf_write_u64 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt64.t #parse_u64_kind #parse_u64 serialize_u64 += (n: _) + (x: _) + (pos: SZ.t) + (#v: Ghost.erased (Seq.seq byte)) +{ + S.pts_to_len x; + serialize_u64_spec_be n; + let pos' = SZ.add pos 8sz; + n_to_be_8 n x pos'; + pos' +} +``` From caea2d38126ee817be35cf3c8dcc680827b2a055 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 29 Jul 2024 17:35:30 -0700 Subject: [PATCH 53/89] WIP l2r_write_dtuple2 --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 68 +++++++++- .../pulse/LowParse.Pulse.Combinators.fst | 123 ++++++++++++++++++ 2 files changed, 186 insertions(+), 5 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 0e3087db5..8e574ee54 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -675,10 +675,11 @@ let l2r_writer (fun res -> exists* v' . S.pts_to out v' ** vmatch x' x ** pure ( let bs = bare_serialize s x in - SZ.v offset + Seq.length bs <= Seq.length v /\ + SZ.v res == SZ.v offset + Seq.length bs /\ + SZ.v res <= Seq.length v /\ Seq.length v' == Seq.length v /\ - Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ - Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + Seq.slice v' 0 (SZ.v offset) `Seq.equal` Seq.slice v 0 (SZ.v offset) /\ + Seq.slice v' (SZ.v offset) (SZ.v res) `Seq.equal` bs )) ```pulse @@ -704,6 +705,38 @@ fn l2r_writer_ext } ``` +let vmatch_and_const + (#tl #th: Type) + (const: slprop) + (vmatch: tl -> th -> slprop) + (xl: tl) + (xh: th) +: Tot slprop += const ** vmatch xl xh + +```pulse +fn l2r_writer_frame + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (const: slprop) + (w: l2r_writer vmatch s1) +: l2r_writer #t' #t (vmatch_and_const const vmatch) #k1 #p1 s1 += (x': t') + (#x: Ghost.erased t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + unfold (vmatch_and_const const vmatch); + let res = w x' out offset; + fold (vmatch_and_const const vmatch); + res +} +``` + inline_for_extraction let vmatch_lens (#t1' #t2' #t: Type0) @@ -816,10 +849,11 @@ let l2r_leaf_writer (fun res -> exists* v' . S.pts_to out v' ** pure ( let bs = bare_serialize s x in - SZ.v offset + Seq.length bs <= Seq.length v /\ + SZ.v res == SZ.v offset + Seq.length bs /\ + SZ.v res <= Seq.length v /\ Seq.length v' == Seq.length v /\ Seq.slice v' 0 (SZ.v offset) == Seq.slice v 0 (SZ.v offset) /\ - Seq.slice v' (SZ.v offset) (SZ.v offset + Seq.length bs) == bs + Seq.slice v' (SZ.v offset) (SZ.v res) == bs )) inline_for_extraction @@ -925,6 +959,30 @@ fn compute_remaining_size_ext } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_frame + (#t' #t: Type0) + (#vmatch: t' -> t -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (const: slprop) + (cr1: compute_remaining_size vmatch s1) +: compute_remaining_size #t' #t (vmatch_and_const const vmatch) #k1 #p1 s1 += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + unfold (vmatch_and_const const vmatch); + let res = cr1 x' out; + fold (vmatch_and_const const vmatch); + res +} +``` + inline_for_extraction ```pulse fn compute_remaining_size_lens diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 85d8afc83..815cd8005 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -948,3 +948,126 @@ fn read_dtuple2 }} } ``` + +inline_for_extraction // because Karamel does not like tuple2 +let cpair (t1 t2: Type) = dtuple2 t1 (fun _ -> t2) + +let vmatch_dep_prod + (#tl1 #tl2 #th1: Type) + (#th2: th1 -> Type) + (vmatch1: tl1 -> th1 -> slprop) + (vmatch2: (x: th1) -> tl2 -> th2 x -> slprop) + (xl: (tl1 `cpair` tl2)) + (xh: dtuple2 th1 th2) +: Tot slprop += vmatch1 (dfst xl) (dfst xh) ** vmatch2 (dfst xh) (dsnd xl) (dsnd xh) + +inline_for_extraction +```pulse +fn size_dtuple2 + (#tl1 #tl2 #th1: Type) + (#th2: th1 -> Type) + (#vmatch1: tl1 -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: compute_remaining_size vmatch1 s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#vmatch2: (x: th1) -> tl2 -> th2 x -> slprop) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xl: tl1) -> (xh: Ghost.erased th1) -> compute_remaining_size (vmatch_and_const (vmatch1 xl xh) (vmatch2 xh)) (s2 xh)) +: compute_remaining_size #(tl1 `cpair` tl2) #(dtuple2 th1 th2) (vmatch_dep_prod vmatch1 vmatch2) #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + unfold (vmatch_dep_prod vmatch1 vmatch2); + let res1 = w1 (dfst x') #(dfst x) out; + if res1 { + fold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); + let res2 = w2 (dfst x') (dfst x) (dsnd x') #(dsnd x) out; + unfold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); + fold (vmatch_dep_prod vmatch1 vmatch2); + res2 + } else { + fold (vmatch_dep_prod vmatch1 vmatch2); + false + } +} +``` + +#push-options "--z3rlimit 32" +#restart-solver + +module S = Pulse.Lib.Slice + +inline_for_extraction +```pulse +fn l2r_write_dtuple2 + (#tl1 #tl2 #th1: Type) + (#th2: th1 -> Type) + (#vmatch1: tl1 -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: l2r_writer vmatch1 s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#vmatch2: (x: th1) -> tl2 -> th2 x -> slprop) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xl: tl1) -> (xh: Ghost.erased th1) -> l2r_writer (vmatch_and_const (vmatch1 xl xh) (vmatch2 xh)) (s2 xh)) +: l2r_writer #(tl1 `cpair` tl2) #(dtuple2 th1 th2) (vmatch_dep_prod vmatch1 vmatch2) #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + unfold (vmatch_dep_prod vmatch1 vmatch2); + let res1 = w1 (dfst x') #(dfst x) out offset; + with v1 . assert (S.pts_to out v1); + Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); + fold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); + let res2 = w2 (dfst x') (dfst x) (dsnd x') #(dsnd x) out res1; + with v2 . assert (S.pts_to out v2); + Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) (SZ.v res1 - SZ.v offset) (SZ.v res2 - SZ.v offset); + Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) (SZ.v res1 - SZ.v offset) (SZ.v res2 - SZ.v offset); + unfold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); + fold (vmatch_dep_prod vmatch1 vmatch2); + res2 +} +``` + +#pop-options + +assume val l2r_write_nondep_then + (#tl1 #tl2 #th1 #th2: Type) + (#vmatch1: tl1 -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: l2r_writer vmatch1 s1 { k1.parser_kind_subkind == Some ParserStrong }) + (#vmatch2: tl2 -> th2 -> slprop) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 th2) + (#s2: serializer p2) + (w2: l2r_writer vmatch2 s2) + (#tl: Type) + (vmatch: tl -> (th1 & th2) -> slprop) + (f1: (xl: tl) -> (xh: Ghost.erased (th1 & th2)) -> stt tl1 + (vmatch xl xh) + (fun xl1 -> vmatch1 xl1 (fst xh) ** trade (vmatch1 xl1 (fst xh)) (vmatch xl xh)) + ) + (f2: (xl: tl) -> (xh: Ghost.erased (th1 & th2)) -> stt tl2 + (vmatch xl xh) + (fun xl2 -> vmatch2 xl2 (snd xh) ** trade (vmatch2 xl2 (snd xh)) (vmatch xl xh)) + ) +: l2r_writer vmatch (serialize_nondep_then s1 s2) From 6d47866066cdb99524e6e36969306927de54c7ea Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 29 Jul 2024 19:37:55 -0700 Subject: [PATCH 54/89] write dtuple2, nondep_then --- .../pulse/LowParse.Pulse.Combinators.fst | 87 ++++++++++++++++--- 1 file changed, 76 insertions(+), 11 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 815cd8005..241c97f19 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1000,9 +1000,6 @@ fn size_dtuple2 } ``` -#push-options "--z3rlimit 32" -#restart-solver - module S = Pulse.Lib.Slice inline_for_extraction @@ -1032,35 +1029,81 @@ fn l2r_write_dtuple2 unfold (vmatch_dep_prod vmatch1 vmatch2); let res1 = w1 (dfst x') #(dfst x) out offset; with v1 . assert (S.pts_to out v1); - Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); fold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); let res2 = w2 (dfst x') (dfst x) (dsnd x') #(dsnd x) out res1; with v2 . assert (S.pts_to out v2); + Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); - Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) (SZ.v res1 - SZ.v offset) (SZ.v res2 - SZ.v offset); - Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) (SZ.v res1 - SZ.v offset) (SZ.v res2 - SZ.v offset); + Seq.slice_slice v2 0 (SZ.v res1) 0 (SZ.v offset); unfold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); fold (vmatch_dep_prod vmatch1 vmatch2); res2 } ``` -#pop-options +inline_for_extraction +```pulse +fn size_nondep_then + (#tl1 #tl2 #th1 #th2: Type) + (#vmatch1: tl1 -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: compute_remaining_size vmatch1 s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#vmatch2: tl2 -> th2 -> slprop) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 th2) + (#s2: serializer p2) + (w2: compute_remaining_size vmatch2 s2) + (#tl: Type0) + (vmatch: tl -> (th1 & th2) -> slprop) + (f1: (xl: tl) -> (xh: Ghost.erased (th1 & th2)) -> stt tl1 + (vmatch xl xh) + (fun xl1 -> vmatch1 xl1 (fst xh) ** trade (vmatch1 xl1 (fst xh)) (vmatch xl xh)) + ) + (f2: (xl: tl) -> (xh: Ghost.erased (th1 & th2)) -> stt tl2 + (vmatch xl xh) + (fun xl2 -> vmatch2 xl2 (snd xh) ** trade (vmatch2 xl2 (snd xh)) (vmatch xl xh)) + ) +: compute_remaining_size #tl #(th1 & th2) vmatch #(and_then_kind k1 k2) #(nondep_then p1 p2) (serialize_nondep_then s1 s2) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_nondep_then_eq s1 s2 x; + let x1 = f1 x' x; + let res1 = w1 x1 #(Ghost.hide (fst x)) out #v; + Trade.elim _ _; + if res1 { + let x2 = f2 x' x; + let res2 = w2 x2 #(Ghost.hide (snd x)) out; + Trade.elim _ _; + res2 + } else { + false + } +} +``` -assume val l2r_write_nondep_then +inline_for_extraction +```pulse +fn l2r_write_nondep_then (#tl1 #tl2 #th1 #th2: Type) (#vmatch1: tl1 -> th1 -> slprop) (#k1: Ghost.erased parser_kind) (#p1: parser k1 th1) (#s1: serializer p1) - (w1: l2r_writer vmatch1 s1 { k1.parser_kind_subkind == Some ParserStrong }) + (w1: l2r_writer vmatch1 s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) (#vmatch2: tl2 -> th2 -> slprop) (#k2: Ghost.erased parser_kind) (#p2: parser k2 th2) (#s2: serializer p2) (w2: l2r_writer vmatch2 s2) - (#tl: Type) + (#tl: Type0) (vmatch: tl -> (th1 & th2) -> slprop) (f1: (xl: tl) -> (xh: Ghost.erased (th1 & th2)) -> stt tl1 (vmatch xl xh) @@ -1070,4 +1113,26 @@ assume val l2r_write_nondep_then (vmatch xl xh) (fun xl2 -> vmatch2 xl2 (snd xh) ** trade (vmatch2 xl2 (snd xh)) (vmatch xl xh)) ) -: l2r_writer vmatch (serialize_nondep_then s1 s2) +: l2r_writer #tl #(th1 & th2) vmatch #(and_then_kind k1 k2) #(nondep_then p1 p2) (serialize_nondep_then s1 s2) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_nondep_then_eq s1 s2 x; + let x1 = f1 x' x; + let res1 = w1 x1 #(Ghost.hide (fst x)) out offset; + with v1 . assert (S.pts_to out v1); + Trade.elim _ _; + let x2 = f2 x' x; + let res2 = w2 x2 #(Ghost.hide (snd x)) out res1; + with v2 . assert (S.pts_to out v2); + Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); + Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 0 (SZ.v res1) 0 (SZ.v offset); + Trade.elim _ _; + res2 +} +``` From 6b86b7ab8c12baca3a31ba3afca2489b1586d3f6 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 1 Aug 2024 14:53:12 -0700 Subject: [PATCH 55/89] l2r_write_synth --- .../pulse/LowParse.Pulse.Combinators.fst | 71 +++++++++++++++++++ 1 file changed, 71 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 241c97f19..91d7aa040 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1136,3 +1136,74 @@ fn l2r_write_nondep_then res2 } ``` + +inline_for_extraction +```pulse +fn l2r_leaf_write_synth + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_leaf_writer u#0 s1) + (#t2: Type0) (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) + (f1': ((x2: t2) -> (x1: t1 { x1 == f1 x2 }))) +: l2r_leaf_writer u#0 #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += + (x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + w (f1' x) out offset +} +``` + +let vmatch_synth + (#tl: Type) + (#th1 #th2: Type) + (vmatch: tl -> th1 -> slprop) + (f21: th2 -> GTot th1) + (xh: tl) + (xl2: th2) +: slprop += vmatch xh (f21 xl2) + +inline_for_extraction +```pulse +fn size_synth + (#t: Type0) (#t1: Type0) (#t2: Type0) + (vmatch: t -> t1 -> slprop) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size vmatch s1) + (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) +: compute_remaining_size #t #t2 (vmatch_synth vmatch f1) #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + unfold (vmatch_synth vmatch f1 x' x); + let res = w x' out; + fold (vmatch_synth vmatch f1 x' x); + res +} +``` + +inline_for_extraction +```pulse +fn l2r_write_synth + (#t: Type0) (#t1: Type0) (#t2: Type0) + (vmatch: t -> t1 -> slprop) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_writer vmatch s1) + (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) +: l2r_writer #t #t2 (vmatch_synth vmatch f1) #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + unfold (vmatch_synth vmatch f1 x' x); + let res = w x' out offset; + fold (vmatch_synth vmatch f1 x' x); + res +} +``` From eca27ba1cf197108cc965ec6012e8360f33f371c Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 1 Aug 2024 15:11:18 -0700 Subject: [PATCH 56/89] l2r_write_filter --- .../pulse/LowParse.Pulse.Combinators.fst | 52 +++++++++++++++++++ 1 file changed, 52 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 91d7aa040..2ad6bf695 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1207,3 +1207,55 @@ fn l2r_write_synth res } ``` + +let vmatch_filter + (#tl: Type0) + (#th: Type0) + (vmatch: tl -> th -> slprop) + (f: th -> GTot bool) +: Tot (tl -> parse_filter_refine f -> slprop) += vmatch + + +#set-options "--print_universes --print_implicits" + +inline_for_extraction +```pulse +fn l2r_write_filter + (#t: Type0) (#t1: Type0) + (vmatch: t -> t1 -> slprop) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_writer #t #t1 vmatch s1) + (f: (t1 -> GTot bool)) +: l2r_writer #t #(parse_filter_refine u#0 f) (vmatch_filter vmatch f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + unfold (vmatch_filter vmatch f x' x); + let res = w x' #(Ghost.hide #t1 (Ghost.reveal x)) out offset; + fold (vmatch_filter vmatch f x' x); + res +} +``` + +inline_for_extraction +```pulse +fn size_filter + (#t: Type0) (#t1: Type0) + (vmatch: t -> t1 -> slprop) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size #t #t1 vmatch s1) + (f: (t1 -> GTot bool)) +: compute_remaining_size #t #(parse_filter_refine u#0 f) (vmatch_filter vmatch f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + unfold (vmatch_filter vmatch f x' x); + let res = w x' #(Ghost.hide #t1 (Ghost.reveal x)) out; + fold (vmatch_filter vmatch f x' x); + res +} +``` From 5917d64e3650b02b59aee2a068c2dd21ae868a1e Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 1 Aug 2024 15:21:36 -0700 Subject: [PATCH 57/89] l2r_leaf_writer_ext --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 8e574ee54..32dd563f2 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -856,6 +856,28 @@ let l2r_leaf_writer Seq.slice v' (SZ.v offset) (SZ.v res) == bs )) +inline_for_extraction +```pulse +fn l2r_leaf_writer_ext + (#t: Type0) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (w1: l2r_leaf_writer s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { forall x . parse p1 x == parse p2 x }) +: l2r_leaf_writer u#0 #t #k2 #p2 s2 += (x: t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + serializer_unique_strong s1 s2 x; + w1 x out offset +} +``` + inline_for_extraction ```pulse fn l2r_writer_of_leaf_writer From 05e9b2211de10324647a40e2fb1e209d54356aa5 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 1 Aug 2024 18:46:59 -0700 Subject: [PATCH 58/89] l2r_leaf_write_filter --- .../pulse/LowParse.Pulse.Combinators.fst | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 2ad6bf695..0cdc74095 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1219,6 +1219,22 @@ let vmatch_filter #set-options "--print_universes --print_implicits" +inline_for_extraction +```pulse +fn l2r_leaf_write_filter + (#t1: Type0) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_leaf_writer #t1 s1) + (f: (t1 -> GTot bool)) +: l2r_leaf_writer u#0 #(parse_filter_refine u#0 f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x: _) + (out: _) + (offset: _) + (#v: _) +{ + w x out offset +} +``` + inline_for_extraction ```pulse fn l2r_write_filter From 8ca9313a87f4d263818013a66664e64607c5daa5 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 21 Aug 2024 21:53:37 -0700 Subject: [PATCH 59/89] WHY WHY WHY are lambdas no longer usable in slprops? --- .../pulse/LowParse.Pulse.Combinators.fst | 58 ++++++++++++++----- 1 file changed, 43 insertions(+), 15 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 0cdc74095..770fa072c 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -509,6 +509,8 @@ let dtuple2_of_pair = match x with | (x1, x2) -> (| x1, x2 |) +let const_fun (#t1: Type) (#t2: Type) (x2: t2) (x1: t1) : Tot t2 = x2 + let nondep_then_eq_dtuple2 (#t1 #t2: Type) (#k1 #k2: parser_kind) @@ -516,9 +518,9 @@ let nondep_then_eq_dtuple2 (p2: parser k2 t2) (input: bytes) : Lemma - (parse (nondep_then p1 p2) input == parse (parse_synth (parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2) input) -= parse_synth_eq (parse_dtuple2 p1 (fun _ -> p2)) pair_of_dtuple2 input; - parse_dtuple2_eq p1 (fun _ -> p2) input; + (parse (nondep_then p1 p2) input == parse (parse_synth (parse_dtuple2 p1 #k2 #(const_fun t2) (const_fun p2)) pair_of_dtuple2) input) += parse_synth_eq (parse_dtuple2 p1 #k2 #(const_fun t2) (const_fun p2)) pair_of_dtuple2 input; + parse_dtuple2_eq p1 #k2 #(const_fun t2) (const_fun p2) input; nondep_then_eq #k1 #t1 p1 #k2 #t2 p2 input inline_for_extraction @@ -810,6 +812,32 @@ let split_nondep_then_post = let (SlicePair left right) = res in split_nondep_then_post' s1 s2 input pm v left right +#set-options "--print_implicits" + +```pulse +ghost +fn pts_to_serialized_ext_trade' + (#t: Type0) + (#k1: parser_kind) + (#p1: parser k1 t) + (s1: serializer p1) + (#k2: parser_kind) + (#p2: parser k2 t) + (s2: serializer p2) + (input: slice byte) + (prf: (x: bytes) -> Lemma + (parse p1 x == parse p2 x) + ) + (#pm: perm) + (#v: t) + requires pts_to_serialized s1 input #pm v + ensures pts_to_serialized s2 input #pm v ** trade (pts_to_serialized s2 input #pm v) (pts_to_serialized s1 input #pm v) +{ + Classical.forall_intro prf; + pts_to_serialized_ext_trade s1 s2 input +} +``` + inline_for_extraction ```pulse fn split_nondep_then @@ -828,27 +856,27 @@ fn split_nondep_then returns res: slice_pair byte ensures split_nondep_then_post s1 s2 input pm v res { - Classical.forall_intro (nondep_then_eq_dtuple2 p1 p2); - pts_to_serialized_ext_trade + pts_to_serialized_ext_trade' (serialize_nondep_then s1 s2) - (serialize_synth - (parse_dtuple2 p1 (fun _ -> p2)) - pair_of_dtuple2 - (serialize_dtuple2 s1 (fun _ -> s2)) + (serialize_synth #(and_then_kind k1 k2) #(_: t1 & t2) #(t1 & t2) + (parse_dtuple2 #k1 #t1 p1 #k2 #(const_fun t2) (const_fun p2)) + (pair_of_dtuple2 #t1 #t2) + (serialize_dtuple2 s1 #k2 #(const_fun t2) #(const_fun p2) (const_fun s2)) dtuple2_of_pair () ) - input; + input + (nondep_then_eq_dtuple2 #t1 #t2 #k1 #k2 p1 p2); pts_to_serialized_synth_l2r_trade - (serialize_dtuple2 s1 (fun _ -> s2)) + (serialize_dtuple2 s1 #k2 #(const_fun t2) #(const_fun p2) (const_fun s2)) pair_of_dtuple2 dtuple2_of_pair input; - Trade.trans (pts_to_serialized (serialize_dtuple2 s1 (fun _ -> s2)) _ #pm _) _ _; - let res = split_dtuple2 s1 j1 (fun _ -> s2) input; + Trade.trans (pts_to_serialized (serialize_dtuple2 s1 #k2 #(const_fun t2) #(const_fun p2) (const_fun s2)) _ #pm _) _ _; + let res = split_dtuple2 #t1 #(const_fun t2) s1 j1 #_ #(const_fun p2) (const_fun s2) input; match res { SlicePair input1 input2 -> { - unfold (split_dtuple2_post s1 (fun _ -> s2) input pm (dtuple2_of_pair v) res); - unfold (split_dtuple2_post' s1 (fun _ -> s2) input pm (dtuple2_of_pair v) input1 input2); + unfold (split_dtuple2_post #t1 #(const_fun t2) s1 #k2 #(const_fun p2) (const_fun s2) input pm (dtuple2_of_pair v) res); + unfold (split_dtuple2_post' #t1 #(const_fun t2) s1 #_ #(const_fun p2) (const_fun s2) input pm (dtuple2_of_pair v) input1 input2); Trade.trans (_ ** _) _ _; fold (split_nondep_then_post' s1 s2 input pm v input1 input2); fold (split_nondep_then_post s1 s2 input pm v (input1 `SlicePair` input2)); From 59e59a04d92b3c9de10276ef2bb1d01e279d3f48 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 5 Sep 2024 18:15:14 -0700 Subject: [PATCH 60/89] lowparse: remove LowParse.Pulse.Util --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 8 +- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 2 +- .../pulse/LowParse.Pulse.Combinators.fst | 8 +- .../pulse/LowParse.Pulse.Endianness.fst | 2 +- src/lowparse/pulse/LowParse.Pulse.Int.fst | 2 +- src/lowparse/pulse/LowParse.Pulse.Util.fst | 159 ------------------ src/lowparse/pulse/LowParse.Pulse.VCList.fst | 2 +- 7 files changed, 12 insertions(+), 171 deletions(-) delete mode 100644 src/lowparse/pulse/LowParse.Pulse.Util.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 32dd563f2..bfe041cce 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -1,6 +1,6 @@ module LowParse.Pulse.Base open FStar.Tactics.V2 -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade open Pulse.Lib.Slice open LowParse.Spec.Base @@ -473,10 +473,10 @@ fn peek_trade_gen parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { - let split123 = slice_split_trade false input offset; + let split123 = split_trade false input offset; match split123 { SlicePair input1 input23 -> { - unfold (slice_split_trade_post input pm v offset split123); - unfold (slice_split_trade_post' input pm v offset input1 input23); + unfold (split_trade_post input pm v offset split123); + unfold (split_trade_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); let consumed = SZ.sub off offset; diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index bb53e70b6..e7cb581a6 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -2,7 +2,7 @@ module LowParse.Pulse.BitSum include LowParse.Spec.BitSum include LowParse.Pulse.Combinators open FStar.Tactics.V2 -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade open Pulse.Lib.Slice #push-options "--print_universes" diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 770fa072c..d2aec6f0a 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -2,7 +2,7 @@ module LowParse.Pulse.Combinators include LowParse.Spec.Combinators include LowParse.Pulse.Base open FStar.Tactics.V2 -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade open Pulse.Lib.Slice module SZ = FStar.SizeT @@ -699,11 +699,11 @@ fn split_dtuple2 (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); let i = j1 input 0sz; - let res = slice_append_split_trade false input i; + let res = append_split_trade false input i; match res { SlicePair input1 input2 -> { - unfold (slice_append_split_trade_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); - unfold (slice_append_split_trade_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); + unfold (append_split_trade_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); + unfold (append_split_trade_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); Trade.trans (_ ** _) _ _; pts_to_serialized_intro_trade s1 input1 (dfst v); pts_to_serialized_intro_trade (s2 (dfst v)) input2 (dsnd v); diff --git a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst index 645dcb6b7..fd7bd677a 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst @@ -1,5 +1,5 @@ module LowParse.Pulse.Endianness -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade include LowParse.Spec.Endianness module U8 = FStar.UInt8 diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 0f690ab46..335d3bc94 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -1,7 +1,7 @@ module LowParse.Pulse.Int include LowParse.Spec.Int include LowParse.Pulse.Base -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade module E = LowParse.Pulse.Endianness module EI = LowParse.Spec.Endianness.Instances diff --git a/src/lowparse/pulse/LowParse.Pulse.Util.fst b/src/lowparse/pulse/LowParse.Pulse.Util.fst deleted file mode 100644 index 9fdf97b15..000000000 --- a/src/lowparse/pulse/LowParse.Pulse.Util.fst +++ /dev/null @@ -1,159 +0,0 @@ -module LowParse.Pulse.Util -include Pulse.Lib.Pervasives -include Pulse.Lib.Trade - -module S = Pulse.Lib.Slice -module SZ = FStar.SizeT -module T = FStar.Tactics - -noextract -let slice_append_split_precond - (#t: Type) (mutb: bool) (p: perm) (v1: Ghost.erased (Seq.seq t)) (i: SZ.t) -: Tot prop -= SZ.v i == Seq.length v1 /\ (mutb == true ==> p == 1.0R) - -let slice_append_split_post' - (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (s1: S.slice t) - (s2: S.slice t) -: Tot slprop -= - S.pts_to s1 #p v1 ** - S.pts_to s2 #p v2 ** - S.is_split s p i s1 s2 - -let slice_append_split_post - (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: S.slice_pair t) -: Tot slprop -= let S.SlicePair s1 s2 = res in - slice_append_split_post' s p v1 v2 i s1 s2 - -inline_for_extraction -noextract -```pulse -fn slice_append_split (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - requires S.pts_to s #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: S.slice_pair t - ensures slice_append_split_post s p v1 v2 i res -{ - let vs = Ghost.hide (Seq.split (Seq.append v1 v2) (SZ.v i)); - assert (pure (fst vs `Seq.equal` v1)); - assert (pure (snd vs `Seq.equal` v2)); - let res = S.split mutb s i; - match res { - S.SlicePair s1 s2 -> { - unfold (S.split_post s p (Seq.append v1 v2) i res); - unfold (S.split_post' s p (Seq.append v1 v2) i s1 s2); - fold (slice_append_split_post' s p v1 v2 i s1 s2); - fold (slice_append_split_post s p v1 v2 i (S.SlicePair s1 s2)); - (S.SlicePair s1 s2) - } - } -} -``` - -let slice_append_split_trade_post' - (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (s1: S.slice t) - (s2: S.slice t) -: Tot slprop -= - S.pts_to s1 #p v1 ** - S.pts_to s2 #p v2 ** - (trade (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) (S.pts_to s #p (v1 `Seq.append` v2))) - -let slice_append_split_trade_post - (#t: Type) (s: S.slice t) (p: perm) (v1 v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: S.slice_pair t) -: Tot slprop -= let S.SlicePair s1 s2 = res in - slice_append_split_trade_post' s p v1 v2 i s1 s2 - -```pulse -ghost -fn slice_append_split_trade_aux - (#t: Type) (input: S.slice t) (p: perm) (v1 v2: (Seq.seq t)) (i: SZ.t) (input1 input2: S.slice t) (_: unit) - requires S.is_split input p i input1 input2 ** (S.pts_to input1 #p v1 ** S.pts_to input2 #p v2) - ensures S.pts_to input #p (v1 `Seq.append` v2) -{ - S.join input1 input2 input -} -``` - -inline_for_extraction -noextract -```pulse -fn slice_append_split_trade (#t: Type) (mutb: bool) (input: S.slice t) (#p: perm) (#v1 #v2: Ghost.erased (Seq.seq t)) (i: SZ.t) - requires S.pts_to input #p (v1 `Seq.append` v2) ** pure (slice_append_split_precond mutb p v1 i) - returns res: S.slice_pair t - ensures slice_append_split_trade_post input p v1 v2 i res -{ - let res = slice_append_split mutb input i; - match res { - S.SlicePair input1 input2 -> { - unfold (slice_append_split_post input p v1 v2 i res); - unfold (slice_append_split_post' input p v1 v2 i input1 input2); - intro_trade _ _ _ (slice_append_split_trade_aux input p v1 v2 i input1 input2); - fold (slice_append_split_trade_post' input p v1 v2 i input1 input2); - fold (slice_append_split_trade_post input p v1 v2 i (S.SlicePair input1 input2)); - (S.SlicePair input1 input2) - } - } -} -``` - -let slice_split_trade_post' - (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) - (s1: S.slice t) - (s2: S.slice t) -: Tot slprop -= exists* v1 v2 . - S.pts_to s1 #p v1 ** - S.pts_to s2 #p v2 ** - (trade (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2) (S.pts_to s #p v)) ** - pure ( - SZ.v i <= Seq.length v /\ - (v1, v2) == Seq.split v (SZ.v i) - ) - -let slice_split_trade_post - (#t: Type) (s: S.slice t) (p: perm) (v: Ghost.erased (Seq.seq t)) (i: SZ.t) - (res: S.slice_pair t) -: Tot slprop -= let (S.SlicePair s1 s2) = res in - slice_split_trade_post' s p v i s1 s2 - -```pulse -ghost -fn slice_split_trade_aux - (#t: Type) (s: S.slice t) (p: perm) (v: Seq.seq t) (i: SZ.t) - (s1 s2: S.slice t) (v1 v2: Seq.seq t) (hyp: squash (v == Seq.append v1 v2)) (_: unit) - requires (S.is_split s p i s1 s2 ** (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2)) - ensures (S.pts_to s #p v) - { - S.join s1 s2 s - } -``` - -inline_for_extraction -noextract -```pulse -fn slice_split_trade (#t: Type) (mutb: bool) (s: S.slice t) (#p: perm) (#v: Ghost.erased (Seq.seq t)) (i: SZ.t) - requires S.pts_to s #p v ** pure (S.split_precond mutb p v i) - returns res: S.slice_pair t - ensures slice_split_trade_post s p v i res -{ - Seq.lemma_split v (SZ.v i); - let res = S.split mutb s i; - match res { S.SlicePair s1 s2 -> { - unfold (S.split_post s p v i res); - unfold (S.split_post' s p v i s1 s2); - with v1 v2 . assert (S.pts_to s1 #p v1 ** S.pts_to s2 #p v2); - intro_trade _ _ _ (slice_split_trade_aux s p v i s1 s2 v1 v2 ()); - fold (slice_split_trade_post' s p v i s1 s2); - fold (slice_split_trade_post s p v i (S.SlicePair s1 s2)); - (S.SlicePair s1 s2) - }} -} -``` diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 008c96f26..560c2d15f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -2,7 +2,7 @@ module LowParse.Pulse.VCList include LowParse.Spec.VCList include LowParse.Pulse.Combinators open FStar.Tactics.V2 -open LowParse.Pulse.Util +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade open Pulse.Lib.Slice module SZ = FStar.SizeT From 24f2db9570910df643d3af2091fe5886168e2cbb Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 27 Sep 2024 16:35:55 -0700 Subject: [PATCH 61/89] WIP more l2r_writer combinators --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 18 ++++++-- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 25 +++++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 41 +++++++++++++++++++ 3 files changed, 81 insertions(+), 3 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index bfe041cce..957f83e3d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -657,14 +657,15 @@ fn reader_of_leaf_reader ``` inline_for_extraction -let l2r_writer +let l2r_writer_for (#t' #t: Type0) (vmatch: t' -> t -> slprop) (#k: parser_kind) (#p: parser k t) (s: serializer p) -= (x': t') -> - (#x: Ghost.erased t) -> + (x': t') + (x: Ghost.erased t) += (out: slice byte) -> (offset: SZ.t) -> (#v: Ghost.erased bytes) -> @@ -682,6 +683,17 @@ let l2r_writer Seq.slice v' (SZ.v offset) (SZ.v res) `Seq.equal` bs )) +inline_for_extraction +let l2r_writer + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x': t') -> + (#x: Ghost.erased t) -> + l2r_writer_for vmatch s x' x + ```pulse fn l2r_writer_ext (#t' #t: Type0) diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index e7cb581a6..6e33df033 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -56,3 +56,28 @@ let read_bitsum' (read_synth_cont_ifthenelse #(bitsum'_type b)) (read_synth_cont_init) ) + +inline_for_extraction +```pulse +fn l2r_write_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (#b: bitsum' cl tot) + (sr: synth_bitsum'_recip_t b) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: l2r_leaf_writer s) +: (l2r_leaf_writer u#0 #(bitsum'_type b) #(parse_filter_kind k) #(parse_bitsum' b p) (serialize_bitsum' b s)) += (x: _) + (out: _) + (offset: _) + (#v: Ghost.erased bytes) +{ + serialize_bitsum'_eq b s x; + synth_bitsum'_injective b; + synth_bitsum'_recip_inverse b; + w (sr x) out offset +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index d2aec6f0a..8f0bdca2b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1183,6 +1183,21 @@ fn l2r_leaf_write_synth } ``` +inline_for_extraction +let mk_synth + (#t1 #t2: Type) + (f: (t1 -> Tot t2)) + (x: t1) +: Tot (y: t2 { y == f x }) += f x + +inline_for_extraction +let l2r_leaf_write_synth' + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_leaf_writer u#0 s1) + (#t2: Type0) (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) +: l2r_leaf_writer u#0 #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += l2r_leaf_write_synth w f2 f1 (mk_synth f1) + let vmatch_synth (#tl: Type) (#th1 #th2: Type) @@ -1236,6 +1251,32 @@ fn l2r_write_synth } ``` +inline_for_extraction +```pulse +fn l2r_write_synth2 + (#t: Type0) (#t1: Type0) (#t2: Type0) + (vmatch: t -> t2 -> slprop) + (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_writer (vmatch_synth vmatch f2) s1) +: l2r_writer #t #t2 vmatch #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (f2 (f1 x))); + fold (vmatch_synth vmatch f2 x' (f1 x)); + let res = w x' out offset; + unfold (vmatch_synth vmatch f2 x' (f1 x)); + Trade.elim _ _; + res +} +``` + let vmatch_filter (#tl: Type0) (#th: Type0) From 88251c39fc62eadd4b2eb610f08b967c7f776b71 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 30 Sep 2024 17:35:44 -0700 Subject: [PATCH 62/89] serialization combinators with reverse handling of vmatch --- .../pulse/LowParse.Pulse.Combinators.fst | 142 +++++++++++++++++- 1 file changed, 140 insertions(+), 2 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 8f0bdca2b..b0f11884a 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -1070,6 +1070,110 @@ fn l2r_write_dtuple2 } ``` +let vmatch_dep_proj2 + (#tl #th1: Type) + (#th2: th1 -> Type) + (vmatch: tl -> dtuple2 th1 th2 -> slprop) + (xh1: th1) + (xl: tl) + (xh2: th2 xh1) +: Tot slprop += vmatch xl (| xh1, xh2 |) + +inline_for_extraction +```pulse +fn l2r_write_dtuple2_recip + (#tl #th1: Type) + (#th2: th1 -> Type) + (#vmatch: tl -> dtuple2 th1 th2 -> slprop) + (#vmatch1: tl -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: l2r_writer vmatch1 s1) + (phi: (xl: tl) -> (xh: Ghost.erased (dtuple2 th1 th2)) -> stt_ghost unit emp_inames + (vmatch xl xh) + (fun _ -> vmatch1 xl (dfst xh) ** trade (vmatch1 xl (dfst xh)) (vmatch xl xh)) + ) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: Ghost.erased th1) -> l2r_writer (vmatch_dep_proj2 vmatch xh1) (s2 xh1)) +: l2r_writer #tl #(dtuple2 th1 th2) vmatch #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + phi x' x; + let res1 = w1 x' out offset; + Trade.elim _ _; + with v1 . assert (S.pts_to out v1); + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (| dfst x, dsnd x |)); + fold (vmatch_dep_proj2 vmatch (dfst x) x' (dsnd x)); + let res2 = w2 (dfst x) x' out res1; + unfold (vmatch_dep_proj2 vmatch (dfst x) x' (dsnd x)); + Trade.elim _ _; + with v2 . assert (S.pts_to out v2); + Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); + Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 0 (SZ.v res1) 0 (SZ.v offset); + res2 +} +``` + +inline_for_extraction +```pulse +fn l2r_write_dtuple2_recip_explicit_header + (#tl #th1: Type) + (#th2: th1 -> Type) + (#vmatch: tl -> dtuple2 th1 th2 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: l2r_leaf_writer s1) + (phi: (xl: tl) -> (xh: Ghost.erased (dtuple2 th1 th2)) -> stt th1 + (vmatch xl xh) + (fun res -> vmatch xl xh ** pure (res == dfst xh)) + ) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> l2r_writer (vmatch_dep_proj2 vmatch xh1) (s2 xh1)) +: l2r_writer #tl #(dtuple2 th1 th2) vmatch #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + let xh1 = phi x' x; + let res1 = w1 xh1 out offset; + with v1 . assert (S.pts_to out v1); + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (| xh1, dsnd x |)); + fold (vmatch_dep_proj2 vmatch xh1 x' (dsnd x)); + let res2 = w2 xh1 x' out res1; + unfold (vmatch_dep_proj2 vmatch xh1 x' (dsnd x)); + Trade.elim _ _; + with v2 . assert (S.pts_to out v2); + Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); + Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); + Seq.slice_slice v2 0 (SZ.v res1) 0 (SZ.v offset); + res2 +} +``` + inline_for_extraction ```pulse fn size_nondep_then @@ -1253,10 +1357,10 @@ fn l2r_write_synth inline_for_extraction ```pulse -fn l2r_write_synth2 +fn l2r_write_synth_recip (#t: Type0) (#t1: Type0) (#t2: Type0) (vmatch: t -> t2 -> slprop) - (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) + (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: l2r_writer (vmatch_synth vmatch f2) s1) : l2r_writer #t #t2 vmatch #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) = (x': _) @@ -1325,6 +1429,40 @@ fn l2r_write_filter } ``` +let vmatch_filter_recip + (#tl: Type0) + (#th: Type0) + (f: th -> GTot bool) + (vmatch: tl -> parse_filter_refine f -> slprop) + (xl: tl) + (xh: th) +: Tot slprop += exists* (xh' : parse_filter_refine f) . vmatch xl xh' ** pure (xh == xh') + +inline_for_extraction +```pulse +fn l2r_write_filter_recip + (#t: Type0) (#t1: Type0) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) + (f: (t1 -> GTot bool)) + (vmatch: t -> parse_filter_refine f -> slprop) + (w: l2r_writer (vmatch_filter_recip f vmatch) s1) +: l2r_writer #t #(parse_filter_refine u#0 f) vmatch #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + fold (vmatch_filter_recip f vmatch x' x); + let res = w x' #(Ghost.hide (Ghost.reveal x)) out offset; + unfold (vmatch_filter_recip f vmatch x' x); + with xh . assert (vmatch x' xh); + rewrite (vmatch x' xh) as (vmatch x' x); + res +} +``` + inline_for_extraction ```pulse fn size_filter From 4d569ba1d2b475cab6c814540fd58c9f702f7164 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 1 Oct 2024 09:46:31 -0700 Subject: [PATCH 63/89] l2r_writer_ifthenelse --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 69 ++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 957f83e3d..bfb55f85d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -717,6 +717,75 @@ fn l2r_writer_ext } ``` +inline_for_extraction +```pulse +fn l2r_writer_ifthenelse + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: bool) + (iftrue: (squash (cond == true) -> l2r_writer vmatch s)) + (iffalse: (squash (cond == false) -> l2r_writer vmatch s)) +: l2r_writer #t' #t vmatch #k #p s += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + if (cond) { + iftrue () x' out offset + } else { + iffalse () x' out offset + } +} +``` + +let vmatch_with_cond + (#tl #th: Type) + (vmatch: tl -> th -> slprop) + (cond: tl -> GTot bool) + (xl: tl) + (xh: th) +: Tot slprop += vmatch xl xh ** pure (cond xl) + +let pnot (#t: Type) (cond: t -> GTot bool) (x: t) : GTot bool = not (cond x) + +inline_for_extraction +```pulse +fn l2r_writer_ifthenelse_low + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: (t' -> bool)) + (iftrue: l2r_writer (vmatch_with_cond vmatch cond) s) + (iffalse: l2r_writer (vmatch_with_cond vmatch (pnot cond)) s) +: l2r_writer #t' #t vmatch #k #p s += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + if (cond x') { + fold (vmatch_with_cond vmatch cond x' x); + let res = iftrue x' out offset; + unfold (vmatch_with_cond vmatch cond x' x); + res + } else { + fold (vmatch_with_cond vmatch (pnot cond) x' x); + let res = iffalse x' out offset; + unfold (vmatch_with_cond vmatch (pnot cond) x' x); + res + } +} +``` + let vmatch_and_const (#tl #th: Type) (const: slprop) From a3ba19f2998d9191aa89fa743ceb0a7562616c49 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 1 Oct 2024 15:47:24 -0700 Subject: [PATCH 64/89] serialize seqbytes by copy --- .../pulse/LowParse.Pulse.SeqBytes.fst | 56 +++++++++++++++++++ 1 file changed, 56 insertions(+) create mode 100644 src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst new file mode 100644 index 000000000..72395daa4 --- /dev/null +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -0,0 +1,56 @@ +module LowParse.Pulse.SeqBytes +include LowParse.Pulse.Base +include LowParse.Spec.SeqBytes +open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade + +module S = Pulse.Lib.Slice +module SZ = FStar.SizeT + +let pts_to_seqbytes + (n: nat) + (p: perm) + (s: S.slice byte) + (v: Seq.lseq byte n) +: Tot slprop += exists* (v': Seq.seq byte) . S.pts_to s #p v' ** pure (v' == v) + +inline_for_extraction +```pulse +fn l2r_write_lseq_bytes_copy + (n: Ghost.erased nat) + (p: perm) +: l2r_writer #_ #_ (pts_to_seqbytes n p) #_ #_ (serialize_lseq_bytes n) += + (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + unfold (pts_to_seqbytes n p x' x); + S.pts_to_len out; + S.pts_to_len x'; + let length = S.len x'; + let sp1 = S.split true out offset; + match sp1 { + SlicePair sp11 sp12 -> { + unfold (S.split_post out 1.0R v offset sp1); + unfold (S.split_post' out 1.0R v offset sp11 sp12); + with v12 . assert (S.pts_to sp12 v12); + let sp2 = S.split true sp12 length; + match sp2 { + SlicePair sp21 sp22 -> { + unfold (S.split_post sp12 1.0R v12 length sp2); + unfold (S.split_post' sp12 1.0R v12 length sp21 sp22); + S.pts_to_len sp21; + S.copy sp21 x'; + fold (pts_to_seqbytes n p x' x); + S.join sp21 sp22 sp12; + S.join sp11 sp12 out; + SZ.add offset length; + } + } + } + } +} +``` From deb455cec40d1eedb65ea0089143f3458b4a91e4 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 1 Oct 2024 15:47:32 -0700 Subject: [PATCH 65/89] extensionality with explicit type equality in vmatch --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 37 ++++++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index bfb55f85d..386195d05 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -717,6 +717,43 @@ fn l2r_writer_ext } ``` +let vmatch_ext + (#t' #t1 t2: Type) + (vmatch: t' -> t1 -> slprop) + (x': t') + (x2: t2) +: Tot slprop += exists* (x1: t1) . vmatch x' x1 ** pure (t1 == t2 /\ x1 == x2) + +```pulse +fn l2r_writer_ext_gen + (#t' #t1 #t2: Type0) + (#vmatch: t' -> t2 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t1) + (#s1: serializer p1) + (w: l2r_writer (vmatch_ext t1 vmatch) s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2 { t1 == t2 /\ (forall b . parse p1 b == parse p2 b) }) +: l2r_writer #t' #t2 vmatch #k2 #p2 s2 += (x': t') + (#x: Ghost.erased t2) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + let x1 : Ghost.erased t1 = Ghost.hide #t1 (Ghost.reveal #t2 x); + serializer_unique_strong s1 s2 x1; + fold (vmatch_ext t1 vmatch x' x1); + let res = w x' out offset; + unfold (vmatch_ext t1 vmatch x' x1); + with x2 . assert (vmatch x' x2); + rewrite (vmatch x' x2) as (vmatch x' x); + res +} +``` + inline_for_extraction ```pulse fn l2r_writer_ifthenelse From ef51cfdb398065ca29224248060275674bb0c4fc Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 1 Oct 2024 17:32:48 -0700 Subject: [PATCH 66/89] vmatch should act on low-level data bundled with perm --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 9 ++++++++- src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst | 18 ++++++++---------- 2 files changed, 16 insertions(+), 11 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 386195d05..3af74eaf1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -892,7 +892,7 @@ fn l2r_writer_lens (#vmatch1: t1' -> t -> slprop) (#vmatch2: t2' -> t -> slprop) (lens: vmatch_lens vmatch1 vmatch2) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (w: l2r_writer vmatch2 s) @@ -1177,3 +1177,10 @@ fn compute_remaining_size_constant_size } } ``` + +inline_for_extraction +noeq +type with_perm (t: Type) = { + v: t; + p: perm +} diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst index 72395daa4..c37ca26a9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -8,18 +8,16 @@ module SZ = FStar.SizeT let pts_to_seqbytes (n: nat) - (p: perm) - (s: S.slice byte) + (s: with_perm (S.slice byte)) (v: Seq.lseq byte n) : Tot slprop -= exists* (v': Seq.seq byte) . S.pts_to s #p v' ** pure (v' == v) += exists* (v': Seq.seq byte) . S.pts_to s.v #s.p v' ** pure (v' == v) inline_for_extraction ```pulse fn l2r_write_lseq_bytes_copy (n: Ghost.erased nat) - (p: perm) -: l2r_writer #_ #_ (pts_to_seqbytes n p) #_ #_ (serialize_lseq_bytes n) +: l2r_writer #_ #_ (pts_to_seqbytes n) #_ #_ (serialize_lseq_bytes n) = (x': _) (#x: _) @@ -27,10 +25,10 @@ fn l2r_write_lseq_bytes_copy (offset: _) (#v: _) { - unfold (pts_to_seqbytes n p x' x); + unfold (pts_to_seqbytes n x' x); S.pts_to_len out; - S.pts_to_len x'; - let length = S.len x'; + S.pts_to_len x'.v; + let length = S.len x'.v; let sp1 = S.split true out offset; match sp1 { SlicePair sp11 sp12 -> { @@ -43,8 +41,8 @@ fn l2r_write_lseq_bytes_copy unfold (S.split_post sp12 1.0R v12 length sp2); unfold (S.split_post' sp12 1.0R v12 length sp21 sp22); S.pts_to_len sp21; - S.copy sp21 x'; - fold (pts_to_seqbytes n p x' x); + S.copy sp21 x'.v; + fold (pts_to_seqbytes n x' x); S.join sp21 sp22 sp12; S.join sp11 sp12 out; SZ.add offset length; From bc979a788be7d8dbad100c49122cc8a5fcb83fdf Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 3 Oct 2024 18:01:55 -0700 Subject: [PATCH 67/89] WIP l2r_write_nlist_as_array --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 88 ++++++++++++++++++++ 1 file changed, 88 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 560c2d15f..d2211161f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -247,3 +247,91 @@ fn pts_to_serialized_nlist_1 _ _ } ``` + +module A = Pulse.Lib.Array +module PM = Pulse.Lib.SeqMatch + +let nlist_match_array + (#tarray: Type0) + (#telem: Type0) + (#t: Type) + (varray: (tarray -> with_perm (A.array telem))) + (vmatch: (tarray -> telem -> t -> slprop)) + (n: nat) + (a: tarray) + (l: LowParse.Spec.VCList.nlist n t) +: Tot slprop += exists* c . + A.pts_to (varray a).v #(varray a).p c ** + PM.seq_list_match c l (vmatch a) + +module GR = Pulse.Lib.GhostReference + +inline_for_extraction +```pulse +fn l2r_write_nlist_as_array + (#tarray: Type0) + (#telem: Type0) + (#t: Type0) + (varray: (tarray -> with_perm (A.array telem))) + (vmatch: (tarray -> telem -> t -> slprop)) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (w: ((a: tarray) -> l2r_writer (vmatch a) s)) + (n: SZ.t) +: l2r_writer #_ #_ (nlist_match_array varray vmatch (SZ.v n)) #_ #_ (LowParse.Spec.VCList.serialize_nlist (SZ.v n) s) += + (arr: _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + let a = varray arr; + unfold (nlist_match_array varray vmatch (SZ.v n) arr x); + with c . assert (A.pts_to a.v #a.p c); + let pl1 : GR.ref (list t) = GR.alloc #(list t) []; + let mut pres = offset; + let mut pi = 0sz; + Trade.refl (PM.seq_list_match c x (vmatch arr)); + PM.seq_list_match_length (vmatch arr) _ _; + while ( + let i = !pi; + SZ.lt i n + ) invariant b . exists* res i l1 c2 l2 v1 . ( + A.pts_to a.v #a.p c ** + R.pts_to pres res ** + R.pts_to pi i ** + GR.pts_to pl1 l1 ** + PM.seq_list_match c2 l2 (vmatch arr) ** + pts_to out v1 ** + trade + (PM.seq_list_match c2 l2 (vmatch arr)) + (PM.seq_list_match c x (vmatch arr)) ** + pure ( + SZ.v i <= SZ.v n /\ + b == (SZ.v i < SZ.v n) /\ + Seq.length c == SZ.v n /\ + Seq.equal c2 (Seq.slice c (SZ.v i) (SZ.v n)) /\ + SZ.v offset <= SZ.v res /\ + SZ.v res <= Seq.length v /\ + Seq.length v1 == Seq.length v /\ + Seq.slice v1 0 (SZ.v offset) `Seq.equal` Seq.slice v 0 (SZ.v offset) /\ + List.Tot.length l1 == SZ.v i /\ + Seq.slice v1 (SZ.v offset) (SZ.v res) `Seq.equal` bare_serialize (serialize_nlist (SZ.v i) s) l1 /\ + List.Tot.append l1 l2 == Ghost.reveal x /\ + True + ) + ) { + () + }; + with l1 . assert (GR.pts_to pl1 l1); + GR.free pl1; + PM.seq_list_match_length (vmatch arr) _ _; + List.Tot.append_l_nil l1; + Trade.elim _ _; + fold (nlist_match_array varray vmatch (SZ.v n) arr x); + !pres +} +``` From a03661daff6157e6293a62dcee9e0250acaa7099 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 3 Oct 2024 20:22:29 -0700 Subject: [PATCH 68/89] l2r_write_nlist_as_array --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 134 ++++++++++++++++++- 1 file changed, 133 insertions(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index d2211161f..22a9611d9 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -9,6 +9,41 @@ module SZ = FStar.SizeT module R = Pulse.Lib.Reference module Trade = Pulse.Lib.Trade.Util +#push-options "--z3rlimit 16" + +let rec serialize_nlist_append + (#k: parser_kind) + (#t: Type) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (n1: nat) + (l1: nlist n1 t) + (n2: nat) + (l2: nlist n2 t) +: Lemma + (ensures (let l = List.Tot.append l1 l2 in + List.Tot.length l == n1 + n2 /\ + serialize (serialize_nlist (n1 + n2) s) l == Seq.append (serialize (serialize_nlist n1 s) l1) (serialize (serialize_nlist n2 s) l2) + )) + (decreases n1) += if n1 = 0 + then begin + serialize_nlist_nil p s; + Seq.append_empty_l (serialize (serialize_nlist n2 s) l2) + end + else begin + let a :: q = l1 in + serialize_nlist_append s (n1 - 1) q n2 l2; + serialize_nlist_cons ((n1 - 1) + n2) s a (List.Tot.append q l2); + serialize_nlist_cons (n1 - 1) s a q; + Seq.append_assoc + (serialize s a) + (serialize (serialize_nlist (n1 - 1) s) q) + (serialize (serialize_nlist n2 s) l2) + end + +#pop-options + inline_for_extraction ```pulse fn jump_nlist @@ -249,7 +284,7 @@ fn pts_to_serialized_nlist_1 ``` module A = Pulse.Lib.Array -module PM = Pulse.Lib.SeqMatch +module PM = Pulse.Lib.SeqMatch.Util let nlist_match_array (#tarray: Type0) @@ -267,6 +302,52 @@ let nlist_match_array module GR = Pulse.Lib.GhostReference +let serialize_nlist_singleton + (#k: parser_kind) + (#t: Type) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (x: t) +: Lemma + (let l = [x] in + List.Tot.length l == 1 /\ + serialize (serialize_nlist 1 s) l == serialize s x) += serialize_nlist_cons 0 s x []; + serialize_nlist_nil p s; + Seq.append_empty_r (serialize s x) + +let serialize_nlist_cons' + (#k: parser_kind) + (#t: Type) + (n: nat) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong } ) + (a: t) + (q: nlist n t) +: Lemma + (ensures ( + let l = a :: q in + List.Tot.length l == n + 1 /\ + serialize (serialize_nlist (n + 1) s) l == Seq.append (serialize s a) (serialize (serialize_nlist n s) q) + )) += serialize_nlist_cons n s a q + +let seq_slice_append_ijk + (#t: Type) + (s: Seq.seq t) + (i j k: nat) +: Lemma + (requires (i <= j /\ j <= k /\ k <= Seq.length s)) + (ensures ( + i <= j /\ j <= k /\ k <= Seq.length s /\ + Seq.slice s i k == Seq.append (Seq.slice s i j) (Seq.slice s j k) + )) += Seq.lemma_split (Seq.slice s i k) (j - i) + +#push-options "--z3rlimit 16" + +#restart-solver + inline_for_extraction ```pulse fn l2r_write_nlist_as_array @@ -324,6 +405,55 @@ fn l2r_write_nlist_as_array True ) ) { + let i = !pi; + let off = !pres; + PM.seq_list_match_length (vmatch arr) _ _; + with l1 . assert (GR.pts_to pl1 l1); + with c2 l2 . assert (PM.seq_list_match c2 l2 (vmatch arr)); + serialize_nlist_append s (SZ.v i) l1 (SZ.v n - SZ.v i) l2; + PM.seq_list_match_cons_elim_trade c2 l2 (vmatch arr); + let e = A.op_Array_Access a.v i; + let c2' : Ghost.erased (Seq.seq telem) = Seq.tail c2; + with ve l2' . assert (vmatch arr (Seq.head c2) ve ** PM.seq_list_match c2' l2' (vmatch arr)); + List.Tot.append_assoc l1 [ve] l2'; + let i' = SZ.add i 1sz; + let ni' : Ghost.erased nat = Ghost.hide (SZ.v n - SZ.v i'); + serialize_nlist_cons' (ni') s ve l2'; + serialize_nlist_singleton s ve; + serialize_nlist_append s (SZ.v i) l1 1 [ve]; + Trade.rewrite_with_trade + (vmatch arr _ _) + (vmatch arr e ve); + with v1 . assert (pts_to out v1); + assert (pure ( + SZ.v off + Seq.length (bare_serialize s ve) <= Seq.length v1 + )); + let res = w arr e out off; + with v1' . assert (pts_to out v1'); + Trade.elim (vmatch arr e ve) _; + pi := i'; + List.Tot.append_length l1 [ve]; + let l1' : Ghost.erased (list t) = List.Tot.append l1 [ve]; + GR.op_Colon_Equals pl1 l1'; + pres := res; + Trade.elim_hyp_l _ _ _; + Trade.trans _ _ (PM.seq_list_match c x (vmatch arr)); + assert (pure (Seq.equal c2' (Seq.slice c (SZ.v i') (SZ.v n)))); + assert (pure (SZ.v offset <= SZ.v res)); + assert (pure (SZ.v res <= Seq.length v)); + assert (pure (Seq.length v1' == Seq.length v)); + Seq.slice_slice v1 0 (SZ.v off) 0 (SZ.v offset); + Seq.slice_slice v1' 0 (SZ.v off) 0 (SZ.v offset); + assert (pure (Seq.slice v1' 0 (SZ.v offset) `Seq.equal` Seq.slice v 0 (SZ.v offset))); + assert (pure (List.Tot.length l1' == SZ.v i')); + Seq.slice_slice v1 0 (SZ.v off) (SZ.v offset) (SZ.v off); + Seq.slice_slice v1' 0 (SZ.v off) (SZ.v offset) (SZ.v off); + seq_slice_append_ijk v1' (SZ.v offset) (SZ.v off) (SZ.v res); + assert (pure (Seq.slice v1' (SZ.v offset) (SZ.v off) == bare_serialize (serialize_nlist (SZ.v i) s) l1)); + assert (pure (Seq.slice v1' (SZ.v off) (SZ.v res) == bare_serialize s ve)); + assert (pure (Seq.slice v1' (SZ.v offset) (SZ.v res) == Seq.append (Seq.slice v1' (SZ.v offset) (SZ.v off)) (Seq.slice v1' (SZ.v off) (SZ.v res)))); + assert (pure (Seq.slice v1' (SZ.v offset) (SZ.v res) `Seq.equal` bare_serialize (serialize_nlist (SZ.v i') s) l1')); + assert (pure (List.Tot.append l1' l2' == Ghost.reveal x)); () }; with l1 . assert (GR.pts_to pl1 l1); @@ -335,3 +465,5 @@ fn l2r_write_nlist_as_array !pres } ``` + +#pop-options From e771aa964c4afa2c61bfa7206d92e05a20a45ff2 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 3 Oct 2024 21:46:47 -0700 Subject: [PATCH 69/89] nlist_match_array: make varray partial --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 22a9611d9..ed9b9b4e5 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -290,15 +290,16 @@ let nlist_match_array (#tarray: Type0) (#telem: Type0) (#t: Type) - (varray: (tarray -> with_perm (A.array telem))) + (varray: (tarray -> GTot (option (with_perm (A.array telem))))) (vmatch: (tarray -> telem -> t -> slprop)) (n: nat) (a: tarray) (l: LowParse.Spec.VCList.nlist n t) : Tot slprop -= exists* c . - A.pts_to (varray a).v #(varray a).p c ** - PM.seq_list_match c l (vmatch a) += exists* (ar: with_perm (A.array telem)) c . + A.pts_to ar.v #ar.p c ** + PM.seq_list_match c l (vmatch a) ** + pure (varray a == Some ar) module GR = Pulse.Lib.GhostReference @@ -354,7 +355,7 @@ fn l2r_write_nlist_as_array (#tarray: Type0) (#telem: Type0) (#t: Type0) - (varray: (tarray -> with_perm (A.array telem))) + (varray: (tarray -> option (with_perm (A.array telem)))) (vmatch: (tarray -> telem -> t -> slprop)) (#k: Ghost.erased parser_kind) (#p: parser k t) @@ -369,8 +370,8 @@ fn l2r_write_nlist_as_array (offset: _) (#v: _) { - let a = varray arr; unfold (nlist_match_array varray vmatch (SZ.v n) arr x); + let a = Some?.v (varray arr); with c . assert (A.pts_to a.v #a.p c); let pl1 : GR.ref (list t) = GR.alloc #(list t) []; let mut pres = offset; From 7dd28f6d06838632c9fdcd45bba5ccef1cfaed21 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 3 Oct 2024 23:28:36 -0700 Subject: [PATCH 70/89] move nlist_match_array_intro to LowParse.Pulse.VCList --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 25 ++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index ed9b9b4e5..78f8b6c89 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -301,6 +301,31 @@ let nlist_match_array PM.seq_list_match c l (vmatch a) ** pure (varray a == Some ar) +```pulse +ghost +fn nlist_match_array_intro + (#tarray: Type0) + (#telem: Type0) + (#t: Type0) + (varray: (tarray -> GTot (option (with_perm (A.array telem))))) + (vmatch: (tarray -> telem -> t -> slprop)) + (n: nat) + (a: tarray) + (l: nlist n t) + (ar: with_perm (A.array telem)) + (c: Seq.seq telem) +requires + (A.pts_to ar.v #ar.p c ** + PM.seq_list_match c l (vmatch a) ** + pure (varray a == Some ar) + ) +ensures + (nlist_match_array varray vmatch n a l) +{ + fold (nlist_match_array varray vmatch n a l) +} +``` + module GR = Pulse.Lib.GhostReference let serialize_nlist_singleton From 19ea2f5523b895f3cdf14c7823b152cd7d40368c Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 9 Oct 2024 10:57:03 -0700 Subject: [PATCH 71/89] nlist_sorted --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 178 +++++++++++++++++++ 1 file changed, 178 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index 78f8b6c89..be427a2ed 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -118,6 +118,74 @@ ensures exists* v' . } ``` +let nlist_hd_tl_post' + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_subkind == Some ParserStrong)) + (n: pos) + (input: slice byte) + (pm: perm) + (v: (nlist n t)) + (hd tl: slice byte) +: slprop += pts_to_serialized s hd #pm (List.Tot.hd v) ** + pts_to_serialized (serialize_nlist (n - 1) s) tl #pm (List.Tot.tl v) ** + Trade.trade + (pts_to_serialized s hd #pm (List.Tot.hd v) ** + pts_to_serialized (serialize_nlist (n - 1) s) tl #pm (List.Tot.tl v)) + (pts_to_serialized (serialize_nlist n s) input #pm v) + +let nlist_hd_tl_post + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_subkind == Some ParserStrong)) + (n: pos) + (input: slice byte) + (pm: perm) + (v: (nlist n t)) + (hd_tl: (slice_pair byte)) +: slprop += nlist_hd_tl_post' s sq n input pm v (hd_tl.fst) (hd_tl.snd) + +inline_for_extraction +```pulse +fn nlist_hd_tl + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_subkind == Some ParserStrong)) + (j: jumper p) + (n: Ghost.erased pos) + (input: slice byte) + (#pm: perm) + (#v: Ghost.erased (nlist n t)) +requires + pts_to_serialized (serialize_nlist n s) input #pm v +returns res : slice_pair byte +ensures + nlist_hd_tl_post s sq n input pm v res +{ + nlist_cons_as_nondep_then s n input; + with v' . assert (pts_to_serialized (serialize_nondep_then s (serialize_nlist (n - 1) s)) input #pm v'); + let res = split_nondep_then #_ #(nlist (n - 1) t) s j #(parse_nlist_kind (n - 1) k) #(coerce_eq () (parse_nlist (n - 1) p)) (coerce_eq () (serialize_nlist (n - 1) s <: serializer (parse_nlist (n - 1) p))) input; // FIXME: same as above + match res { + SlicePair s1 s2 -> { + unfold (split_nondep_then_post s (serialize_nlist (n - 1) s) input pm v' res); + unfold (split_nondep_then_post' s (serialize_nlist (n - 1) s) input pm v' s1 s2); + Trade.trans _ _ (pts_to_serialized (serialize_nlist n s) input #pm v); + fold (nlist_hd_tl_post' s sq n input pm v s1 s2); + fold (nlist_hd_tl_post s sq n input pm v res); + res + } + } +} +``` + inline_for_extraction ```pulse fn nlist_hd @@ -233,6 +301,116 @@ ensures exists* v . } ``` +let impl_order_t + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (order: (t -> t -> bool)) += + (a1: slice byte) -> + (a2: slice byte) -> + (#p1: perm) -> + (#p2: perm) -> + (#v1: Ghost.erased t) -> + (#v2: Ghost.erased t) -> + stt bool + (pts_to_serialized s a1 #p1 v1 ** pts_to_serialized s a2 #p2 v2) + (fun res -> pts_to_serialized s a1 #p1 v1 ** pts_to_serialized s a2 #p2 v2 ** pure (res == order v1 v2)) + +#push-options "--z3rlimit 16" + +#restart-solver + +inline_for_extraction +```pulse +fn nlist_sorted + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_subkind == Some ParserStrong)) + (j: jumper p) + (order: Ghost.erased (t -> t -> bool)) + (impl_order: impl_order_t s order) + (n: SZ.t) + (a: slice byte) + (#pm: perm) + (#v: Ghost.erased (nlist (SZ.v n) t)) +requires + (pts_to_serialized (serialize_nlist (SZ.v n) s) a #pm v) +returns res: bool +ensures + (pts_to_serialized (serialize_nlist (SZ.v n) s) a #pm v ** pure (res == List.Tot.sorted order v)) +{ + if (n = 0sz) { + true + } else { + let pl = nlist_hd_tl s sq j (SZ.v n) a; + match pl { + SlicePair s1 s2 -> { + unfold (nlist_hd_tl_post s sq (SZ.v n) a pm v pl); + unfold (nlist_hd_tl_post' s sq (SZ.v n) a pm v s1 s2); + let mut phd = s1; + let mut ptl = s2; + let n' : SZ.t = SZ.sub n 1sz; + let mut pi = n'; + let mut pres = true; + while ( + let i = !pi; + let res = !pres; + (res && SZ.gt i 0sz) + ) invariant cont . exists* shd stl i res hd tl . + R.pts_to phd shd ** + R.pts_to ptl stl ** + R.pts_to pi i ** + R.pts_to pres res ** + pts_to_serialized s shd #pm hd ** + pts_to_serialized (serialize_nlist (SZ.v i) s) stl #pm tl ** + Trade.trade + (pts_to_serialized s shd #pm hd ** + pts_to_serialized (serialize_nlist (SZ.v i) s) stl #pm tl) + (pts_to_serialized (serialize_nlist (SZ.v n) s) a #pm v) ** + pure ( + List.Tot.sorted order v == (res && List.Tot.sorted order (hd :: tl)) /\ + cont == (res && SZ.gt i 0sz) + ) + { + with gi . assert (R.pts_to pi gi); + let stl = !ptl; + with tl . assert (pts_to_serialized (serialize_nlist (SZ.v gi) s) stl #pm tl); + let pl = nlist_hd_tl s sq j (SZ.v gi) stl; + match pl { + SlicePair s1 s2 -> { + unfold (nlist_hd_tl_post s sq (SZ.v gi) stl pm tl pl); + unfold (nlist_hd_tl_post' s sq (SZ.v gi) stl pm tl s1 s2); + let shd = !phd; + let res = impl_order shd s1; + if (res) { + Trade.elim_hyp_l _ _ (pts_to_serialized (serialize_nlist (SZ.v n) s) a #pm v); + Trade.trans _ _ (pts_to_serialized (serialize_nlist (SZ.v n) s) a #pm v); + phd := s1; + ptl := s2; + let i = !pi; + let i' : SZ.t = SZ.sub i 1sz; + pi := i'; + } else { + Trade.elim _ (pts_to_serialized (serialize_nlist (SZ.v gi) s) stl #pm tl); + pres := false; + } + } + } + }; + Trade.elim _ _; + !pres + } + } + } +} +``` + +#pop-options + let synth_nlist_1 (#t: Type) (x: t) From a74cccc0ec1f5a0c718dab6af14f5214a1c1df12 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 9 Oct 2024 17:35:34 -0700 Subject: [PATCH 72/89] l2r_leaf_writer_ifthenelse --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 24 ++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 3af74eaf1..cfc58dc57 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -996,6 +996,30 @@ fn l2r_leaf_writer_ext } ``` +inline_for_extraction +```pulse +fn l2r_leaf_writer_ifthenelse + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: bool) + (iftrue: (squash (cond == true) -> l2r_leaf_writer s)) + (iffalse: (squash (cond == false) -> l2r_leaf_writer s)) +: l2r_leaf_writer u#0 #t #k #p s += (x: t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + if (cond) { + iftrue () x out offset; + } else { + iffalse () x out offset; + } +} +``` + inline_for_extraction ```pulse fn l2r_writer_of_leaf_writer From c9e131f67fe778e874e71f073bf8624f855ba686 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 10 Oct 2024 11:00:47 -0700 Subject: [PATCH 73/89] test extraction of CBOR serializer --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index cfc58dc57..619e404d0 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -694,6 +694,7 @@ let l2r_writer (#x: Ghost.erased t) -> l2r_writer_for vmatch s x' x +inline_for_extraction ```pulse fn l2r_writer_ext (#t' #t: Type0) @@ -725,6 +726,7 @@ let vmatch_ext : Tot slprop = exists* (x1: t1) . vmatch x' x1 ** pure (t1 == t2 /\ x1 == x2) +inline_for_extraction ```pulse fn l2r_writer_ext_gen (#t' #t1 #t2: Type0) @@ -832,6 +834,7 @@ let vmatch_and_const : Tot slprop = const ** vmatch xl xh +inline_for_extraction ```pulse fn l2r_writer_frame (#t' #t: Type0) From 430fac57816f47361b82442becba45d10e05751e Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Thu, 10 Oct 2024 16:06:46 -0700 Subject: [PATCH 74/89] write_header --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 24 ++- .../pulse/LowParse.Pulse.Combinators.fst | 143 ++++++++++++++++++ 2 files changed, 165 insertions(+), 2 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 619e404d0..c977ba77f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -1027,7 +1027,7 @@ inline_for_extraction ```pulse fn l2r_writer_of_leaf_writer (#t: Type) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (w: l2r_leaf_writer s) @@ -1039,8 +1039,28 @@ fn l2r_writer_of_leaf_writer (#v: Ghost.erased bytes) { unfold (eq_as_slprop t x' x); - let res = w x' out offset; fold (eq_as_slprop t x' x); + w x' out offset +} +``` + +inline_for_extraction +```pulse +fn l2r_leaf_writer_of_writer + (#t: Type) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: l2r_writer #t #t (eq_as_slprop t) #k #p s) +: l2r_leaf_writer u#0 #t #k #p s += (x: t) + (out: slice byte) + (offset: SZ.t) + (#v: Ghost.erased bytes) +{ + fold (eq_as_slprop t x x); + let res = w x out offset; + unfold (eq_as_slprop t x x); res } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index b0f11884a..3a8c6be44 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -40,6 +40,22 @@ fn leaf_read_ret } ``` +inline_for_extraction +```pulse +fn l2r_leaf_write_ret + (#t: Type0) + (x: t) + (v_unique: (v' : t) -> Lemma (x == v')) +: l2r_leaf_writer u#0 #t #_ #_ (serialize_ret x v_unique) += (x: _) + (out: _) + (offset: _) + (#v: _) +{ + offset +} +``` + inline_for_extraction let read_ret (#t: Type0) @@ -63,6 +79,9 @@ let read_empty : reader serialize_empty = reader_of_leaf_reader leaf_read_empty inline_for_extraction let jump_empty : jumper parse_empty = jump_ret () +inline_for_extraction +let l2r_leaf_write_empty : l2r_leaf_writer serialize_empty = l2r_leaf_write_ret () (fun _ -> ()) + let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) @@ -1128,6 +1147,17 @@ fn l2r_write_dtuple2_recip } ``` +let lemma_seq_append_ijk + (#t: Type) + (s: Seq.seq t) + (i j k: nat) +: Lemma + (requires (i <= j /\ j <= k /\ k <= Seq.length s)) + (ensures ( + Seq.slice s i k == Seq.append (Seq.slice s i j) (Seq.slice s j k) + )) += assert (Seq.equal (Seq.slice s i k) (Seq.append (Seq.slice s i j) (Seq.slice s j k))) + inline_for_extraction ```pulse fn l2r_write_dtuple2_recip_explicit_header @@ -1170,10 +1200,123 @@ fn l2r_write_dtuple2_recip_explicit_header Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 0 (SZ.v res1) 0 (SZ.v offset); + lemma_seq_append_ijk v2 (SZ.v offset) (SZ.v res1) (SZ.v res2); res2 } ``` +inline_for_extraction +```pulse +fn l2r_leaf_write_dtuple2_phi + (th1: Type0) + (th2: (th1 -> Type0)) + (xl: dtuple2 th1 th2) + (xh: Ghost.erased (dtuple2 th1 th2)) +requires + eq_as_slprop (dtuple2 th1 th2) xl xh +returns res: th1 +ensures + eq_as_slprop (dtuple2 th1 th2) xl xh ** pure (res == dfst xh) +{ + unfold (eq_as_slprop (dtuple2 th1 th2) xl xh); + fold (eq_as_slprop (dtuple2 th1 th2) xl xh); + dfst xl +} +``` + +```pulse +ghost +fn l2r_leaf_write_dtuple2_body_lens_aux + (#th1: Type0) + (#th2: (th1 -> Tot Type0)) + (xh1: th1) + (x': dtuple2 th1 th2) + (x: th2 xh1) +requires + vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1 x' x +returns res: Ghost.erased (th2 xh1) +ensures + eq_as_slprop (th2 xh1) res x ** + Trade.trade + (eq_as_slprop (th2 xh1) res x) + (vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1 x' x) ** + pure ( + dfst x' == xh1 /\ + Ghost.reveal res == dsnd x' + ) +{ + unfold (vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1 x' x); + unfold (eq_as_slprop (dtuple2 th1 th2) x' (| xh1, x |)); + let res : Ghost.erased (th2 xh1) = dsnd x'; + fold (eq_as_slprop (th2 xh1) res x); + ghost fn aux (_: unit) + requires emp ** eq_as_slprop (th2 xh1) res x + ensures vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1 x' x + { + unfold (eq_as_slprop (th2 xh1) res x); + fold (eq_as_slprop (dtuple2 th1 th2) x' (| xh1, x |)); + fold (vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1 x' x); + }; + Trade.intro _ _ _ aux; + res +} +``` + +inline_for_extraction +```pulse +fn l2r_leaf_write_dtuple2_body_lens + (#th1: Type0) + (#th2: (th1 -> Tot Type0)) + (xh1: th1) +: vmatch_lens #(dtuple2 th1 th2) + #(th2 xh1) + #(th2 xh1) + (vmatch_dep_proj2 #(dtuple2 th1 th2) #th1 #th2 (eq_as_slprop (dtuple2 th1 th2)) xh1) + (eq_as_slprop (th2 xh1)) += (x': dtuple2 th1 th2) + (x: Ghost.erased (th2 xh1)) +{ + let _ = l2r_leaf_write_dtuple2_body_lens_aux xh1 x' x; + dsnd x' +} +``` + +inline_for_extraction +let l2r_leaf_write_dtuple2_body + (#th1: Type0) + (#th2: (th1 -> Type0)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> l2r_leaf_writer (s2 xh1)) + (xh1: th1) +: l2r_writer (vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1) (s2 xh1) += l2r_writer_lens + (l2r_leaf_write_dtuple2_body_lens #th1 #th2 xh1) + (l2r_writer_of_leaf_writer (w2 xh1)) + +inline_for_extraction +let l2r_leaf_write_dtuple2 + (#th1: Type0) + (#th2: (th1 -> Type0)) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: l2r_leaf_writer s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> l2r_leaf_writer (s2 xh1)) +: l2r_leaf_writer (serialize_dtuple2 s1 s2) += l2r_leaf_writer_of_writer + (l2r_write_dtuple2_recip_explicit_header + w1 + (l2r_leaf_write_dtuple2_phi th1 th2) + sq + (l2r_leaf_write_dtuple2_body w2) + ) + inline_for_extraction ```pulse fn size_nondep_then From ac651f19dfc69b7cec611f2bc18d7af03d1e75ba Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 11 Oct 2024 16:37:38 -0700 Subject: [PATCH 75/89] adjust proofs to new official Pulse slices --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 34 ++++++++++--------- .../pulse/LowParse.Pulse.Combinators.fst | 20 +++++------ .../pulse/LowParse.Pulse.Endianness.fst | 12 +++---- src/lowparse/pulse/LowParse.Pulse.Int.fst | 8 ++--- .../pulse/LowParse.Pulse.SeqBytes.fst | 18 ++++------ 5 files changed, 44 insertions(+), 48 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index c977ba77f..9fc68666e 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -337,10 +337,13 @@ let peek_post' (consumed: SZ.t) (left right: slice byte) : Tot slprop -= exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** is_split input pm consumed left right ** pure ( += exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** is_split input left right ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ Seq.length (bare_serialize s v1) == SZ.v consumed /\ - parse p v == Some (v1, SZ.v consumed) + begin match parse p v with + | Some (v2, consumed2) -> v1 == v2 /\ SZ.v consumed == consumed2 + | _ -> False + end ) let peek_post @@ -366,14 +369,12 @@ fn peek returns res: slice_pair byte ensures peek_post s input pm v consumed res { - let s1s2 = split false input #pm #v consumed; + let s1s2 = split input consumed; match s1s2 { SlicePair s1 s2 -> { Seq.lemma_split v (SZ.v consumed); let v1 = Ghost.hide (fst (Some?.v (parse p v))); parse_injective #k p (bare_serialize s v1) v; - unfold (split_post input pm v consumed (SlicePair s1 s2)); - unfold (split_post' input pm v consumed s1 s2); with v1' . assert (pts_to s1 #pm v1'); rewrite (pts_to s1 #pm v1') as (pts_to_serialized s s1 #pm v1); fold (peek_post' s input pm v consumed s1 s2); @@ -395,7 +396,10 @@ let peek_trade_post' = exists* v1 v2 . pts_to_serialized s left #pm v1 ** pts_to right #pm v2 ** trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) ** pure ( bare_serialize s v1 `Seq.append` v2 == v /\ Seq.length (bare_serialize s v1) == SZ.v consumed /\ - parse p v == Some (v1, SZ.v consumed) + begin match parse p v with + | Some (v2, consumed2) -> v1 == v2 /\ SZ.v consumed == consumed2 + | _ -> False + end ) let peek_trade_post @@ -424,7 +428,7 @@ fn peek_trade_aux bare_serialize s v1 `Seq.append` v2 == v )) (_: unit) - requires (is_split input pm consumed left right ** (pts_to_serialized s left #pm v1 ** pts_to right #pm v2)) + requires (is_split input left right ** (pts_to_serialized s left #pm v1 ** pts_to right #pm v2)) ensures pts_to input #pm v { unfold (pts_to_serialized s left #pm v1); @@ -449,7 +453,7 @@ fn peek_trade unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); - intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input pm consumed left right) (peek_trade_aux s input pm consumed v left right v1 v2 ()); + intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input left right) (peek_trade_aux s input pm consumed v left right v1 v2 ()); fold (peek_trade_post' s input pm v consumed left right); fold (peek_trade_post s input pm v consumed (left `SlicePair` right)); (left `SlicePair` right) @@ -473,12 +477,10 @@ fn peek_trade_gen parse p (Seq.slice v (SZ.v offset) (Seq.length v)) == Some (v', SZ.v off - SZ.v offset) ) { - let split123 = split_trade false input offset; + let split123 = split_trade input offset; match split123 { SlicePair input1 input23 -> { - unfold (split_trade_post input pm v offset split123); - unfold (split_trade_post' input pm v offset input1 input23); with v23 . assert (pts_to input23 #pm v23); - Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) (pts_to input #pm v); + Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) _; let consumed = SZ.sub off offset; let split23 = peek_trade s input23 consumed; match split23 { SlicePair input2 input3 -> { @@ -670,11 +672,11 @@ let l2r_writer_for (offset: SZ.t) -> (#v: Ghost.erased bytes) -> stt SZ.t - (S.pts_to out v ** vmatch x' x ** pure ( + (pts_to out v ** vmatch x' x ** pure ( SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v )) (fun res -> exists* v' . - S.pts_to out v' ** vmatch x' x ** pure ( + pts_to out v' ** vmatch x' x ** pure ( let bs = bare_serialize s x in SZ.v res == SZ.v offset + Seq.length bs /\ SZ.v res <= Seq.length v /\ @@ -964,11 +966,11 @@ let l2r_leaf_writer (offset: SZ.t) -> (#v: Ghost.erased bytes) -> stt SZ.t - (S.pts_to out v ** pure ( + (pts_to out v ** pure ( SZ.v offset + Seq.length (bare_serialize s x) <= Seq.length v )) (fun res -> exists* v' . - S.pts_to out v' ** pure ( + pts_to out v' ** pure ( let bs = bare_serialize s x in SZ.v res == SZ.v offset + Seq.length bs /\ SZ.v res <= Seq.length v /\ diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 3a8c6be44..9fdeb73ed 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -718,11 +718,9 @@ fn split_dtuple2 (pts_to input #pm (bare_serialize s1 (dfst v) `Seq.append` bare_serialize (s2 (dfst v)) (dsnd v))); parse_serialize_strong_prefix s1 (dfst v) (bare_serialize (s2 (dfst v)) (dsnd v)); let i = j1 input 0sz; - let res = append_split_trade false input i; + let res = append_split_trade input i; match res { SlicePair input1 input2 -> { - unfold (append_split_trade_post input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i res); - unfold (append_split_trade_post' input pm (bare_serialize s1 (dfst v)) (bare_serialize (s2 (dfst v)) (dsnd v)) i input1 input2); Trade.trans (_ ** _) _ _; pts_to_serialized_intro_trade s1 input1 (dfst v); pts_to_serialized_intro_trade (s2 (dfst v)) input2 (dsnd v); @@ -1075,10 +1073,10 @@ fn l2r_write_dtuple2 serialize_dtuple2_eq s1 s2 x; unfold (vmatch_dep_prod vmatch1 vmatch2); let res1 = w1 (dfst x') #(dfst x) out offset; - with v1 . assert (S.pts_to out v1); + with v1 . assert (pts_to out v1); fold (vmatch_and_const (vmatch1 (dfst x') (dfst x)) (vmatch2 (dfst x)) (dsnd x') (dsnd x)); let res2 = w2 (dfst x') (dfst x) (dsnd x') #(dsnd x) out res1; - with v2 . assert (S.pts_to out v2); + with v2 . assert (pts_to out v2); Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); @@ -1130,7 +1128,7 @@ fn l2r_write_dtuple2_recip phi x' x; let res1 = w1 x' out offset; Trade.elim _ _; - with v1 . assert (S.pts_to out v1); + with v1 . assert (pts_to out v1); Trade.rewrite_with_trade (vmatch x' x) (vmatch x' (| dfst x, dsnd x |)); @@ -1138,7 +1136,7 @@ fn l2r_write_dtuple2_recip let res2 = w2 (dfst x) x' out res1; unfold (vmatch_dep_proj2 vmatch (dfst x) x' (dsnd x)); Trade.elim _ _; - with v2 . assert (S.pts_to out v2); + with v2 . assert (pts_to out v2); Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); @@ -1187,7 +1185,7 @@ fn l2r_write_dtuple2_recip_explicit_header serialize_dtuple2_eq s1 s2 x; let xh1 = phi x' x; let res1 = w1 xh1 out offset; - with v1 . assert (S.pts_to out v1); + with v1 . assert (pts_to out v1); Trade.rewrite_with_trade (vmatch x' x) (vmatch x' (| xh1, dsnd x |)); @@ -1195,7 +1193,7 @@ fn l2r_write_dtuple2_recip_explicit_header let res2 = w2 xh1 x' out res1; unfold (vmatch_dep_proj2 vmatch xh1 x' (dsnd x)); Trade.elim _ _; - with v2 . assert (S.pts_to out v2); + with v2 . assert (pts_to out v2); Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); @@ -1398,11 +1396,11 @@ fn l2r_write_nondep_then serialize_nondep_then_eq s1 s2 x; let x1 = f1 x' x; let res1 = w1 x1 #(Ghost.hide (fst x)) out offset; - with v1 . assert (S.pts_to out v1); + with v1 . assert (pts_to out v1); Trade.elim _ _; let x2 = f2 x' x; let res2 = w2 x2 #(Ghost.hide (snd x)) out res1; - with v2 . assert (S.pts_to out v2); + with v2 . assert (pts_to out v2); Seq.slice_slice v1 0 (SZ.v res1) (SZ.v offset) (SZ.v res1); Seq.slice_slice v1 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); Seq.slice_slice v2 (SZ.v offset) (SZ.v res2) 0 (SZ.v res1 - SZ.v offset); diff --git a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst index fd7bd677a..b53c4277d 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Endianness.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Endianness.fst @@ -20,11 +20,11 @@ let be_to_n_t (#v: Ghost.erased (Seq.seq U8.t)) -> (pos: SZ.t) -> stt t - (S.pts_to x #pm v ** pure ( + (pts_to x #pm v ** pure ( SZ.v pos == len /\ len <= Seq.length v )) - (fun res -> S.pts_to x #pm v ** pure ( + (fun res -> pts_to x #pm v ** pure ( SZ.v pos == len /\ len <= Seq.length v /\ u.v res == E.be_to_n (Seq.slice v 0 len) @@ -135,12 +135,12 @@ let n_to_be_t (#v: Ghost.erased (Seq.seq U8.t)) -> (pos: SZ.t) -> stt unit - (S.pts_to x v ** pure ( + (pts_to x v ** pure ( len <= SZ.v pos /\ SZ.v pos <= Seq.length v /\ u.v n < pow2 (8 * len) )) - (fun _ -> exists* v' . S.pts_to x v' ** pure ( + (fun _ -> exists* v' . pts_to x v' ** pure ( len <= SZ.v pos /\ SZ.v pos <= Seq.length v /\ u.v n < pow2 (8 * len) /\ @@ -208,11 +208,11 @@ fn n_to_be_S let lo = u.to_byte n; let hi = u.div256 n; let pos' = pos `SZ.sub` 1sz; - with v1 . assert (S.pts_to x v1); + with v1 . assert (pts_to x v1); Seq.lemma_split (Seq.slice v1 (SZ.v pos - 1) (Seq.length v1)) 1; let _ = ih hi x pos'; S.op_Array_Assignment x pos' lo; - with v2 . assert (S.pts_to x v2); + with v2 . assert (pts_to x v2); Seq.lemma_split (Seq.slice v2 (SZ.v pos - 1) (Seq.length v2)) 1; } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Int.fst b/src/lowparse/pulse/LowParse.Pulse.Int.fst index 335d3bc94..e4a3a77a3 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Int.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Int.fst @@ -56,7 +56,7 @@ fn l2r_leaf_write_u8 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt8.t #parse_u8_ki (pos: SZ.t) (#v: Ghost.erased (Seq.seq byte)) { - S.pts_to_len x; + pts_to_len x; serialize_u8_spec_be n; let pos' = SZ.add pos 1sz; n_to_be_1 n x pos'; @@ -112,7 +112,7 @@ fn l2r_leaf_write_u16 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt16.t #parse_u16 (pos: SZ.t) (#v: Ghost.erased (Seq.seq byte)) { - S.pts_to_len x; + pts_to_len x; serialize_u16_spec_be n; let pos' = SZ.add pos 2sz; n_to_be_2 n x pos'; @@ -168,7 +168,7 @@ fn l2r_leaf_write_u32 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt32.t #parse_u32 (pos: SZ.t) (#v: Ghost.erased (Seq.seq byte)) { - S.pts_to_len x; + pts_to_len x; serialize_u32_spec_be n; let pos' = SZ.add pos 4sz; n_to_be_4 n x pos'; @@ -224,7 +224,7 @@ fn l2r_leaf_write_u64 (_: unit) : l2r_leaf_writer u#0 #FStar.UInt64.t #parse_u64 (pos: SZ.t) (#v: Ghost.erased (Seq.seq byte)) { - S.pts_to_len x; + pts_to_len x; serialize_u64_spec_be n; let pos' = SZ.add pos 8sz; n_to_be_8 n x pos'; diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst index c37ca26a9..9bdc07e35 100644 --- a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -11,7 +11,7 @@ let pts_to_seqbytes (s: with_perm (S.slice byte)) (v: Seq.lseq byte n) : Tot slprop -= exists* (v': Seq.seq byte) . S.pts_to s.v #s.p v' ** pure (v' == v) += exists* (v': Seq.seq byte) . pts_to s.v #s.p v' ** pure (v' == v) inline_for_extraction ```pulse @@ -26,21 +26,17 @@ fn l2r_write_lseq_bytes_copy (#v: _) { unfold (pts_to_seqbytes n x' x); - S.pts_to_len out; - S.pts_to_len x'.v; + pts_to_len out; + pts_to_len x'.v; let length = S.len x'.v; - let sp1 = S.split true out offset; + let sp1 = S.split out offset; match sp1 { SlicePair sp11 sp12 -> { - unfold (S.split_post out 1.0R v offset sp1); - unfold (S.split_post' out 1.0R v offset sp11 sp12); - with v12 . assert (S.pts_to sp12 v12); - let sp2 = S.split true sp12 length; + with v12 . assert (pts_to sp12 v12); + let sp2 = S.split sp12 length; match sp2 { SlicePair sp21 sp22 -> { - unfold (S.split_post sp12 1.0R v12 length sp2); - unfold (S.split_post' sp12 1.0R v12 length sp21 sp22); - S.pts_to_len sp21; + pts_to_len sp21; S.copy sp21 x'.v; fold (pts_to_seqbytes n x' x); S.join sp21 sp22 sp12; From da4929e35a7b6245a9692c03e7e00f05b836387d Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 15 Oct 2024 14:52:50 -0700 Subject: [PATCH 76/89] move lemmas to LowParse.Pulse --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 106 ++++++++++++++++++ .../pulse/LowParse.Pulse.Combinators.fst | 49 ++++++++ .../pulse/LowParse.Pulse.SeqBytes.fst | 32 ++++++ 3 files changed, 187 insertions(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 9fc68666e..5808933f1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -728,6 +728,34 @@ let vmatch_ext : Tot slprop = exists* (x1: t1) . vmatch x' x1 ** pure (t1 == t2 /\ x1 == x2) +```pulse +ghost +fn vmatch_ext_elim_trade + (#t' #t1 t2: Type0) + (vmatch: t' -> t1 -> slprop) + (x': t') + (x2: t2) +requires + vmatch_ext t2 vmatch x' x2 +returns x1: Ghost.erased t1 +ensures + vmatch x' x1 ** + Trade.trade (vmatch x' x1) (vmatch_ext t2 vmatch x' x2) ** + pure (t1 == t2 /\ x1 == x2) +{ + unfold (vmatch_ext t2 vmatch x' x2); + with x1 . assert (vmatch x' x1); + ghost fn aux (_: unit) + requires emp ** vmatch x' x1 + ensures vmatch_ext t2 vmatch x' x2 + { + fold (vmatch_ext t2 vmatch x' x2); + }; + Trade.intro _ _ _ aux; + x1 +} +``` + inline_for_extraction ```pulse fn l2r_writer_ext_gen @@ -793,6 +821,32 @@ let vmatch_with_cond : Tot slprop = vmatch xl xh ** pure (cond xl) +```pulse +ghost +fn vmatch_with_cond_elim_trade + (#tl #th: Type0) + (vmatch: tl -> th -> slprop) + (cond: tl -> GTot bool) + (xl: tl) + (xh: th) +requires + vmatch_with_cond vmatch cond xl xh +ensures + vmatch xl xh ** + Trade.trade (vmatch xl xh) (vmatch_with_cond vmatch cond xl xh) ** + pure (cond xl) +{ + unfold (vmatch_with_cond vmatch cond xl xh); + ghost fn aux (_: unit) + requires emp ** vmatch xl xh + ensures vmatch_with_cond vmatch cond xl xh + { + fold (vmatch_with_cond vmatch cond xl xh) + }; + Trade.intro _ _ _ aux +} +``` + let pnot (#t: Type) (cond: t -> GTot bool) (x: t) : GTot bool = not (cond x) inline_for_extraction @@ -1233,3 +1287,55 @@ type with_perm (t: Type) = { v: t; p: perm } + +let pts_to_serialized_with_perm + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) + (input: with_perm (S.slice byte)) + (v: t) +: Tot slprop += pts_to_serialized s input.v #input.p v + +inline_for_extraction +```pulse +fn l2r_write_copy + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) +: l2r_writer #_ #_ (pts_to_serialized_with_perm s) #_ #_ s += (x: _) + (#v: _) + (out: _) + (offset: SZ.t) + (#w: _) +{ + Trade.rewrite_with_trade + (pts_to_serialized_with_perm s x v) + (pts_to_serialized s x.v #x.p v); + pts_to_serialized_elim_trade s x.v; + Trade.trans _ _ (pts_to_serialized_with_perm s x v); + S.pts_to_len out; + S.pts_to_len x.v; + let length = S.len x.v; + let sp1 = S.split out offset; + match sp1 { + S.SlicePair sp11 sp12 -> { + with v12 . assert (pts_to sp12 v12); + let sp2 = S.split sp12 length; + match sp2 { + S.SlicePair sp21 sp22 -> { + S.pts_to_len sp21; + S.copy sp21 x.v; + S.join sp21 sp22 sp12; + S.join sp11 sp12 out; + Trade.elim _ _; + SZ.add offset length; + } + } + } + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 9fdeb73ed..a80361d41 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -82,6 +82,55 @@ let jump_empty : jumper parse_empty = jump_ret () inline_for_extraction let l2r_leaf_write_empty : l2r_leaf_writer serialize_empty = l2r_leaf_write_ret () (fun _ -> ()) +```pulse +ghost +fn l2r_write_empty_lens_aux + (#tl: Type0) + (vmatch: tl -> unit -> slprop) + (xl: tl) + (v: unit) +requires + vmatch xl v +ensures + eq_as_slprop unit () v ** + Trade.trade + (eq_as_slprop unit () v) + (vmatch xl v) +{ + fold (eq_as_slprop unit () v); + ghost fn aux (_: unit) + requires + vmatch xl v ** eq_as_slprop unit () v + ensures + vmatch xl v + { + unfold (eq_as_slprop unit () v) + }; + Trade.intro _ _ _ aux +} +``` + +inline_for_extraction +```pulse +fn l2r_write_empty_lens + (#tl: Type0) + (vmatch: tl -> unit -> slprop) +: vmatch_lens #_ #_ #_ vmatch (eq_as_slprop unit) += (xl: _) + (v: _) +{ + l2r_write_empty_lens_aux vmatch xl v; + () +} +``` + +inline_for_extraction +let l2r_write_empty + (#tl: Type0) + (vmatch: tl -> unit -> slprop) +: l2r_writer vmatch serialize_empty += l2r_writer_lens (l2r_write_empty_lens vmatch) (l2r_writer_of_leaf_writer l2r_leaf_write_empty) + let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst index 9bdc07e35..4cfcd9b31 100644 --- a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -5,6 +5,7 @@ open Pulse.Lib.Pervasives open Pulse.Lib.Slice.Util open Pulse.Lib.Trade module S = Pulse.Lib.Slice module SZ = FStar.SizeT +module Trade = Pulse.Lib.Trade.Util let pts_to_seqbytes (n: nat) @@ -13,6 +14,37 @@ let pts_to_seqbytes : Tot slprop = exists* (v': Seq.seq byte) . pts_to s.v #s.p v' ** pure (v' == v) +```pulse +ghost +fn pts_to_seqbytes_intro + (n: nat) + (p: perm) + (s: S.slice byte) + (v: bytes) + (res: with_perm (S.slice byte)) +requires + pts_to s #p v ** pure (Seq.length v == n /\ res.v == s /\ res.p == p) +returns v': Ghost.erased (Seq.lseq byte n) +ensures + pts_to_seqbytes n res v' ** + Trade.trade + (pts_to_seqbytes n res v') + (pts_to s #p v) ** + pure (v == Ghost.reveal v') +{ + let v' : Seq.lseq byte n = v; + fold (pts_to_seqbytes n res v'); + ghost fn aux (_: unit) + requires emp ** pts_to_seqbytes n res v' + ensures pts_to s #p v + { + unfold (pts_to_seqbytes n res v') + }; + Trade.intro _ _ _ aux; + v' +} +``` + inline_for_extraction ```pulse fn l2r_write_lseq_bytes_copy From 4caa9b99e1f04791fbd22123ae3581cb0df171d1 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 15 Oct 2024 17:30:29 -0700 Subject: [PATCH 77/89] WIP serialize tagged --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 5808933f1..2de3a0f03 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -854,7 +854,7 @@ inline_for_extraction fn l2r_writer_ifthenelse_low (#t' #t: Type0) (vmatch: t' -> t -> slprop) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (cond: (t' -> bool)) From 6bd6456054fbbc88719d7f04b03cc5a2a39d5922 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 16 Oct 2024 00:03:47 -0700 Subject: [PATCH 78/89] compute_remaining_size combinators --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 304 +++++++++++++++++- src/lowparse/pulse/LowParse.Pulse.BitSum.fst | 24 ++ .../pulse/LowParse.Pulse.Combinators.fst | 256 ++++++++++++++- .../pulse/LowParse.Pulse.SeqBytes.fst | 25 ++ src/lowparse/pulse/LowParse.Pulse.VCList.fst | 85 +++++ 5 files changed, 682 insertions(+), 12 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 2de3a0f03..38eabc5e4 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -791,7 +791,7 @@ inline_for_extraction fn l2r_writer_ifthenelse (#t' #t: Type0) (vmatch: t' -> t -> slprop) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (s: serializer p) (cond: bool) @@ -812,6 +812,27 @@ fn l2r_writer_ifthenelse } ``` +inline_for_extraction +```pulse +fn l2r_writer_zero_size + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_high == Some 0)) +: l2r_writer #t' #t vmatch #k #p s += (x': _) + (#x: _) + (out: _) + (offset: _) + (#v: _) +{ + serialize_length s x; + offset +} +``` + let vmatch_with_cond (#tl #th: Type) (vmatch: tl -> th -> slprop) @@ -1121,6 +1142,17 @@ fn l2r_leaf_writer_of_writer } ``` +inline_for_extraction +let l2r_leaf_writer_zero_size + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash (k.parser_kind_high == Some 0)) +: l2r_leaf_writer u#0 #t #k #p s += l2r_leaf_writer_of_writer + (l2r_writer_zero_size _ s ()) + inline_for_extraction let compute_remaining_size (#t' #t: Type0) @@ -1202,6 +1234,35 @@ fn compute_remaining_size_ext } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_ext_gen + (#t' #t1 #t2: Type0) + (#vmatch: t' -> t2 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t1) + (#s1: serializer p1) + (w: compute_remaining_size (vmatch_ext t1 vmatch) s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t2) + (s2: serializer p2 { t1 == t2 /\ (forall b . parse p1 b == parse p2 b) }) +: compute_remaining_size #t' #t2 vmatch #k2 #p2 s2 += (x': t') + (#x: Ghost.erased t2) + (out: _) + (#v: _) +{ + let x1 : Ghost.erased t1 = Ghost.hide #t1 (Ghost.reveal #t2 x); + serializer_unique_strong s1 s2 x1; + fold (vmatch_ext t1 vmatch x' x1); + let res = w x' out; + unfold (vmatch_ext t1 vmatch x' x1); + with x2 . assert (vmatch x' x2); + rewrite (vmatch x' x2) as (vmatch x' x); + res +} +``` + inline_for_extraction ```pulse fn compute_remaining_size_frame @@ -1233,7 +1294,7 @@ fn compute_remaining_size_lens (#vmatch1: t1' -> t -> slprop) (#vmatch2: t2' -> t -> slprop) (lens: vmatch_lens vmatch1 vmatch2) - (#k: parser_kind) + (#k: Ghost.erased parser_kind) (#p: parser k t) (#s: serializer p) (w: compute_remaining_size vmatch2 s) @@ -1281,6 +1342,214 @@ fn compute_remaining_size_constant_size } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_zero_size + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash ( + k.parser_kind_high == Some 0 + )) +: compute_remaining_size #t' #t vmatch #k #p s += + (x': t') + (#x: Ghost.erased t) + (out: R.ref SZ.t) + (#v: Ghost.erased SZ.t) +{ + serialize_length s x; + true +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_ifthenelse + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: bool) + (iftrue: (squash (cond == true) -> compute_remaining_size vmatch s)) + (iffalse: (squash (cond == false) -> compute_remaining_size vmatch s)) +: compute_remaining_size #t' #t vmatch #k #p s += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + if (cond) { + iftrue () x' out + } else { + iffalse () x' out + } +} +``` + +inline_for_extraction +```pulse +fn compute_remaining_size_ifthenelse_low + (#t' #t: Type0) + (vmatch: t' -> t -> slprop) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: (t' -> bool)) + (iftrue: compute_remaining_size (vmatch_with_cond vmatch cond) s) + (iffalse: compute_remaining_size (vmatch_with_cond vmatch (pnot cond)) s) +: compute_remaining_size #t' #t vmatch #k #p s += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + if (cond x') { + fold (vmatch_with_cond vmatch cond x' x); + let res = iftrue x' out; + unfold (vmatch_with_cond vmatch cond x' x); + res + } else { + fold (vmatch_with_cond vmatch (pnot cond) x' x); + let res = iffalse x' out; + unfold (vmatch_with_cond vmatch (pnot cond) x' x); + res + } +} +``` + +inline_for_extraction +let leaf_compute_remaining_size + (#t: Type0) + (#k: parser_kind) + (#p: parser k t) + (s: serializer p) += (x: t) -> + (out: R.ref SZ.t) -> + (#v: Ghost.erased SZ.t) -> + stt bool + (R.pts_to out v) + (fun res -> exists* v' . + R.pts_to out v' ** pure ( + let bs = Seq.length (bare_serialize s x) in + (res == true <==> bs <= SZ.v v) /\ + (res == true ==> bs + SZ.v v' == SZ.v v) + )) + +inline_for_extraction +```pulse +fn compute_remaining_size_of_leaf_compute_remaining_size + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: leaf_compute_remaining_size s) +: compute_remaining_size #_ #_ (eq_as_slprop _) #_ #_ s += (x: _) + (#v: _) + (out: _) + (#outv: _) +{ + unfold (eq_as_slprop t x v); + fold (eq_as_slprop t x v); + w x out; +} +``` + +inline_for_extraction +```pulse +fn leaf_compute_remaining_size_of_compute_remaining_size + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: compute_remaining_size (eq_as_slprop _) s) +: leaf_compute_remaining_size #_ #_ #_ s += (x: _) + (out: _) + (#outv: _) +{ + fold (eq_as_slprop t x x); + let res = w x out; + unfold (eq_as_slprop t x x); + res +} +``` + +inline_for_extraction +```pulse +fn leaf_compute_remaining_size_ext + (#t: Type0) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 t) + (#s1: serializer p1) + (w1: leaf_compute_remaining_size s1) + (#k2: Ghost.erased parser_kind) + (#p2: parser k2 t) + (s2: serializer p2 { forall x . parse p1 x == parse p2 x }) +: leaf_compute_remaining_size #t #k2 #p2 s2 += (x: t) + (out: _) + (#v: _) +{ + serializer_unique_strong s1 s2 x; + w1 x out +} +``` + +inline_for_extraction +```pulse +fn leaf_compute_remaining_size_ifthenelse + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (cond: bool) + (iftrue: (squash (cond == true) -> leaf_compute_remaining_size s)) + (iffalse: (squash (cond == false) -> leaf_compute_remaining_size s)) +: leaf_compute_remaining_size #t #k #p s += (x: t) + (out: _) + (#v: _) +{ + if (cond) { + iftrue () x out; + } else { + iffalse () x out; + } +} +``` + +inline_for_extraction +let leaf_compute_remaining_size_constant_size + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sz: SZ.t { + k.parser_kind_high == Some k.parser_kind_low /\ + k.parser_kind_low == SZ.v sz + }) +: leaf_compute_remaining_size s += leaf_compute_remaining_size_of_compute_remaining_size + (compute_remaining_size_constant_size _ s sz) + +inline_for_extraction +let leaf_compute_remaining_size_zero_size + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) + (sq: squash ( + k.parser_kind_high == Some 0 + )) +: leaf_compute_remaining_size s += leaf_compute_remaining_size_of_compute_remaining_size + (compute_remaining_size_zero_size _ s sq) + inline_for_extraction noeq type with_perm (t: Type) = { @@ -1339,3 +1608,34 @@ fn l2r_write_copy } } ``` + +inline_for_extraction +```pulse +fn compute_remaining_size_copy + (#t: Type0) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p) +: compute_remaining_size #_ #_ (pts_to_serialized_with_perm s) #_ #_ s += (x: _) + (#v: _) + (out: _) + (#w: _) +{ + Trade.rewrite_with_trade + (pts_to_serialized_with_perm s x v) + (pts_to_serialized s x.v #x.p v); + pts_to_serialized_elim_trade s x.v; + Trade.trans _ _ (pts_to_serialized_with_perm s x v); + S.pts_to_len x.v; + Trade.elim _ _; + let length = S.len x.v; + let cur = !out; + if (SZ.lt cur length) { + false + } else { + out := SZ.sub cur length; + true + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst index 6e33df033..f2ab73d49 100644 --- a/src/lowparse/pulse/LowParse.Pulse.BitSum.fst +++ b/src/lowparse/pulse/LowParse.Pulse.BitSum.fst @@ -81,3 +81,27 @@ fn l2r_write_bitsum' w (sr x) out offset } ``` + +inline_for_extraction +```pulse +fn compute_remaining_size_bitsum' + (#t: eqtype) + (#tot: pos) + (#cl: uint_t tot t) + (#b: bitsum' cl tot) + (sr: synth_bitsum'_recip_t b) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (#s: serializer p) + (w: leaf_compute_remaining_size s) +: (leaf_compute_remaining_size #(bitsum'_type b) #(parse_filter_kind k) #(parse_bitsum' b p) (serialize_bitsum' b s)) += (x: _) + (out: _) + (#v: Ghost.erased bytes) +{ + serialize_bitsum'_eq b s x; + synth_bitsum'_injective b; + synth_bitsum'_recip_inverse b; + w (sr x) out +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index a80361d41..90b96d26c 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -56,6 +56,21 @@ fn l2r_leaf_write_ret } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_ret + (#t: Type0) + (x: t) + (v_unique: (v' : t) -> Lemma (x == v')) +: leaf_compute_remaining_size #t #_ #_ (serialize_ret x v_unique) += (x: _) + (out: _) + (#v: _) +{ + true +} +``` + inline_for_extraction let read_ret (#t: Type0) @@ -82,6 +97,9 @@ let jump_empty : jumper parse_empty = jump_ret () inline_for_extraction let l2r_leaf_write_empty : l2r_leaf_writer serialize_empty = l2r_leaf_write_ret () (fun _ -> ()) +inline_for_extraction +let leaf_compute_remaining_size_empty : leaf_compute_remaining_size serialize_empty = compute_remaining_size_ret () (fun _ -> ()) + ```pulse ghost fn l2r_write_empty_lens_aux @@ -131,6 +149,13 @@ let l2r_write_empty : l2r_writer vmatch serialize_empty = l2r_writer_lens (l2r_write_empty_lens vmatch) (l2r_writer_of_leaf_writer l2r_leaf_write_empty) +inline_for_extraction +let compute_remaining_size_empty + (#tl: Type0) + (vmatch: tl -> unit -> slprop) +: compute_remaining_size vmatch serialize_empty += compute_remaining_size_lens (l2r_write_empty_lens vmatch) (compute_remaining_size_of_leaf_compute_remaining_size leaf_compute_remaining_size_empty) + let parse_serialize_strong_prefix (#t: Type) (#k: parser_kind) @@ -1058,7 +1083,7 @@ let vmatch_dep_prod inline_for_extraction ```pulse -fn size_dtuple2 +fn compute_remaining_size_dtuple2 (#tl1 #tl2 #th1: Type) (#th2: th1 -> Type) (#vmatch1: tl1 -> th1 -> slprop) @@ -1194,6 +1219,51 @@ fn l2r_write_dtuple2_recip } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_dtuple2_recip + (#tl #th1: Type) + (#th2: th1 -> Type) + (#vmatch: tl -> dtuple2 th1 th2 -> slprop) + (#vmatch1: tl -> th1 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: compute_remaining_size vmatch1 s1) + (phi: (xl: tl) -> (xh: Ghost.erased (dtuple2 th1 th2)) -> stt_ghost unit emp_inames + (vmatch xl xh) + (fun _ -> vmatch1 xl (dfst xh) ** trade (vmatch1 xl (dfst xh)) (vmatch xl xh)) + ) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: Ghost.erased th1) -> compute_remaining_size (vmatch_dep_proj2 vmatch xh1) (s2 xh1)) +: compute_remaining_size #tl #(dtuple2 th1 th2) vmatch #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + phi x' x; + let res1 = w1 x' out; + Trade.elim _ _; + if (res1) { + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (| dfst x, dsnd x |)); + fold (vmatch_dep_proj2 vmatch (dfst x) x' (dsnd x)); + let res2 = w2 (dfst x) x' out; + unfold (vmatch_dep_proj2 vmatch (dfst x) x' (dsnd x)); + Trade.elim _ _; + res2 + } else { + false + } +} +``` + let lemma_seq_append_ijk (#t: Type) (s: Seq.seq t) @@ -1252,6 +1322,49 @@ fn l2r_write_dtuple2_recip_explicit_header } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_dtuple2_recip_explicit_header + (#tl #th1: Type) + (#th2: th1 -> Type) + (#vmatch: tl -> dtuple2 th1 th2 -> slprop) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: leaf_compute_remaining_size s1) + (phi: (xl: tl) -> (xh: Ghost.erased (dtuple2 th1 th2)) -> stt th1 + (vmatch xl xh) + (fun res -> vmatch xl xh ** pure (res == dfst xh)) + ) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> compute_remaining_size (vmatch_dep_proj2 vmatch xh1) (s2 xh1)) +: compute_remaining_size #tl #(dtuple2 th1 th2) vmatch #(and_then_kind k1 k2) #(parse_dtuple2 p1 p2) (serialize_dtuple2 s1 s2) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_dtuple2_eq s1 s2 x; + let xh1 = phi x' x; + let res1 = w1 xh1 out; + if (res1) { + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (| xh1, dsnd x |)); + fold (vmatch_dep_proj2 vmatch xh1 x' (dsnd x)); + let res2 = w2 xh1 x' out; + unfold (vmatch_dep_proj2 vmatch xh1 x' (dsnd x)); + Trade.elim _ _; + res2 + } else { + false + } +} +``` + inline_for_extraction ```pulse fn l2r_leaf_write_dtuple2_phi @@ -1342,6 +1455,20 @@ let l2r_leaf_write_dtuple2_body (l2r_leaf_write_dtuple2_body_lens #th1 #th2 xh1) (l2r_writer_of_leaf_writer (w2 xh1)) +inline_for_extraction +let leaf_compute_remaining_size_dtuple2_body + (#th1: Type0) + (#th2: (th1 -> Type0)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> leaf_compute_remaining_size (s2 xh1)) + (xh1: th1) +: compute_remaining_size (vmatch_dep_proj2 (eq_as_slprop (dtuple2 th1 th2)) xh1) (s2 xh1) += compute_remaining_size_lens + (l2r_leaf_write_dtuple2_body_lens #th1 #th2 xh1) + (compute_remaining_size_of_leaf_compute_remaining_size (w2 xh1)) + inline_for_extraction let l2r_leaf_write_dtuple2 (#th1: Type0) @@ -1364,9 +1491,31 @@ let l2r_leaf_write_dtuple2 (l2r_leaf_write_dtuple2_body w2) ) +inline_for_extraction +let leaf_compute_remaining_size_dtuple2 + (#th1: Type0) + (#th2: (th1 -> Type0)) + (#k1: Ghost.erased parser_kind) + (#p1: parser k1 th1) + (#s1: serializer p1) + (w1: leaf_compute_remaining_size s1) + (sq: squash (k1.parser_kind_subkind == Some ParserStrong)) + (#k2: Ghost.erased parser_kind) + (#p2: (x: th1) -> parser k2 (th2 x)) + (#s2: (x: th1) -> serializer (p2 x)) + (w2: (xh1: th1) -> leaf_compute_remaining_size (s2 xh1)) +: leaf_compute_remaining_size (serialize_dtuple2 s1 s2) += leaf_compute_remaining_size_of_compute_remaining_size + (compute_remaining_size_dtuple2_recip_explicit_header + w1 + (l2r_leaf_write_dtuple2_phi th1 th2) + sq + (leaf_compute_remaining_size_dtuple2_body w2) + ) + inline_for_extraction ```pulse -fn size_nondep_then +fn compute_remaining_size_nondep_then (#tl1 #tl2 #th1 #th2: Type) (#vmatch1: tl1 -> th1 -> slprop) (#k1: Ghost.erased parser_kind) @@ -1477,6 +1626,23 @@ fn l2r_leaf_write_synth } ``` +inline_for_extraction +```pulse +fn leaf_compute_remaining_size_synth + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (w: leaf_compute_remaining_size s1) + (#t2: Type0) (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) + (f1': ((x2: t2) -> (x1: t1 { x1 == f1 x2 }))) +: leaf_compute_remaining_size #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += + (x: _) + (out: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + w (f1' x) out +} +``` + inline_for_extraction let mk_synth (#t1 #t2: Type) @@ -1492,6 +1658,13 @@ let l2r_leaf_write_synth' : l2r_leaf_writer u#0 #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) = l2r_leaf_write_synth w f2 f1 (mk_synth f1) +inline_for_extraction +let leaf_compute_remaining_size_synth' + (#k1: Ghost.erased parser_kind) (#t1: Type0) (#p1: parser k1 t1) (#s1: serializer p1) (w: leaf_compute_remaining_size s1) + (#t2: Type0) (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> Tot t1) { synth_inverse f2 f1 }) +: leaf_compute_remaining_size #t2 #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += leaf_compute_remaining_size_synth w f2 f1 (mk_synth f1) + let vmatch_synth (#tl: Type) (#th1 #th2: Type) @@ -1504,7 +1677,7 @@ let vmatch_synth inline_for_extraction ```pulse -fn size_synth +fn compute_remaining_size_synth (#t: Type0) (#t1: Type0) (#t2: Type0) (vmatch: t -> t1 -> slprop) (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size vmatch s1) @@ -1571,6 +1744,31 @@ fn l2r_write_synth_recip } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_synth_recip + (#t: Type0) (#t1: Type0) (#t2: Type0) + (vmatch: t -> t2 -> slprop) + (f2: (t1 -> GTot t2) { synth_injective f2 }) (f1: (t2 -> GTot t1) { synth_inverse f2 f1 }) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size (vmatch_synth vmatch f2) s1) +: compute_remaining_size #t #t2 vmatch #k1 #(parse_synth p1 f2) (serialize_synth p1 f2 s1 f1 ()) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + serialize_synth_eq p1 f2 s1 f1 () x; + Trade.rewrite_with_trade + (vmatch x' x) + (vmatch x' (f2 (f1 x))); + fold (vmatch_synth vmatch f2 x' (f1 x)); + let res = w x' out; + unfold (vmatch_synth vmatch f2 x' (f1 x)); + Trade.elim _ _; + res +} +``` + let vmatch_filter (#tl: Type0) (#th: Type0) @@ -1598,6 +1796,21 @@ fn l2r_leaf_write_filter } ``` +inline_for_extraction +```pulse +fn leaf_compute_remaining_size_filter + (#t1: Type0) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: leaf_compute_remaining_size #t1 s1) + (f: (t1 -> GTot bool)) +: leaf_compute_remaining_size #(parse_filter_refine u#0 f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x: _) + (out: _) + (#v: _) +{ + w x out +} +``` + inline_for_extraction ```pulse fn l2r_write_filter @@ -1619,6 +1832,26 @@ fn l2r_write_filter } ``` +inline_for_extraction +```pulse +fn compute_remaining_size_filter + (#t: Type0) (#t1: Type0) + (vmatch: t -> t1 -> slprop) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size #t #t1 vmatch s1) + (f: (t1 -> GTot bool)) +: compute_remaining_size #t #(parse_filter_refine u#0 f) (vmatch_filter vmatch f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) += (x': _) + (#x: _) + (out: _) + (#v: _) +{ + unfold (vmatch_filter vmatch f x' x); + let res = w x' #(Ghost.hide #t1 (Ghost.reveal x)) out; + fold (vmatch_filter vmatch f x' x); + res +} +``` + let vmatch_filter_recip (#tl: Type0) (#th: Type0) @@ -1655,20 +1888,23 @@ fn l2r_write_filter_recip inline_for_extraction ```pulse -fn size_filter +fn compute_remaining_size_filter_recip (#t: Type0) (#t1: Type0) - (vmatch: t -> t1 -> slprop) - (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (w: compute_remaining_size #t #t1 vmatch s1) + (#k1: Ghost.erased parser_kind) (#p1: parser k1 t1) (#s1: serializer p1) (f: (t1 -> GTot bool)) -: compute_remaining_size #t #(parse_filter_refine u#0 f) (vmatch_filter vmatch f) #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) + (vmatch: t -> parse_filter_refine f -> slprop) + (w: compute_remaining_size (vmatch_filter_recip f vmatch) s1) +: compute_remaining_size #t #(parse_filter_refine u#0 f) vmatch #(parse_filter_kind k1) #(parse_filter p1 f) (serialize_filter s1 f) = (x': _) (#x: _) (out: _) (#v: _) { - unfold (vmatch_filter vmatch f x' x); - let res = w x' #(Ghost.hide #t1 (Ghost.reveal x)) out; - fold (vmatch_filter vmatch f x' x); + fold (vmatch_filter_recip f vmatch x' x); + let res = w x' #(Ghost.hide (Ghost.reveal x)) out; + unfold (vmatch_filter_recip f vmatch x' x); + with xh . assert (vmatch x' xh); + rewrite (vmatch x' xh) as (vmatch x' x); res } ``` diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst index 4cfcd9b31..1bdfa5096 100644 --- a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -80,3 +80,28 @@ fn l2r_write_lseq_bytes_copy } } ``` + +inline_for_extraction +```pulse +fn compute_remaining_size_lseq_bytes_copy + (n: Ghost.erased nat) +: compute_remaining_size #_ #_ (pts_to_seqbytes n) #_ #_ (serialize_lseq_bytes n) += + (x': _) + (#x: _) + (out: _) + (#v: _) +{ + unfold (pts_to_seqbytes n x' x); + pts_to_len x'.v; + fold (pts_to_seqbytes n x' x); + let length = S.len x'.v; + let cur = !out; + if (SZ.lt cur length) { + false + } else { + out := SZ.sub cur length; + true + } +} +``` diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index be427a2ed..f702bfc1b 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -552,6 +552,91 @@ let seq_slice_append_ijk #restart-solver +inline_for_extraction +```pulse +fn compute_remaining_size_nlist_as_array + (#tarray: Type0) + (#telem: Type0) + (#t: Type0) + (varray: (tarray -> option (with_perm (A.array telem)))) + (vmatch: (tarray -> telem -> t -> slprop)) + (#k: Ghost.erased parser_kind) + (#p: parser k t) + (s: serializer p { k.parser_kind_subkind == Some ParserStrong }) + (w: ((a: tarray) -> compute_remaining_size (vmatch a) s)) + (n: SZ.t) +: compute_remaining_size #_ #_ (nlist_match_array varray vmatch (SZ.v n)) #_ #_ (LowParse.Spec.VCList.serialize_nlist (SZ.v n) s) += + (arr: _) + (#x: _) + (out: _) + (#v: _) +{ + unfold (nlist_match_array varray vmatch (SZ.v n) arr x); + let a = Some?.v (varray arr); + with c . assert (A.pts_to a.v #a.p c); + let mut pres = true; + let mut pi = 0sz; + Trade.refl (PM.seq_list_match c x (vmatch arr)); + PM.seq_list_match_length (vmatch arr) _ _; + while ( + let res = !pres; + let i = !pi; + (res && (SZ.lt i n)) + ) invariant b . exists* res i c2 l2 v1 . ( + A.pts_to a.v #a.p c ** + R.pts_to pres res ** + R.pts_to pi i ** + PM.seq_list_match c2 l2 (vmatch arr) ** + pts_to out v1 ** + trade + (PM.seq_list_match c2 l2 (vmatch arr)) + (PM.seq_list_match c x (vmatch arr)) + ** pure ( + SZ.v i <= SZ.v n /\ + b == (res && (SZ.v i < SZ.v n)) /\ + Seq.length c == SZ.v n /\ + (res == false ==> SZ.v v < Seq.length (serialize (serialize_nlist (SZ.v n) s) x)) /\ + (res == true ==> ( + Seq.equal c2 (Seq.slice c (SZ.v i) (SZ.v n)) /\ + List.Tot.length l2 == SZ.v n - SZ.v i /\ + SZ.v v - Seq.length (serialize (serialize_nlist (SZ.v n) s) x) == SZ.v v1 - Seq.length (serialize (serialize_nlist (SZ.v n - SZ.v i) s) l2) + )) /\ + True + ) + ) { + let i = !pi; + PM.seq_list_match_length (vmatch arr) _ _; + with c2 l2 . assert (PM.seq_list_match c2 l2 (vmatch arr)); + PM.seq_list_match_cons_elim_trade c2 l2 (vmatch arr); + let e = A.op_Array_Access a.v i; + let c2' : Ghost.erased (Seq.seq telem) = Seq.tail c2; + with ve l2' . assert (vmatch arr (Seq.head c2) ve ** PM.seq_list_match c2' l2' (vmatch arr)); + let ni' : Ghost.erased nat = Ghost.hide (SZ.v n - SZ.v i - 1); + serialize_nlist_cons' (ni') s ve l2'; + Trade.rewrite_with_trade + (vmatch arr _ _) + (vmatch arr e ve); + let res = w arr e out; + Trade.elim (vmatch arr e ve) _; + if (res) { + let i' = SZ.add i 1sz; + pi := i'; + Trade.elim_hyp_l _ _ _; + Trade.trans _ _ (PM.seq_list_match c x (vmatch arr)); + } else { + Trade.elim _ (PM.seq_list_match c2 l2 (vmatch arr)); + pres := false; + } + }; + Trade.elim _ _; + fold (nlist_match_array varray vmatch (SZ.v n) arr x); + !pres +} +``` + +#restart-solver + inline_for_extraction ```pulse fn l2r_write_nlist_as_array From 1c6683f65c2c368cc6cbdcd0977b3a3ffae83876 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Wed, 16 Oct 2024 11:02:00 -0700 Subject: [PATCH 79/89] extract the validator, the parser and the serializer --- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 1 + 1 file changed, 1 insertion(+) diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index f702bfc1b..f1cce3935 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -301,6 +301,7 @@ ensures exists* v . } ``` +inline_for_extraction let impl_order_t (#t: Type0) (#k: parser_kind) From 367716a8ccb67c5e1561cb51b361c85344b41196 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 15:34:19 -0700 Subject: [PATCH 80/89] towards sharing of Makefiles --- src/common.Makefile | 94 ++++++++++++++++++++++++++++++++++++++++++ src/everparse.Makefile | 9 ++++ src/fstar.Makefile | 29 +++++++++++++ src/karamel.Makefile | 17 ++++++++ src/pulse.Makefile | 33 +++++++++++++++ 5 files changed, 182 insertions(+) create mode 100644 src/common.Makefile create mode 100644 src/everparse.Makefile create mode 100644 src/fstar.Makefile create mode 100644 src/karamel.Makefile create mode 100644 src/pulse.Makefile diff --git a/src/common.Makefile b/src/common.Makefile new file mode 100644 index 000000000..5edd14536 --- /dev/null +++ b/src/common.Makefile @@ -0,0 +1,94 @@ +ifeq (,$(EVERPARSE_SRC_PATH)) + $(error "EVERPARSE_SRC_PATH must be set to the absolute path of the src/ subdirectory of the EverParse repository") +endif + +include $(EVERPARSE_SRC_PATH)/fstar.Makefile + +# List the directories of all root files +SRC_DIRS += . + +# List additional include paths +INCLUDE_PATHS += $(SRC_DIRS) + +# A place to put all build artifacts +ifneq (,$(OUTPUT_DIRECTORY)) + FSTAR_OPTIONS += --odir $(OUTPUT_DIRECTORY) +endif + +# A place to put .checked files. If this variable is left empty, then +# each .checked file will be generated right next to its corresponding +# source file. +ifneq (,$(CACHE_DIRECTORY)) + FSTAR_OPTIONS += --cache_dir $(CACHE_DIRECTORY) + INCLUDE_PATHS+=$(CACHE_DIRECTORY) +endif + +# Used only for OCaml extraction, not krml extraction +# OCaml or Plugin +FSTAR_ML_CODEGEN ?= OCaml + +# Uncomment the definition of PROFILE below, if you want some basic +# profiling of F* runs It will report the time spent +# on typechecking your file And the time spent in SMT, which is +# included in the total typechecking time + +# PROFILE=--profile YOUR_FILE --profile_component 'FStar.Universal.tc_source_file FStar.SMTEncoding' +FSTAR_OPTIONS += $(PROFILE) + +# List the roots from where all dependencies are computed +FSTAR_FILES ?= $(wildcard $(addsuffix /*.fst,$(SRC_DIRS)) $(addsuffix /*.fsti,$(SRC_DIRS))) + +# `ALREADY_CACHED` expected to be empty or to end with a comma +FSTAR_OPTIONS += $(OTHERFLAGS) $(addprefix --include ,$(INCLUDE_PATHS)) --cache_checked_modules --warn_error @241 --already_cached $(ALREADY_CACHED)Prims,FStar,LowStar --cmi + +# Passing RESOURCEMONITOR=1 will create .runlim files through the source tree with +# information about the time and space taken by each F* invocation. +ifneq ($(RESOURCEMONITOR),) + ifeq ($(shell which runlim),) + _ := $(error $(NO_RUNLIM_ERR))) + endif + ifneq ($(MONID),) + MONPREFIX=$(MONID). + endif + RUNLIM=runlim -p -o $@.$(MONPREFIX)runlim +endif + +FSTAR=$(RUNLIM) $(FSTAR_HOME)/bin/fstar.exe $(SIL) $(FSTAR_OPTIONS) + +FSTAR_DEP_FILE ?= .depend + +$(FSTAR_DEP_FILE): $(FSTAR_FILES) + $(call msg, "DEPEND") +ifneq (,$(OUTPUT_DIRECTORY)) + mkdir -p $(OUTPUT_DIRECTORY) +endif +ifneq (,$(CACHE_DIRECTORY)) + mkdir -p $(CACHE_DIRECTORY) +endif + $(Q)true $(shell mkdir -p $(dir $@)) $(shell rm -f $@.rsp) $(foreach f,$(FSTAR_FILES),$(shell echo $(f) >> $@.rsp)) + $(Q)$(FSTAR) $(FSTAR_DEP_OPTIONS) --dep full @$@.rsp --output_deps_to $@.aux + mv $@.aux $@ + +include $(FSTAR_DEP_FILE) + +$(ALL_CHECKED_FILES): %.checked: + $(call msg, "CHECK", $(basename $(notdir $@))) + $(Q)$(RUNLIM) $(FSTAR) $(SIL) $(COMPAT_INDEXED_EFFECTS) $< + touch -c $@ + +verify: $(ALL_CHECKED_FILES) + +%.fst-in %.fsti-in: + @echo $(FSTAR_OPTIONS) + +# Extraction + +$(ALL_ML_FILES): %.ml: + $(FSTAR) $(subst .checked,,$(notdir $<)) --codegen $(FSTAR_ML_CODEGEN) --extract_module $(basename $(notdir $(subst .checked,,$<))) + +$(ALL_KRML_FILES): %.krml: + $(FSTAR) $(notdir $(subst .checked,,$<)) --codegen krml \ + --extract_module $(basename $(notdir $(subst .checked,,$<))) + touch -c $@ + +.PHONY: all verify %.fst-in %.fsti-in diff --git a/src/everparse.Makefile b/src/everparse.Makefile new file mode 100644 index 000000000..a5c86ef17 --- /dev/null +++ b/src/everparse.Makefile @@ -0,0 +1,9 @@ +ifeq (,$(EVERPARSE_SRC_PATH)) + $(error "EVERPARSE_SRC_PATH must be set to the absolute path of the src/ subdirectory of the EverParse repository") +endif + +include $(EVERPARSE_SRC_PATH)/karamel.Makefile + +ALREADY_CACHED := LowParse,$(ALREADY_CACHED) + +INCLUDE_PATHS += $(EVERPARSE_SRC_PATH)/lowparse diff --git a/src/fstar.Makefile b/src/fstar.Makefile new file mode 100644 index 000000000..58b5096e7 --- /dev/null +++ b/src/fstar.Makefile @@ -0,0 +1,29 @@ +ifeq (,$(EVERPARSE_SRC_PATH)) + $(error "EVERPARSE_SRC_PATH must be set to the absolute path of the src/ subdirectory of the EverParse repository") +endif + +# Find fstar.exe and the fstar.lib OCaml package +ifeq (,$(FSTAR_HOME)) + _check_fstar := $(shell which fstar.exe) + ifeq ($(.SHELLSTATUS),0) + FSTAR_HOME := $(realpath $(dir $(_check_fstar))/..) + else +# $(error "FSTAR_HOME is not defined and I cannot find fstar.exe in $(PATH). Please make sure fstar.exe is properly installed and in your PATH or FSTAR_HOME points to its prefix directory or the F* source repository.") + # assuming Everest layout + FSTAR_HOME := $(realpath $(EVERPARSE_SRC_PATH)/../../FStar) + endif +endif +ifeq ($(OS),Windows_NT) + FSTAR_HOME := $(shell cygpath -m $(FSTAR_HOME)) +endif +export FSTAR_HOME +ifeq ($(OS),Windows_NT) + OCAMLPATH := $(shell cygpath -m $(FSTAR_HOME)/lib);$(OCAMLPATH) +else + OCAMLPATH := $(FSTAR_HOME)/lib:$(OCAMLPATH) +endif +export OCAMLPATH +_check_fstar_lib_package := $(shell env OCAMLPATH="$(OCAMLPATH)" ocamlfind query fstar.lib) +ifneq ($(.SHELLSTATUS),0) + $(error "Cannot find fstar.lib in $(OCAMLPATH). Please make sure fstar.exe is properly installed and in your PATH or FSTAR_HOME points to its prefix directory or the F* source repository.") +endif diff --git a/src/karamel.Makefile b/src/karamel.Makefile new file mode 100644 index 000000000..2bdbd70a5 --- /dev/null +++ b/src/karamel.Makefile @@ -0,0 +1,17 @@ +ifeq (,$(EVERPARSE_SRC_PATH)) + $(error "EVERPARSE_SRC_PATH must be set to the absolute path of the src/ subdirectory of the EverParse repository") +endif + +ifeq (,$(KRML_HOME)) + # assuming Everest layout + KRML_HOME := $(realpath $(EVERPARSE_SRC_PATH)/../../karamel) + ifeq (,$(KRML_HOME)) + $(error "KRML_HOME must be defined and set to the root directory of the Karamel repository") + endif +endif + +ALREADY_CACHED := C,LowStar,$(ALREADY_CACHED) + +INCLUDE_PATHS += $(KRML_HOME)/krmllib $(KRML_HOME)/krmllib/obj + +CFLAGS += -I $(KRML_HOME)/include -I $(KRML_HOME)/krmllib/dist/generic diff --git a/src/pulse.Makefile b/src/pulse.Makefile new file mode 100644 index 000000000..df3dd96a2 --- /dev/null +++ b/src/pulse.Makefile @@ -0,0 +1,33 @@ +ifeq (,$(EVERPARSE_SRC_PATH)) + $(error "EVERPARSE_SRC_PATH must be set to the absolute path of the src/ subdirectory of the EverParse repository") +endif + +ifeq (,$(PULSE_LIB)) + ifeq (,$(PULSE_HOME)) + PULSE_LIB := $(shell ocamlfind query pulse) + ifeq (,$(PULSE_LIB)) +# $(error "Pulse should be installed and its lib/ subdirectory should be in ocamlpath; or PULSE_HOME should be defined in the enclosing Makefile as the prefix directory where Pulse was installed, or the root directory of its source repository") + # assuming Everest layout + PULSE_LIB := $(realpath $(EVERPARSE_SRC_PATH)/../../pulse/lib/pulse) + endif + PULSE_HOME := $(realpath $(PULSE_LIB)/../..) + else + PULSE_LIB := $(PULSE_HOME)/lib/pulse + endif +endif +ifeq ($(OS),Windows_NT) + OCAMLPATH := $(PULSE_LIB);$(OCAMLPATH) +else + OCAMLPATH := $(PULSE_LIB):$(OCAMLPATH) +endif +export OCAMLPATH + +# Which modules or namespaces are already cached? If all of your +# source files in your project are under the same namespace, say +# MyNamespace, then you can set this variable to *,-MyNamespace +ALREADY_CACHED := PulseCore,Pulse,$(ALREADY_CACHED) + +# FIXME: do we still need separate subdirectories for pledges, classes? +INCLUDE_PATHS += $(PULSE_LIB) $(PULSE_LIB)/lib $(PULSE_LIB)/lib/class $(PULSE_LIB)/lib/pledge $(PULSE_LIB)/core + +FSTAR_OPTIONS += --load_cmxs pulse From 9a15bae32d50a8933ec75bcee6e5da3c6d5a72bf Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 15:50:50 -0700 Subject: [PATCH 81/89] ignore response files --- src/lowparse/.gitignore | 1 + 1 file changed, 1 insertion(+) diff --git a/src/lowparse/.gitignore b/src/lowparse/.gitignore index bf858426f..d01040b27 100644 --- a/src/lowparse/.gitignore +++ b/src/lowparse/.gitignore @@ -1,3 +1,4 @@ *.checked* *.source *.depend +*.rsp From 1b08afce84c509f90c1b9bbb02393bc501de2051 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 15:50:57 -0700 Subject: [PATCH 82/89] recover build of lowparse/pulse --- src/lowparse/pulse/Makefile | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/lowparse/pulse/Makefile b/src/lowparse/pulse/Makefile index 8267bae2b..db4fe73b8 100644 --- a/src/lowparse/pulse/Makefile +++ b/src/lowparse/pulse/Makefile @@ -1,13 +1,13 @@ all: verify -EVERCBOR_SRC_PATH = $(realpath ../..) +EVERPARSE_SRC_PATH = $(realpath ../..) INCLUDE_PATHS += $(realpath ..) -include $(EVERCBOR_SRC_PATH)/pulse.Makefile -include $(EVERCBOR_SRC_PATH)/karamel.Makefile -include $(EVERCBOR_SRC_PATH)/everparse.Makefile +include $(EVERPARSE_SRC_PATH)/pulse.Makefile +include $(EVERPARSE_SRC_PATH)/karamel.Makefile +include $(EVERPARSE_SRC_PATH)/everparse.Makefile ALREADY_CACHED := *,-LowParse.Pulse, OUTPUT_DIRECTORY := _output -include $(EVERCBOR_SRC_PATH)/common.Makefile +include $(EVERPARSE_SRC_PATH)/common.Makefile From 9b78c608b4ce5ae29c2c33f781d448b9a6f9d2db Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 15:51:20 -0700 Subject: [PATCH 83/89] use --ext context_pruning with new Makefiles --- src/common.Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/common.Makefile b/src/common.Makefile index 5edd14536..99c27a252 100644 --- a/src/common.Makefile +++ b/src/common.Makefile @@ -39,7 +39,7 @@ FSTAR_OPTIONS += $(PROFILE) FSTAR_FILES ?= $(wildcard $(addsuffix /*.fst,$(SRC_DIRS)) $(addsuffix /*.fsti,$(SRC_DIRS))) # `ALREADY_CACHED` expected to be empty or to end with a comma -FSTAR_OPTIONS += $(OTHERFLAGS) $(addprefix --include ,$(INCLUDE_PATHS)) --cache_checked_modules --warn_error @241 --already_cached $(ALREADY_CACHED)Prims,FStar,LowStar --cmi +FSTAR_OPTIONS += $(OTHERFLAGS) $(addprefix --include ,$(INCLUDE_PATHS)) --cache_checked_modules --warn_error @241 --already_cached $(ALREADY_CACHED)Prims,FStar,LowStar --cmi --ext context_pruning # Passing RESOURCEMONITOR=1 will create .runlim files through the source tree with # information about the time and space taken by each F* invocation. From 2c3edccc2291d14e525caad0efb897ec302f31b9 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 15:56:18 -0700 Subject: [PATCH 84/89] CI: install pulse, build lowparse/pulse --- .docker/build/config.json | 1 + .docker/build/install-other-deps.sh | 8 ++++++++ .docker/hierarchic.Dockerfile | 1 + .docker/standalone.Dockerfile | 1 + Makefile | 7 ++++++- 5 files changed, 17 insertions(+), 1 deletion(-) diff --git a/.docker/build/config.json b/.docker/build/config.json index 6db0d193e..76c7ef4ee 100644 --- a/.docker/build/config.json +++ b/.docker/build/config.json @@ -29,6 +29,7 @@ "RepoVersions" : { "mitls_version" : "origin/dev", + "pulse_version" : "origin/main", "karamel_version" : "origin/master" } } diff --git a/.docker/build/install-other-deps.sh b/.docker/build/install-other-deps.sh index 1f6b2351f..5fd71cc38 100755 --- a/.docker/build/install-other-deps.sh +++ b/.docker/build/install-other-deps.sh @@ -16,4 +16,12 @@ pushd $KRML_HOME popd OTHERFLAGS='--admit_smt_queries true' make -j 24 -C "$KRML_HOME" +# Identify the Pulse branch +PULSE_BRANCH=$(jq -c -r '.RepoVersions.pulse_version' "$build_home"/config.json | sed 's!^origin/!!') + +# Install Pulse and its dependencies +[[ -n "$PULSE_HOME" ]] +git clone --branch $PULSE_BRANCH https://github.com/FStarLang/pulse "$PULSE_HOME" +OTHERFLAGS='--admit_smt_queries true' make -j 24 -C "$PULSE_HOME" + opam install hex re ctypes sha sexplib diff --git a/.docker/hierarchic.Dockerfile b/.docker/hierarchic.Dockerfile index 24ddad49d..5bf4ca36c 100644 --- a/.docker/hierarchic.Dockerfile +++ b/.docker/hierarchic.Dockerfile @@ -8,6 +8,7 @@ WORKDIR $HOME/everparse # Dependencies (opam packages) ENV KRML_HOME=$HOME/everparse/karamel +ENV PULSE_HOME=$HOME/everparse/pulse RUN sudo apt-get update && eval $(opam env) && .docker/build/install-other-deps.sh # CI dependencies: sphinx (for the docs) diff --git a/.docker/standalone.Dockerfile b/.docker/standalone.Dockerfile index 2c0e9d728..5ae5c463c 100644 --- a/.docker/standalone.Dockerfile +++ b/.docker/standalone.Dockerfile @@ -15,6 +15,7 @@ RUN sudo apt-get update && sudo apt-get install -y --no-install-recommends \ # Dependencies (F*, Karamel and opam packages) ENV FSTAR_HOME=$HOME/FStar ENV KRML_HOME=$HOME/karamel +ENV PULSE_HOME=$HOME/karamel RUN eval $(opam env) && .docker/build/install-deps.sh # CI dependencies: sphinx (for the docs) diff --git a/Makefile b/Makefile index 5d99e91e5..6861d6b0c 100644 --- a/Makefile +++ b/Makefile @@ -52,7 +52,12 @@ quackyducky-test: quackyducky-unit-test quackyducky-sample-test quackyducky-samp test: all lowparse-test quackyducky-test 3d-test asn1-test -ci: test +lowparse-pulse: lowparse + +$(MAKE) -C src/lowparse/pulse + +.PHONY: lowparse-pulse + +ci: test lowparse-pulse clean-3d: +$(MAKE) -C src/3d clean From 2ae077b029542e98ebc96df02ea2eef442cc3350 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Fri, 25 Oct 2024 18:03:40 -0700 Subject: [PATCH 85/89] also install Pulse to build the binary package (even though it's not used yet) --- .docker/produce-binary.Dockerfile | 1 + 1 file changed, 1 insertion(+) diff --git a/.docker/produce-binary.Dockerfile b/.docker/produce-binary.Dockerfile index 501cf727c..855bc5f9b 100644 --- a/.docker/produce-binary.Dockerfile +++ b/.docker/produce-binary.Dockerfile @@ -17,6 +17,7 @@ RUN sudo apt-get update && sudo apt-get install -y --no-install-recommends \ # Dependencies (F*, Karamel and opam packages) ENV FSTAR_HOME=$HOME/FStar ENV KRML_HOME=$HOME/karamel +ENV PULSE_HOME=$HOME/pulse RUN eval $(opam env) && .docker/build/install-deps.sh # CI proper From bfeaa303f038e987fb0d1aa3404ec4a35a6a7eac Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Mon, 28 Oct 2024 16:01:56 -0700 Subject: [PATCH 86/89] fix standalone CI --- .docker/standalone.Dockerfile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.docker/standalone.Dockerfile b/.docker/standalone.Dockerfile index 5ae5c463c..f1f510518 100644 --- a/.docker/standalone.Dockerfile +++ b/.docker/standalone.Dockerfile @@ -15,7 +15,7 @@ RUN sudo apt-get update && sudo apt-get install -y --no-install-recommends \ # Dependencies (F*, Karamel and opam packages) ENV FSTAR_HOME=$HOME/FStar ENV KRML_HOME=$HOME/karamel -ENV PULSE_HOME=$HOME/karamel +ENV PULSE_HOME=$HOME/pulse RUN eval $(opam env) && .docker/build/install-deps.sh # CI dependencies: sphinx (for the docs) From 3593766c38c5edc294b958d3ec13c6e567fed947 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 29 Oct 2024 15:32:44 -0700 Subject: [PATCH 87/89] (TEMP) use FStarLang/pulse#247 --- .docker/build/config.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.docker/build/config.json b/.docker/build/config.json index 76c7ef4ee..3bfbe7f38 100644 --- a/.docker/build/config.json +++ b/.docker/build/config.json @@ -29,7 +29,7 @@ "RepoVersions" : { "mitls_version" : "origin/dev", - "pulse_version" : "origin/main", + "pulse_version" : "origin/_taramana_no_slice_pair", "karamel_version" : "origin/master" } } From 28f4e3b124ab141a21070e76ff1d88aaa70ae76c Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 29 Oct 2024 16:12:16 -0700 Subject: [PATCH 88/89] remove slice_pair --- src/lowparse/pulse/LowParse.Pulse.Base.fst | 32 ++++++++--------- .../pulse/LowParse.Pulse.Combinators.fst | 34 +++++++++---------- .../pulse/LowParse.Pulse.SeqBytes.fst | 4 +-- src/lowparse/pulse/LowParse.Pulse.VCList.fst | 12 +++---- 4 files changed, 41 insertions(+), 41 deletions(-) diff --git a/src/lowparse/pulse/LowParse.Pulse.Base.fst b/src/lowparse/pulse/LowParse.Pulse.Base.fst index 38eabc5e4..5572aee18 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Base.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Base.fst @@ -352,9 +352,9 @@ let peek_post (pm: perm) (v: bytes) (consumed: SZ.t) - (res: slice_pair byte) + (res: (slice byte & slice byte)) : Tot slprop -= let SlicePair left right = res in += let (left, right) = res in peek_post' s input pm v consumed left right inline_for_extraction @@ -366,20 +366,20 @@ fn peek (#v: Ghost.erased bytes) (consumed: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) - returns res: slice_pair byte + returns res: (slice byte & slice byte) ensures peek_post s input pm v consumed res { let s1s2 = split input consumed; match s1s2 { - SlicePair s1 s2 -> { + Mktuple2 s1 s2 -> { Seq.lemma_split v (SZ.v consumed); let v1 = Ghost.hide (fst (Some?.v (parse p v))); parse_injective #k p (bare_serialize s v1) v; with v1' . assert (pts_to s1 #pm v1'); rewrite (pts_to s1 #pm v1') as (pts_to_serialized s s1 #pm v1); fold (peek_post' s input pm v consumed s1 s2); - fold (peek_post s input pm v consumed (SlicePair s1 s2)); - (SlicePair s1 s2) + fold (peek_post s input pm v consumed (s1, s2)); + (s1, s2) } } } @@ -408,9 +408,9 @@ let peek_trade_post (pm: perm) (v: bytes) (consumed: SZ.t) - (res: slice_pair byte) + (res: (slice byte & slice byte)) : Tot slprop -= let (SlicePair left right) = res in += let (left, right) = res in peek_trade_post' s input pm v consumed left right ```pulse @@ -445,18 +445,18 @@ fn peek_trade (#v: Ghost.erased bytes) (consumed: SZ.t) requires (pts_to input #pm v ** pure (validator_success #k #t p 0sz v (consumed))) - returns res: (slice_pair byte) + returns res: (slice byte & slice byte) ensures peek_trade_post s input pm v consumed res { let res = peek s input consumed; - match res { SlicePair left right -> { + match res { Mktuple2 left right -> { unfold (peek_post s input pm v consumed res); unfold (peek_post' s input pm v consumed left right); with v1 v2 . assert (pts_to_serialized s left #pm v1 ** pts_to right #pm v2); intro_trade (pts_to_serialized s left #pm v1 ** pts_to right #pm v2) (pts_to input #pm v) (is_split input left right) (peek_trade_aux s input pm consumed v left right v1 v2 ()); fold (peek_trade_post' s input pm v consumed left right); - fold (peek_trade_post s input pm v consumed (left `SlicePair` right)); - (left `SlicePair` right) + fold (peek_trade_post s input pm v consumed (left, right)); + (left, right) }} } ``` @@ -478,12 +478,12 @@ fn peek_trade_gen ) { let split123 = split_trade input offset; - match split123 { SlicePair input1 input23 -> { + match split123 { Mktuple2 input1 input23 -> { with v23 . assert (pts_to input23 #pm v23); Trade.elim_hyp_l (pts_to input1 #pm _) (pts_to input23 #pm v23) _; let consumed = SZ.sub off offset; let split23 = peek_trade s input23 consumed; - match split23 { SlicePair input2 input3 -> { + match split23 { Mktuple2 input2 input3 -> { unfold (peek_trade_post s input23 pm v23 consumed split23); unfold (peek_trade_post' s input23 pm v23 consumed input2 input3); with v' . assert (pts_to_serialized s input2 #pm v'); @@ -1591,11 +1591,11 @@ fn l2r_write_copy let length = S.len x.v; let sp1 = S.split out offset; match sp1 { - S.SlicePair sp11 sp12 -> { + Mktuple2 sp11 sp12 -> { with v12 . assert (pts_to sp12 v12); let sp2 = S.split sp12 length; match sp2 { - S.SlicePair sp21 sp22 -> { + Mktuple2 sp21 sp22 -> { S.pts_to_len sp21; S.copy sp21 x.v; S.join sp21 sp22 sp12; diff --git a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst index 90b96d26c..1d49997a1 100644 --- a/src/lowparse/pulse/LowParse.Pulse.Combinators.fst +++ b/src/lowparse/pulse/LowParse.Pulse.Combinators.fst @@ -762,9 +762,9 @@ let split_dtuple2_post (input: slice byte) (pm: perm) (v: Ghost.erased (dtuple2 t1 t2)) - (res: slice_pair byte) + (res: (slice byte & slice byte)) : Tot slprop -= let (SlicePair left right) = res in += let (left, right) = res in split_dtuple2_post' s1 s2 input pm v left right inline_for_extraction @@ -783,7 +783,7 @@ fn split_dtuple2 (#pm: perm) (#v: Ghost.erased (dtuple2 t1 t2)) requires pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v - returns res: slice_pair byte + returns res: (slice byte & slice byte) ensures split_dtuple2_post s1 s2 input pm v res { serialize_dtuple2_eq s1 s2 v; @@ -794,15 +794,15 @@ fn split_dtuple2 let i = j1 input 0sz; let res = append_split_trade input i; match res { - SlicePair input1 input2 -> { + Mktuple2 input1 input2 -> { Trade.trans (_ ** _) _ _; pts_to_serialized_intro_trade s1 input1 (dfst v); pts_to_serialized_intro_trade (s2 (dfst v)) input2 (dsnd v); Trade.prod (pts_to_serialized s1 input1 #pm _) (pts_to input1 #pm _) (pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input2 #pm _); Trade.trans (pts_to_serialized s1 input1 #pm _ ** pts_to_serialized (s2 (dfst v)) input2 #pm _) (pts_to input1 #pm _ ** pts_to input2 #pm _) _; fold (split_dtuple2_post' s1 s2 input pm v input1 input2); - fold (split_dtuple2_post s1 s2 input pm v (input1 `SlicePair` input2)); - (input1 `SlicePair` input2) + fold (split_dtuple2_post s1 s2 input pm v (input1, input2)); + (input1, input2) } } } @@ -829,7 +829,7 @@ fn dtuple2_dfst trade (pts_to_serialized s1 res #pm (dfst v)) (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; - match spl { SlicePair input1 input2 -> { + match spl { Mktuple2 input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v spl); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); Trade.elim_hyp_r _ _ _; @@ -859,7 +859,7 @@ fn dtuple2_dsnd trade (pts_to_serialized (s2 (dfst v)) res #pm (dsnd v)) (pts_to_serialized (serialize_dtuple2 s1 s2) input #pm v) { let spl = split_dtuple2 s1 j1 s2 input; - match spl { SlicePair input1 input2 -> { + match spl { Mktuple2 input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v spl); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); Trade.elim_hyp_l _ _ _; @@ -898,9 +898,9 @@ let split_nondep_then_post (input: slice byte) (pm: perm) (v: Ghost.erased (t1 & t2)) - (res: slice_pair byte) + (res: (slice byte & slice byte)) : Tot slprop -= let (SlicePair left right) = res in += let (left, right) = res in split_nondep_then_post' s1 s2 input pm v left right #set-options "--print_implicits" @@ -944,7 +944,7 @@ fn split_nondep_then (#pm: perm) (#v: Ghost.erased (t1 & t2)) requires pts_to_serialized (serialize_nondep_then s1 s2) input #pm v - returns res: slice_pair byte + returns res: (slice byte & slice byte) ensures split_nondep_then_post s1 s2 input pm v res { pts_to_serialized_ext_trade' @@ -965,13 +965,13 @@ fn split_nondep_then input; Trade.trans (pts_to_serialized (serialize_dtuple2 s1 #k2 #(const_fun t2) #(const_fun p2) (const_fun s2)) _ #pm _) _ _; let res = split_dtuple2 #t1 #(const_fun t2) s1 j1 #_ #(const_fun p2) (const_fun s2) input; - match res { SlicePair input1 input2 -> { + match res { Mktuple2 input1 input2 -> { unfold (split_dtuple2_post #t1 #(const_fun t2) s1 #k2 #(const_fun p2) (const_fun s2) input pm (dtuple2_of_pair v) res); unfold (split_dtuple2_post' #t1 #(const_fun t2) s1 #_ #(const_fun p2) (const_fun s2) input pm (dtuple2_of_pair v) input1 input2); Trade.trans (_ ** _) _ _; fold (split_nondep_then_post' s1 s2 input pm v input1 input2); - fold (split_nondep_then_post s1 s2 input pm v (input1 `SlicePair` input2)); - (input1 `SlicePair` input2) + fold (split_nondep_then_post s1 s2 input pm v (input1, input2)); + (input1, input2) }} } ``` @@ -996,7 +996,7 @@ fn nondep_then_fst trade (pts_to_serialized s1 res #pm (fst v)) (pts_to_serialized (serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; - match spl { SlicePair input1 input2 -> { + match spl { Mktuple2 input1 input2 -> { unfold (split_nondep_then_post s1 s2 input pm v spl); unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); Trade.elim_hyp_r _ _ _; @@ -1025,7 +1025,7 @@ fn nondep_then_snd trade (pts_to_serialized s2 res #pm (snd v)) (pts_to_serialized (serialize_nondep_then s1 s2) input #pm v) { let spl = split_nondep_then s1 j1 s2 input; - match spl { SlicePair input1 input2 -> { + match spl { Mktuple2 input1 input2 -> { unfold (split_nondep_then_post s1 s2 input pm v spl); unfold (split_nondep_then_post' s1 s2 input pm v input1 input2); Trade.elim_hyp_l _ _ _; @@ -1057,7 +1057,7 @@ fn read_dtuple2 (f: _) { let split12 = split_dtuple2 s1 j1 s2 input; - match split12 { SlicePair input1 input2 -> { + match split12 { Mktuple2 input1 input2 -> { unfold (split_dtuple2_post s1 s2 input pm v split12); unfold (split_dtuple2_post' s1 s2 input pm v input1 input2); let x1 = leaf_reader_of_reader r1 input1; diff --git a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst index 1bdfa5096..65423161f 100644 --- a/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst +++ b/src/lowparse/pulse/LowParse.Pulse.SeqBytes.fst @@ -63,11 +63,11 @@ fn l2r_write_lseq_bytes_copy let length = S.len x'.v; let sp1 = S.split out offset; match sp1 { - SlicePair sp11 sp12 -> { + Mktuple2 sp11 sp12 -> { with v12 . assert (pts_to sp12 v12); let sp2 = S.split sp12 length; match sp2 { - SlicePair sp21 sp22 -> { + Mktuple2 sp21 sp22 -> { pts_to_len sp21; S.copy sp21 x'.v; fold (pts_to_seqbytes n x' x); diff --git a/src/lowparse/pulse/LowParse.Pulse.VCList.fst b/src/lowparse/pulse/LowParse.Pulse.VCList.fst index f1cce3935..9d2c2a1b0 100644 --- a/src/lowparse/pulse/LowParse.Pulse.VCList.fst +++ b/src/lowparse/pulse/LowParse.Pulse.VCList.fst @@ -147,9 +147,9 @@ let nlist_hd_tl_post (input: slice byte) (pm: perm) (v: (nlist n t)) - (hd_tl: (slice_pair byte)) + (hd_tl: (slice byte & slice byte)) : slprop -= nlist_hd_tl_post' s sq n input pm v (hd_tl.fst) (hd_tl.snd) += nlist_hd_tl_post' s sq n input pm v (fst hd_tl) (snd hd_tl) inline_for_extraction ```pulse @@ -166,7 +166,7 @@ fn nlist_hd_tl (#v: Ghost.erased (nlist n t)) requires pts_to_serialized (serialize_nlist n s) input #pm v -returns res : slice_pair byte +returns res : (slice byte & slice byte) ensures nlist_hd_tl_post s sq n input pm v res { @@ -174,7 +174,7 @@ ensures with v' . assert (pts_to_serialized (serialize_nondep_then s (serialize_nlist (n - 1) s)) input #pm v'); let res = split_nondep_then #_ #(nlist (n - 1) t) s j #(parse_nlist_kind (n - 1) k) #(coerce_eq () (parse_nlist (n - 1) p)) (coerce_eq () (serialize_nlist (n - 1) s <: serializer (parse_nlist (n - 1) p))) input; // FIXME: same as above match res { - SlicePair s1 s2 -> { + Mktuple2 s1 s2 -> { unfold (split_nondep_then_post s (serialize_nlist (n - 1) s) input pm v' res); unfold (split_nondep_then_post' s (serialize_nlist (n - 1) s) input pm v' s1 s2); Trade.trans _ _ (pts_to_serialized (serialize_nlist n s) input #pm v); @@ -349,7 +349,7 @@ ensures } else { let pl = nlist_hd_tl s sq j (SZ.v n) a; match pl { - SlicePair s1 s2 -> { + Mktuple2 s1 s2 -> { unfold (nlist_hd_tl_post s sq (SZ.v n) a pm v pl); unfold (nlist_hd_tl_post' s sq (SZ.v n) a pm v s1 s2); let mut phd = s1; @@ -382,7 +382,7 @@ ensures with tl . assert (pts_to_serialized (serialize_nlist (SZ.v gi) s) stl #pm tl); let pl = nlist_hd_tl s sq j (SZ.v gi) stl; match pl { - SlicePair s1 s2 -> { + Mktuple2 s1 s2 -> { unfold (nlist_hd_tl_post s sq (SZ.v gi) stl pm tl pl); unfold (nlist_hd_tl_post' s sq (SZ.v gi) stl pm tl s1 s2); let shd = !phd; From 848ec7a814b2977baa5f33e5ca2e7944733e8694 Mon Sep 17 00:00:00 2001 From: Tahina Ramananandro Date: Tue, 29 Oct 2024 19:36:21 -0700 Subject: [PATCH 89/89] Revert "(TEMP) use FStarLang/pulse#247" This reverts commit 3593766c38c5edc294b958d3ec13c6e567fed947. --- .docker/build/config.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.docker/build/config.json b/.docker/build/config.json index 3bfbe7f38..76c7ef4ee 100644 --- a/.docker/build/config.json +++ b/.docker/build/config.json @@ -29,7 +29,7 @@ "RepoVersions" : { "mitls_version" : "origin/dev", - "pulse_version" : "origin/_taramana_no_slice_pair", + "pulse_version" : "origin/main", "karamel_version" : "origin/master" } }