diff options
Diffstat (limited to 'theories/nix')
| -rw-r--r-- | theories/nix/floats.v | 85 | ||||
| -rw-r--r-- | theories/nix/interp.v | 351 | ||||
| -rw-r--r-- | theories/nix/interp_proofs.v | 2690 | ||||
| -rw-r--r-- | theories/nix/notations.v | 43 | ||||
| -rw-r--r-- | theories/nix/operational.v | 527 | ||||
| -rw-r--r-- | theories/nix/operational_props.v | 680 | ||||
| -rw-r--r-- | theories/nix/tests.v | 185 | ||||
| -rw-r--r-- | theories/nix/wp.v | 143 | ||||
| -rw-r--r-- | theories/nix/wp_examples.v | 164 |
9 files changed, 4868 insertions, 0 deletions
diff --git a/theories/nix/floats.v b/theories/nix/floats.v new file mode 100644 index 0000000..246e0c3 --- /dev/null +++ b/theories/nix/floats.v | |||
| @@ -0,0 +1,85 @@ | |||
| 1 | From stdpp Require Import prelude ssreflect. | ||
| 2 | From Flocq.IEEE754 Require Import | ||
| 3 | Binary BinarySingleNaN (mode_NE, mode_DN, mode_UP) Bits. | ||
| 4 | From stdpp Require Import options. | ||
| 5 | |||
| 6 | Global Arguments B754_zero {_ _}. | ||
| 7 | Global Arguments B754_infinity {_ _}. | ||
| 8 | Global Arguments B754_nan {_ _}. | ||
| 9 | Global Arguments B754_finite {_ _}. | ||
| 10 | |||
| 11 | (** The bit representation of floats is not observable in Nix, and it appears | ||
| 12 | that only negative NaNs are ever produced. So we setup the Flocq floats in | ||
| 13 | the way that it always produces [-NaN], i.e., [indef_nan]. *) | ||
| 14 | Definition float := binary64. | ||
| 15 | |||
| 16 | Variant round_mode := Floor | Ceil | NearestEven. | ||
| 17 | Global Instance round_mode_eq_dec : EqDecision round_mode. | ||
| 18 | Proof. solve_decision. Defined. | ||
| 19 | |||
| 20 | Module Float. | ||
| 21 | Definition prec : Z := 53. | ||
| 22 | Definition emax : Z := 1024. | ||
| 23 | |||
| 24 | Lemma Hprec : (0 < 53)%Z. | ||
| 25 | Proof. done. Qed. | ||
| 26 | Lemma Hprec_emax : (53 < 1024)%Z. | ||
| 27 | Proof. done. Qed. | ||
| 28 | |||
| 29 | Global Instance inhabited : Inhabited float := populate (B754_zero false). | ||
| 30 | |||
| 31 | Global Instance eq_dec : EqDecision float. | ||
| 32 | Proof. | ||
| 33 | refine (λ f1 f2, | ||
| 34 | match f1, f2 with | ||
| 35 | | B754_zero s1, B754_zero s2 => cast_if (decide (s1 = s2)) | ||
| 36 | | B754_infinity s1, B754_infinity s2 => cast_if (decide (s1 = s2)) | ||
| 37 | | B754_nan s1 pl1 _, B754_nan s2 pl2 _ => | ||
| 38 | cast_if_and (decide (s1 = s2)) (decide (pl1 = pl2)) | ||
| 39 | | B754_finite s1 m1 e1 _, B754_finite s2 m2 e2 _ => | ||
| 40 | cast_if_and3 (decide (s1 = s2)) (decide (m1 = m2)) (decide (e1 = e2)) | ||
| 41 | | _, _ => right _ | ||
| 42 | end); abstract naive_solver (f_equal; apply (proof_irrel _)). | ||
| 43 | Defined. | ||
| 44 | |||
| 45 | Definition of_Z (x : Z) : float := | ||
| 46 | binary_normalize prec emax (refl_equal _) (refl_equal _) mode_NE x 0 false. | ||
| 47 | |||
| 48 | Definition to_Z (f : float) : option Z := | ||
| 49 | match f with | ||
| 50 | | B754_zero _ => Some 0 | ||
| 51 | | B754_finite s m e _ => Some (Zaux.cond_Zopp s (Zpos m) ≪ e) | ||
| 52 | | _ => None | ||
| 53 | end%Z. | ||
| 54 | |||
| 55 | (** QNaN Floating-Point Indefinite; see Table 4-3. Floating-Point Number and | ||
| 56 | NaN Encodings. *) | ||
| 57 | Definition indef_nan : { f | is_nan prec emax f = true } := | ||
| 58 | @B754_nan prec emax true (2^51) (refl_equal _) ↾ eq_refl _. | ||
| 59 | |||
| 60 | Definition to_flocq_round_mode (m : round_mode) : BinarySingleNaN.mode := | ||
| 61 | match m with Floor => mode_DN | Ceil => mode_UP | NearestEven => mode_NE end. | ||
| 62 | |||
| 63 | Definition round (m : round_mode) : float → float := | ||
| 64 | Bnearbyint prec emax (refl_equal _) (λ _, indef_nan) (to_flocq_round_mode m). | ||
| 65 | |||
| 66 | (* For add: not [mode_DN]; otherwise [+0.0 + -0.0] would yield [-0.0], but | ||
| 67 | [inf / (+0.0 + -0.0)] yields [inf] in C++, so this cannot be the case. *) | ||
| 68 | (* C++ compiles floating point addition to the x86 ADDSD instruction. Looking | ||
| 69 | at the Intel x86 Software Developer's Manual, it seems that the default | ||
| 70 | rounding mode on x86 is Round to Nearest (even); see table 4-8. (In §4.8.4.) *) | ||
| 71 | Definition add : float → float → float := | ||
| 72 | Bplus _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
| 73 | Definition sub : float → float → float := | ||
| 74 | Bminus _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
| 75 | Definition mul : float → float → float := | ||
| 76 | Bmult _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
| 77 | Definition div : float → float → float := | ||
| 78 | Bdiv _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
| 79 | |||
| 80 | Definition eqb (f1 f2 : float) : bool := | ||
| 81 | bool_decide (b64_compare f1 f2 = Some Eq). | ||
| 82 | |||
| 83 | Definition ltb (f1 f2 : float) : bool := | ||
| 84 | bool_decide (b64_compare f1 f2 = Some Lt). | ||
| 85 | End Float. | ||
diff --git a/theories/nix/interp.v b/theories/nix/interp.v new file mode 100644 index 0000000..bb4e815 --- /dev/null +++ b/theories/nix/interp.v | |||
| @@ -0,0 +1,351 @@ | |||
| 1 | From Coq Require Import Ascii. | ||
| 2 | From mininix Require Export res nix.operational_props. | ||
| 3 | From stdpp Require Import options. | ||
| 4 | |||
| 5 | Section val. | ||
| 6 | Local Unset Elimination Schemes. | ||
| 7 | Inductive val := | ||
| 8 | | VLit (bl : base_lit) (Hbl : base_lit_ok bl) | ||
| 9 | | VClo (x : string) (E : gmap string (kind * thunk)) (e : expr) | ||
| 10 | | VCloMatch (E : gmap string (kind * thunk)) | ||
| 11 | (ms : gmap string (option expr)) | ||
| 12 | (strict : bool) (e : expr) | ||
| 13 | | VList (ts : list thunk) | ||
| 14 | | VAttr (ts : gmap string thunk) | ||
| 15 | with thunk := | ||
| 16 | | Forced (v : val) : thunk | ||
| 17 | | Thunk (E : gmap string (kind * thunk)) (e : expr) : thunk | ||
| 18 | | Indirect (x : string) | ||
| 19 | (E : gmap string (kind * thunk)) | ||
| 20 | (tαs : gmap string (expr + thunk)). | ||
| 21 | End val. | ||
| 22 | |||
| 23 | Notation VLitI bl := (VLit bl I). | ||
| 24 | |||
| 25 | Notation tattr := (expr + thunk)%type. | ||
| 26 | Notation env := (gmap string (kind * thunk)). | ||
| 27 | |||
| 28 | Definition maybe_VLit (v : val) : option base_lit := | ||
| 29 | if v is VLit bl _ then Some bl else None. | ||
| 30 | Global Instance maybe_VList : Maybe VList := λ v, | ||
| 31 | if v is VList ts then Some ts else None. | ||
| 32 | Global Instance maybe_VAttr : Maybe VAttr := λ v, | ||
| 33 | if v is VAttr ts then Some ts else None. | ||
| 34 | |||
| 35 | Fixpoint interp_eq_list_body (i : nat) (ts1 ts2 : list thunk) : expr := | ||
| 36 | match ts1, ts2 with | ||
| 37 | | [], [] => ELit (LitBool true) | ||
| 38 | | _ :: ts1, _ :: ts2 => | ||
| 39 | EIf (EBinOp EqOp (EId' ("1" +:+ pretty i)) (EId' ("2" +:+ pretty i))) | ||
| 40 | (interp_eq_list_body (S i) ts1 ts2) (ELit (LitBool false)) | ||
| 41 | | _, _ => ELit (LitBool false) | ||
| 42 | end. | ||
| 43 | |||
| 44 | Definition interp_eq_list (ts1 ts2 : list thunk) : thunk := | ||
| 45 | Thunk (kmap (λ n : nat, String "1" (pretty n)) ((ABS,.) <$> map_seq 0 ts1) ∪ | ||
| 46 | kmap (λ n : nat, String "2" (pretty n)) ((ABS,.) <$> map_seq 0 ts2)) $ | ||
| 47 | interp_eq_list_body 0 ts1 ts2. | ||
| 48 | |||
| 49 | Fixpoint interp_lt_list_body (i : nat) (ts1 ts2 : list thunk) : expr := | ||
| 50 | match ts1, ts2 with | ||
| 51 | | [], _ => ELit (LitBool true) | ||
| 52 | | _ :: ts1, _ :: ts2 => | ||
| 53 | EIf (EBinOp LtOp (EId' ("1" +:+ pretty i)) (EId' ("2" +:+ pretty i))) | ||
| 54 | (ELit (LitBool true)) | ||
| 55 | (EIf (EBinOp EqOp (EId' ("1" +:+ pretty i)) (EId' ("2" +:+ pretty i))) | ||
| 56 | (interp_lt_list_body (S i) ts1 ts2) (ELit (LitBool false))) | ||
| 57 | | _ :: _, [] => ELit (LitBool false) | ||
| 58 | end. | ||
| 59 | |||
| 60 | Definition interp_lt_list (ts1 ts2 : list thunk) : thunk := | ||
| 61 | Thunk (kmap (λ n : nat, String "1" (pretty n)) ((ABS,.) <$> map_seq 0 ts1) ∪ | ||
| 62 | kmap (λ n : nat, String "2" (pretty n)) ((ABS,.) <$> map_seq 0 ts2)) $ | ||
| 63 | interp_lt_list_body 0 ts1 ts2. | ||
| 64 | |||
| 65 | Definition interp_eq_attr (ts1 ts2 : gmap string thunk) : thunk := | ||
| 66 | Thunk (kmap (String "1") ((ABS,.) <$> ts1) ∪ | ||
| 67 | kmap (String "2") ((ABS,.) <$> ts2)) $ | ||
| 68 | sem_and_attr $ map_imap (λ x _, | ||
| 69 | Some (EBinOp EqOp (EId' ("1" +:+ x)) (EId' ("2" +:+ x)))) ts1. | ||
| 70 | |||
| 71 | Definition interp_eq (v1 v2 : val) : option thunk := | ||
| 72 | match v1, v2 with | ||
| 73 | | VLit bl1 _, VLit bl2 _ => | ||
| 74 | Some $ Forced $ VLitI (LitBool $ sem_eq_base_lit bl1 bl2) | ||
| 75 | | VClo _ _ _, VClo _ _ _ => None | ||
| 76 | | VList ts1, VList ts2 => Some $ | ||
| 77 | if decide (length ts1 = length ts2) then interp_eq_list ts1 ts2 | ||
| 78 | else Forced $ VLitI (LitBool false) | ||
| 79 | | VAttr ts1, VAttr ts2 => Some $ | ||
| 80 | if decide (dom ts1 = dom ts2) then interp_eq_attr ts1 ts2 | ||
| 81 | else Forced $ VLitI (LitBool false) | ||
| 82 | | _, _ => Some $ Forced $ VLitI (LitBool false) | ||
| 83 | end. | ||
| 84 | |||
| 85 | Definition type_of_val (v : val) : string := | ||
| 86 | match v with | ||
| 87 | | VLit bl _ => type_of_base_lit bl | ||
| 88 | | VClo _ _ _ | VCloMatch _ _ _ _ => "lambda" | ||
| 89 | | VList _ => "list" | ||
| 90 | | VAttr _ => "set" | ||
| 91 | end. | ||
| 92 | |||
| 93 | Global Instance val_inhabited : Inhabited val := populate (VLitI inhabitant). | ||
| 94 | Global Instance thunk_inhabited : Inhabited thunk := populate (Forced inhabitant). | ||
| 95 | |||
| 96 | Definition interp_bin_op (op : bin_op) (v1 : val) : option (val → option thunk) := | ||
| 97 | if decide (op = EqOp) then | ||
| 98 | Some (interp_eq v1) | ||
| 99 | else if decide (op = TypeOfOp) then | ||
| 100 | Some $ λ v2, | ||
| 101 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 102 | Some $ Forced $ VLitI (LitString $ type_of_val v1) | ||
| 103 | else | ||
| 104 | match v1 with | ||
| 105 | | VLit (LitNum n1) Hn1 => | ||
| 106 | if maybe RoundOp op is Some m then | ||
| 107 | Some $ λ v2, | ||
| 108 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 109 | Some $ Forced $ VLit | ||
| 110 | (LitNum $ NInt $ float_to_bounded_Z $ Float.round m $ num_to_float n1) | ||
| 111 | (float_to_bounded_Z_ok _) | ||
| 112 | else | ||
| 113 | '(f ↾ Hf) ← option_to_eq_Some (sem_bin_op_num op n1); | ||
| 114 | Some $ λ v2, | ||
| 115 | if v2 is VLit (LitNum n2) Hn2 then | ||
| 116 | '(bl ↾ Hbl) ← option_to_eq_Some (f n2); | ||
| 117 | Some $ Forced $ VLit bl (sem_bin_op_num_ok Hn1 Hn2 Hf Hbl) | ||
| 118 | else None | ||
| 119 | | VLit (LitString s1) _ => | ||
| 120 | match op with | ||
| 121 | | SingletonAttrOp => Some $ λ v2, | ||
| 122 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 123 | Some $ Forced $ VClo "t" ∅ (EAttr {[ s1 := AttrN (EId' "t") ]}) | ||
| 124 | | MatchStringOp => Some $ λ v2, | ||
| 125 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 126 | match s1 with | ||
| 127 | | EmptyString => Some $ Forced $ VAttr {[ | ||
| 128 | "empty" := Forced (VLitI (LitBool true)); | ||
| 129 | "head" := Forced (VLitI LitNull); | ||
| 130 | "tail" := Forced (VLitI LitNull) ]} | ||
| 131 | | String a s1 => Some $ Forced $ VAttr {[ | ||
| 132 | "empty" := Forced (VLitI (LitBool false)); | ||
| 133 | "head" := Forced (VLitI (LitString (String a EmptyString))); | ||
| 134 | "tail" := Forced (VLitI (LitString s1)) ]} | ||
| 135 | end | ||
| 136 | | _ => | ||
| 137 | '(f ↾ Hf) ← option_to_eq_Some (sem_bin_op_string op); | ||
| 138 | Some $ λ v2, | ||
| 139 | bl2 ← maybe_VLit v2; | ||
| 140 | s2 ← maybe LitString bl2; | ||
| 141 | Some $ Forced $ VLit (f s1 s2) (sem_bin_op_string_ok Hf) | ||
| 142 | end | ||
| 143 | | VClo _ _ _ => | ||
| 144 | match op with | ||
| 145 | | FunctionArgsOp => Some $ λ v2, | ||
| 146 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 147 | Some (Forced (VAttr ∅)) | ||
| 148 | | _ => None | ||
| 149 | end | ||
| 150 | | VCloMatch _ ms _ _ => | ||
| 151 | match op with | ||
| 152 | | FunctionArgsOp => Some $ λ v2, | ||
| 153 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 154 | Some $ Forced $ VAttr $ | ||
| 155 | (λ m, Forced $ VLitI (LitBool (from_option (λ _, true) false m))) <$> ms | ||
| 156 | | _ => None | ||
| 157 | end | ||
| 158 | | VList ts1 => | ||
| 159 | match op with | ||
| 160 | | LtOp => Some $ λ v2, | ||
| 161 | ts2 ← maybe VList v2; | ||
| 162 | Some (interp_lt_list ts1 ts2) | ||
| 163 | | MatchListOp => Some $ λ v2, | ||
| 164 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 165 | match ts1 with | ||
| 166 | | [] => Some $ Forced $ VAttr {[ | ||
| 167 | "empty" := Forced (VLitI (LitBool true)); | ||
| 168 | "head" := Forced (VLitI LitNull); | ||
| 169 | "tail" := Forced (VLitI LitNull) ]} | ||
| 170 | | t :: ts1 => Some $ Forced $ VAttr {[ | ||
| 171 | "empty" := Forced (VLitI (LitBool false)); | ||
| 172 | "head" := t; | ||
| 173 | "tail" := Forced (VList ts1) ]} | ||
| 174 | end | ||
| 175 | | AppendListOp => Some $ λ v2, | ||
| 176 | ts2 ← maybe VList v2; | ||
| 177 | Some (Forced (VList (ts1 ++ ts2))) | ||
| 178 | | _ => None | ||
| 179 | end | ||
| 180 | | VAttr ts1 => | ||
| 181 | match op with | ||
| 182 | | SelectAttrOp => Some $ λ v2, | ||
| 183 | bl ← maybe_VLit v2; | ||
| 184 | x ← maybe LitString bl; | ||
| 185 | ts1 !! x | ||
| 186 | | UpdateAttrOp => Some $ λ v2, | ||
| 187 | ts2 ← maybe VAttr v2; | ||
| 188 | Some $ Forced $ VAttr $ ts2 ∪ ts1 | ||
| 189 | | HasAttrOp => Some $ λ v2, | ||
| 190 | bl ← maybe_VLit v2; | ||
| 191 | x ← maybe LitString bl; | ||
| 192 | Some $ Forced $ VLitI (LitBool $ bool_decide (is_Some (ts1 !! x))) | ||
| 193 | | DeleteAttrOp => Some $ λ v2, | ||
| 194 | bl ← maybe_VLit v2; | ||
| 195 | x ← maybe LitString bl; | ||
| 196 | Some $ Forced $ VAttr $ delete x ts1 | ||
| 197 | | MatchAttrOp => Some $ λ v2, | ||
| 198 | guard (maybe_VLit v2 = Some LitNull);; | ||
| 199 | match map_minimal_key attr_le ts1 with | ||
| 200 | | None => Some $ Forced $ VAttr {[ | ||
| 201 | "empty" := Forced (VLitI (LitBool true)); | ||
| 202 | "key" := Forced (VLitI LitNull); | ||
| 203 | "head" := Forced (VLitI LitNull); | ||
| 204 | "tail" := Forced (VLitI LitNull) ]} | ||
| 205 | | Some x => Some $ Forced $ VAttr {[ | ||
| 206 | "empty" := Forced (VLitI (LitBool false)); | ||
| 207 | "key" := Forced (VLitI (LitString x)); | ||
| 208 | "head" := ts1 !!! x; | ||
| 209 | "tail" := Forced (VAttr (delete x ts1)) ]} | ||
| 210 | end | ||
| 211 | | _ => None | ||
| 212 | end | ||
| 213 | | _ => None | ||
| 214 | end. | ||
| 215 | |||
| 216 | Definition interp_match | ||
| 217 | (ts : gmap string thunk) (mds : gmap string (option expr)) | ||
| 218 | (strict : bool) : option (gmap string tattr) := | ||
| 219 | map_mapM id $ merge (λ mt mmd, | ||
| 220 | (* Some (Some _) means keep, Some None means fail, None means skip *) | ||
| 221 | match mt, mmd with | ||
| 222 | | Some t, Some _ => Some $ Some (inr t) | ||
| 223 | | None, Some (Some e) => Some $ Some (inl e) | ||
| 224 | | None, Some _ => Some None (* bad *) | ||
| 225 | | Some _, None => guard strict;; Some None | ||
| 226 | | _, _ => None (* skip *) | ||
| 227 | end) ts mds. | ||
| 228 | |||
| 229 | Definition force_deep1 | ||
| 230 | (force_deep : val → res val) | ||
| 231 | (interp_thunk : thunk → res val) (v : val) : res val := | ||
| 232 | match v with | ||
| 233 | | VList ts => VList ∘ fmap Forced <$> | ||
| 234 | mapM (mbind force_deep ∘ interp_thunk) ts | ||
| 235 | | VAttr ts => VAttr ∘ fmap Forced <$> | ||
| 236 | map_mapM_sorted attr_le (mbind force_deep ∘ interp_thunk) ts | ||
| 237 | | _ => mret v | ||
| 238 | end. | ||
| 239 | |||
| 240 | Definition indirects_env (E : env) (tαs : gmap string tattr) : env := | ||
| 241 | map_imap (λ y _, Some (ABS, Indirect y E tαs)) tαs ∪ E. | ||
| 242 | |||
| 243 | Definition attr_to_tattr (E : env) (α : attr) : tattr := | ||
| 244 | from_attr inl (inr ∘ Thunk E) α. | ||
| 245 | |||
| 246 | Definition interp1 | ||
| 247 | (force_deep : val → res val) | ||
| 248 | (interp : env → expr → res val) | ||
| 249 | (interp_thunk : thunk → res val) | ||
| 250 | (interp_app : val → thunk → res val) | ||
| 251 | (E : env) (e : expr) : res val := | ||
| 252 | match e with | ||
| 253 | | ELit bl => | ||
| 254 | bl_ok ← guard (base_lit_ok bl); | ||
| 255 | mret (VLit bl bl_ok) | ||
| 256 | | EId x mke => | ||
| 257 | '(_,t) ← Res $ union_kinded (E !! x) (prod_map id (Thunk ∅) <$> mke); | ||
| 258 | interp_thunk t | ||
| 259 | | EAbs x e => mret (VClo x E e) | ||
| 260 | | EAbsMatch ms strict e => mret (VCloMatch E ms strict e) | ||
| 261 | | EApp e1 e2 => | ||
| 262 | v1 ← interp E e1; | ||
| 263 | interp_app v1 (Thunk E e2) | ||
| 264 | | ESeq μ' e1 e2 => | ||
| 265 | v ← interp E e1; | ||
| 266 | (if μ' is DEEP then force_deep else mret) v;; | ||
| 267 | interp E e2 | ||
| 268 | | EList es => mret (VList (Thunk E <$> es)) | ||
| 269 | | EAttr αs => | ||
| 270 | let E' := indirects_env E (attr_to_tattr E <$> αs) in | ||
| 271 | mret (VAttr (from_attr (Thunk E') (Thunk E) <$> αs)) | ||
| 272 | | ELetAttr k e1 e2 => | ||
| 273 | v1 ← interp E e1; | ||
| 274 | ts ← Res (maybe VAttr v1); | ||
| 275 | interp (union_kinded ((k,.) <$> ts) E) e2 | ||
| 276 | | EBinOp op e1 e2 => | ||
| 277 | v1 ← interp E e1; | ||
| 278 | f ← Res (interp_bin_op op v1); | ||
| 279 | v2 ← interp E e2; | ||
| 280 | t2 ← Res (f v2); | ||
| 281 | interp_thunk t2 | ||
| 282 | | EIf e1 e2 e3 => | ||
| 283 | v1 ← interp E e1; | ||
| 284 | '(b : bool) ← Res (maybe_VLit v1 ≫= maybe LitBool); | ||
| 285 | interp E (if b then e2 else e3) | ||
| 286 | end. | ||
| 287 | |||
| 288 | Definition interp_thunk1 | ||
| 289 | (interp : env → expr → res val) | ||
| 290 | (interp_thunk : thunk → res val) | ||
| 291 | (t : thunk) : res val := | ||
| 292 | match t with | ||
| 293 | | Forced v => mret v | ||
| 294 | | Thunk E e => interp E e | ||
| 295 | | Indirect x E tαs => | ||
| 296 | tα ← Res $ tαs !! x; | ||
| 297 | match tα with | ||
| 298 | | inl e => interp (indirects_env E tαs) e | ||
| 299 | | inr t => interp_thunk t | ||
| 300 | end | ||
| 301 | end. | ||
| 302 | |||
| 303 | Definition interp_app1 | ||
| 304 | (interp : env → expr → res val) | ||
| 305 | (interp_thunk : thunk → res val) | ||
| 306 | (interp_app : val → thunk → res val) | ||
| 307 | (v1 : val) (t2 : thunk) : res val := | ||
| 308 | match v1 with | ||
| 309 | | VClo x E e => | ||
| 310 | interp (<[x:=(ABS, t2)]> E) e | ||
| 311 | | VCloMatch E ms strict e => | ||
| 312 | vt ← interp_thunk t2; | ||
| 313 | ts ← Res (maybe VAttr vt); | ||
| 314 | tαs ← Res $ interp_match ts ms strict; | ||
| 315 | interp (indirects_env E tαs) e | ||
| 316 | | VAttr ts => | ||
| 317 | t ← Res (ts !! "__functor"); | ||
| 318 | vt ← interp_thunk t; | ||
| 319 | v ← interp_app vt (Forced v1); | ||
| 320 | interp_app v t2 | ||
| 321 | | _ => mfail | ||
| 322 | end. | ||
| 323 | |||
| 324 | Fixpoint force_deep (n : nat) (v : val) : res val := | ||
| 325 | match n with | ||
| 326 | | O => NoFuel | ||
| 327 | | S n => force_deep1 (force_deep n) (interp_thunk n) v | ||
| 328 | end | ||
| 329 | with interp (n : nat) (E : env) (e : expr) : res val := | ||
| 330 | match n with | ||
| 331 | | O => NoFuel | ||
| 332 | | S n => interp1 (force_deep n) (interp n) (interp_thunk n) (interp_app n) E e | ||
| 333 | end | ||
| 334 | with interp_thunk (n : nat) (t : thunk) : res val := | ||
| 335 | match n with | ||
| 336 | | O => NoFuel | ||
| 337 | | S n => interp_thunk1 (interp n) (interp_thunk n) t | ||
| 338 | end | ||
| 339 | with interp_app (n : nat) (v1 : val) (t2 : thunk) : res val := | ||
| 340 | match n with | ||
| 341 | | O => NoFuel | ||
| 342 | | S n => interp_app1 (interp n) (interp_thunk n) (interp_app n) v1 t2 | ||
| 343 | end. | ||
| 344 | |||
| 345 | Definition force_deep' (n : nat) (μ : mode) : val → res val := | ||
| 346 | match μ with SHALLOW => mret | DEEP => force_deep n end. | ||
| 347 | |||
| 348 | Definition interp' (n : nat) (μ : mode) (E : env) (e : expr) : res val := | ||
| 349 | interp n E e ≫= force_deep' n μ. | ||
| 350 | |||
| 351 | Global Opaque force_deep interp interp_thunk interp_app. | ||
diff --git a/theories/nix/interp_proofs.v b/theories/nix/interp_proofs.v new file mode 100644 index 0000000..5780e48 --- /dev/null +++ b/theories/nix/interp_proofs.v | |||
| @@ -0,0 +1,2690 @@ | |||
| 1 | From Coq Require Import Ascii. | ||
| 2 | From mininix Require Export nix.interp. | ||
| 3 | From stdpp Require Import options. | ||
| 4 | |||
| 5 | Lemma force_deep_S n : | ||
| 6 | force_deep (S n) = force_deep1 (force_deep n) (interp_thunk n). | ||
| 7 | Proof. done. Qed. | ||
| 8 | Lemma interp_S n : | ||
| 9 | interp (S n) = interp1 (force_deep n) (interp n) (interp_thunk n) (interp_app n). | ||
| 10 | Proof. done. Qed. | ||
| 11 | Lemma interp_thunk_S n : | ||
| 12 | interp_thunk (S n) = interp_thunk1 (interp n) (interp_thunk n). | ||
| 13 | Proof. done. Qed. | ||
| 14 | Lemma interp_app_S n : | ||
| 15 | interp_app (S n) = interp_app1 (interp n) (interp_thunk n) (interp_app n). | ||
| 16 | Proof. done. Qed. | ||
| 17 | |||
| 18 | Lemma interp_shallow' m E e : interp' m SHALLOW E e = interp m E e. | ||
| 19 | Proof. rewrite /interp'. by destruct (interp _ _ _) as [[]|]. Qed. | ||
| 20 | |||
| 21 | Lemma interp_lit n E bl (Hbl : base_lit_ok bl) : | ||
| 22 | interp (S n) E (ELit bl) = mret (VLit bl Hbl). | ||
| 23 | Proof. | ||
| 24 | rewrite interp_S /=. case_guard; simpl; [|done]. | ||
| 25 | do 2 f_equal. apply (proof_irrel _). | ||
| 26 | Qed. | ||
| 27 | |||
| 28 | (** Induction *) | ||
| 29 | Fixpoint val_size (v : val) : nat := | ||
| 30 | match v with | ||
| 31 | | VLit _ _ => 1 | ||
| 32 | | VClo _ E _ | VCloMatch E _ _ _ => S (map_sum_with (thunk_size ∘ snd) E) | ||
| 33 | | VList ts => S (sum_list_with thunk_size ts) | ||
| 34 | | VAttr ts => S (map_sum_with thunk_size ts) | ||
| 35 | end | ||
| 36 | with thunk_size (t : thunk) : nat := | ||
| 37 | match t with | ||
| 38 | | Forced v => S (val_size v) | ||
| 39 | | Thunk E _ => S (map_sum_with (thunk_size ∘ snd) E) | ||
| 40 | | Indirect _ E tαs => S (map_sum_with (thunk_size ∘ snd) E + | ||
| 41 | map_sum_with (from_sum (λ _, 1) thunk_size) tαs) | ||
| 42 | end. | ||
| 43 | Notation env_size := (map_sum_with (thunk_size ∘ snd)). | ||
| 44 | |||
| 45 | Definition from_thunk {A} (f : val → A) (g : env → expr → A) | ||
| 46 | (h : string → env → gmap string tattr → A) (t : thunk) : A := | ||
| 47 | match t with | ||
| 48 | | Forced v => f v | ||
| 49 | | Thunk E e => g E e | ||
| 50 | | Indirect x E tαs => h x E tαs | ||
| 51 | end. | ||
| 52 | |||
| 53 | Lemma env_val_ind (P : env → Prop) (Q : val → Prop) : | ||
| 54 | (∀ E, | ||
| 55 | map_Forall (λ _, from_thunk Q (λ E _, P E) (λ _ E _, P E) ∘ snd) E → P E) → | ||
| 56 | (∀ b Hbl, Q (VLit b Hbl)) → | ||
| 57 | (∀ x E e, P E → Q (VClo x E e)) → | ||
| 58 | (∀ E ms strict e, P E → Q (VCloMatch E ms strict e)) → | ||
| 59 | (∀ ts, Forall (from_thunk Q (λ E _, P E) (λ _ E _, P E)) ts → Q (VList ts)) → | ||
| 60 | (∀ ts, map_Forall (λ _, from_thunk Q (λ E _, P E) (λ _ E _, P E)) ts → Q (VAttr ts)) → | ||
| 61 | (∀ E, P E) ∧ (∀ v, Q v). | ||
| 62 | Proof. | ||
| 63 | intros Penv Qlit Qclo Qmatch Qlist Qattr. | ||
| 64 | cut (∀ n, (∀ E, env_size E < n → P E) ∧ (∀ v, val_size v < n → Q v)). | ||
| 65 | { intros Hhelp; split. | ||
| 66 | - intros E. apply (Hhelp (S (env_size E))); lia. | ||
| 67 | - intros v. apply (Hhelp (S (val_size v))); lia. } | ||
| 68 | intros n. induction n as [|n IH]; [by auto with lia|]. split. | ||
| 69 | - intros E ?. apply Penv, map_Forall_lookup=> y [k ei] Hy. | ||
| 70 | apply (map_sum_with_lookup_le (thunk_size ∘ snd)) in Hy; simpl in *. | ||
| 71 | destruct ei as [v|E' e'|x E' tαs]; simplify_eq/=; try apply IH; eauto with lia. | ||
| 72 | - intros [] ?; simpl in *. | ||
| 73 | + apply Qlit. | ||
| 74 | + apply Qclo, IH. lia. | ||
| 75 | + apply Qmatch, IH. lia. | ||
| 76 | + apply Qlist, Forall_forall=> t Hy. | ||
| 77 | apply (sum_list_with_in _ thunk_size) in Hy. | ||
| 78 | destruct t; simpl in *; try apply IH; lia. | ||
| 79 | + apply Qattr, map_Forall_lookup=> y t Hy. | ||
| 80 | apply (map_sum_with_lookup_le thunk_size) in Hy. | ||
| 81 | destruct t; simpl in *; try apply IH; lia. | ||
| 82 | Qed. | ||
| 83 | |||
| 84 | Lemma env_ind (P : env → Prop) : | ||
| 85 | (∀ E, | ||
| 86 | map_Forall (λ i, from_thunk (λ _, True) (λ E _, P E) (λ _ E _, P E) ∘ snd) E → | ||
| 87 | P E) → | ||
| 88 | ∀ E : env, P E. | ||
| 89 | Proof. intros. apply (env_val_ind P (λ _, True)); auto. Qed. | ||
| 90 | |||
| 91 | Lemma val_ind (Q : val → Prop) : | ||
| 92 | (∀ bl Hbl, Q (VLit bl Hbl)) → | ||
| 93 | (∀ x E e, Q (VClo x E e)) → | ||
| 94 | (∀ ms strict E e, Q (VCloMatch ms strict E e)) → | ||
| 95 | (∀ ts, Forall (from_thunk Q (λ _ _, True) (λ _ _ _, True)) ts → Q (VList ts)) → | ||
| 96 | (∀ ts, | ||
| 97 | map_Forall (λ _, from_thunk Q (λ _ _, True) (λ _ _ _, True)) ts → Q (VAttr ts)) → | ||
| 98 | (∀ v, Q v). | ||
| 99 | Proof. intros. apply (env_val_ind (λ _, True) Q); auto. Qed. | ||
| 100 | (** Correspondence to operational semantics *) | ||
| 101 | Definition subst_env' (thunk_to_expr : thunk → expr) (E : env) : expr → expr := | ||
| 102 | subst (prod_map id thunk_to_expr <$> E). | ||
| 103 | |||
| 104 | Definition tattr_to_attr' | ||
| 105 | (thunk_to_expr : thunk → expr) (subst_env : env → expr → expr) | ||
| 106 | (E : env) (α : tattr) : attr := | ||
| 107 | from_sum (AttrR ∘ subst_env E) (AttrN ∘ thunk_to_expr) α. | ||
| 108 | |||
| 109 | Fixpoint thunk_to_expr (t : thunk) : expr := | ||
| 110 | match t with | ||
| 111 | | Forced v => val_to_expr v | ||
| 112 | | Thunk E e => subst_env' thunk_to_expr E e | ||
| 113 | | Indirect x E tαs => ESelect | ||
| 114 | (EAttr (tattr_to_attr' thunk_to_expr (subst_env' thunk_to_expr) E <$> tαs)) x | ||
| 115 | end | ||
| 116 | with val_to_expr (v : val) : expr := | ||
| 117 | match v with | ||
| 118 | | VLit bl _ => ELit bl | ||
| 119 | | VClo x E e => EAbs x (subst_env' thunk_to_expr E e) | ||
| 120 | | VCloMatch E ms strict e => EAbsMatch | ||
| 121 | (fmap (M:=option) (subst_env' thunk_to_expr E) <$> ms) | ||
| 122 | strict | ||
| 123 | (subst_env' thunk_to_expr E e) | ||
| 124 | | VList ts => EList (thunk_to_expr <$> ts) | ||
| 125 | | VAttr ts => EAttr (AttrN ∘ thunk_to_expr <$> ts) | ||
| 126 | end. | ||
| 127 | |||
| 128 | Notation subst_env := (subst_env' thunk_to_expr). | ||
| 129 | Notation tattr_to_attr := (tattr_to_attr' thunk_to_expr subst_env). | ||
| 130 | Notation attr_subst_env E := (attr_map (subst_env E)). | ||
| 131 | |||
| 132 | Lemma subst_env_eq e E : | ||
| 133 | subst_env E e = | ||
| 134 | match e with | ||
| 135 | | ELit n => ELit n | ||
| 136 | | EId x mkd => EId x $ | ||
| 137 | union_kinded (prod_map id thunk_to_expr <$> E !! x) mkd | ||
| 138 | | EAbs x e => EAbs x (subst_env E e) | ||
| 139 | | EAbsMatch ms strict e => | ||
| 140 | EAbsMatch (fmap (M:=option) (subst_env E) <$> ms) strict (subst_env E e) | ||
| 141 | | EApp e1 e2 => EApp (subst_env E e1) (subst_env E e2) | ||
| 142 | | ESeq μ e1 e2 => ESeq μ (subst_env E e1) (subst_env E e2) | ||
| 143 | | EList es => EList (subst_env E <$> es) | ||
| 144 | | EAttr αs => EAttr (attr_subst_env E <$> αs) | ||
| 145 | | ELetAttr k e1 e2 => ELetAttr k (subst_env E e1) (subst_env E e2) | ||
| 146 | | EBinOp op e1 e2 => EBinOp op (subst_env E e1) (subst_env E e2) | ||
| 147 | | EIf e1 e2 e3 => EIf (subst_env E e1) (subst_env E e2) (subst_env E e3) | ||
| 148 | end. | ||
| 149 | Proof. rewrite /subst_env. destruct e; by rewrite /= ?lookup_fmap. Qed. | ||
| 150 | |||
| 151 | Lemma subst_env_alt E e : subst_env E e = subst (prod_map id thunk_to_expr <$> E) e. | ||
| 152 | Proof. done. Qed. | ||
| 153 | |||
| 154 | (* Use the unfolding lemmas, don't rely on conversion *) | ||
| 155 | Opaque subst_env'. | ||
| 156 | |||
| 157 | Lemma subst_env_empty e : subst_env ∅ e = e. | ||
| 158 | Proof. rewrite subst_env_alt. apply subst_empty. Qed. | ||
| 159 | |||
| 160 | Lemma final_val_to_expr v : final SHALLOW (val_to_expr v). | ||
| 161 | Proof. induction v; simpl; constructor; auto. Qed. | ||
| 162 | Local Hint Resolve final_val_to_expr | 0 : core. | ||
| 163 | Lemma step_not_val_to_expr v e : val_to_expr v -{SHALLOW}-> e → False. | ||
| 164 | Proof. intros []%step_not_final. done. Qed. | ||
| 165 | |||
| 166 | Lemma final_force_deep n t v : | ||
| 167 | force_deep n t = mret v → final DEEP (val_to_expr v). | ||
| 168 | Proof. | ||
| 169 | revert t v. induction n as [|n IH]; intros v w; [done|]. | ||
| 170 | rewrite force_deep_S /=. | ||
| 171 | intros; destruct v; simplify_res; eauto using final. | ||
| 172 | + destruct (mapM _ _) as [[vs|]|] eqn:Hmap; simplify_res; eauto. | ||
| 173 | constructor. revert vs Hmap. | ||
| 174 | induction ts as [|t ts IHts]; intros; simplify_res; [by constructor|..]. | ||
| 175 | destruct (interp_thunk _ _) as [[w|]|]; simplify_res. | ||
| 176 | destruct (force_deep _ _) as [[w'|]|] eqn:?; simplify_res. | ||
| 177 | destruct (mapM _ _) as [[ws|]|]; simplify_res; eauto. | ||
| 178 | + destruct (map_mapM_sorted _ _ _) as [[vs|]|] eqn:Hmap; simplify_res. | ||
| 179 | constructor; [done|]. | ||
| 180 | revert vs Hmap. induction ts as [|x t ts ?? IHts] | ||
| 181 | using (map_sorted_ind attr_le); intros ts' Hts. | ||
| 182 | { rewrite map_mapM_sorted_empty in Hts; simplify_res. done. } | ||
| 183 | rewrite map_mapM_sorted_insert //= in Hts. | ||
| 184 | destruct (interp_thunk _ _) as [[w|]|] eqn:?; simplify_res. | ||
| 185 | destruct (force_deep _ _) as [[w'|]|] eqn:?; simplify_res. | ||
| 186 | destruct (map_mapM_sorted _ _ _) as [[ws|]|] eqn:Hmap; simplify_res. | ||
| 187 | rewrite !fmap_insert /=. apply map_Forall_insert_2, IHts; eauto. | ||
| 188 | Qed. | ||
| 189 | |||
| 190 | Lemma interp_bin_op_Some_1 op v1 f : | ||
| 191 | interp_bin_op op v1 = Some f → | ||
| 192 | ∃ Φ, sem_bin_op op (val_to_expr v1) Φ. | ||
| 193 | Proof. | ||
| 194 | intros Hinterp. unfold interp_bin_op, interp_eq in *. | ||
| 195 | repeat (case_match || simplify_option_eq); | ||
| 196 | eexists; by repeat econstructor; eauto using final. | ||
| 197 | Qed. | ||
| 198 | |||
| 199 | Lemma interp_bin_op_Some_2 op v1 Φ : | ||
| 200 | sem_bin_op op (val_to_expr v1) Φ → | ||
| 201 | is_Some (interp_bin_op op v1). | ||
| 202 | Proof. | ||
| 203 | unfold interp_bin_op; destruct v1; inv 1; | ||
| 204 | repeat (case_match || simplify_option_eq); eauto. | ||
| 205 | destruct (option_to_eq_Some _) as [[??]|] eqn:Ho; simplify_eq/=; eauto. | ||
| 206 | by rewrite H2 in Ho. | ||
| 207 | Qed. | ||
| 208 | |||
| 209 | Lemma interp_eq_list_correct ts1 ts2 : | ||
| 210 | thunk_to_expr (interp_eq_list ts1 ts2) =D=> | ||
| 211 | sem_eq_list (thunk_to_expr <$> ts1) (thunk_to_expr <$> ts2). | ||
| 212 | Proof. | ||
| 213 | simpl. | ||
| 214 | remember (kmap (λ n : nat, String "1" (pretty n)) ((ABS,.) <$> map_seq 0 ts1) ∪ | ||
| 215 | kmap (λ n : nat, String "2" (pretty n)) ((ABS,.) <$> map_seq 0 ts2)) | ||
| 216 | as E eqn:HE. | ||
| 217 | assert (∀ (i : nat) t, ts1 !! i = Some t → | ||
| 218 | E !! String "1" (pretty (i + 0)) = Some (ABS, t)) as Hts1. | ||
| 219 | { intros x t Hxt. rewrite Nat.add_0_r. | ||
| 220 | rewrite HE lookup_union (lookup_kmap _) lookup_fmap. | ||
| 221 | rewrite lookup_map_seq_0 Hxt /= union_Some_l. done. } | ||
| 222 | assert (∀ (i : nat) t, ts2 !! i = Some t → | ||
| 223 | E !! String "2" (pretty (i + 0)) = Some (ABS, t)) as Hts2. | ||
| 224 | { intros x t Hxt. rewrite Nat.add_0_r. | ||
| 225 | rewrite HE lookup_union_r; last by apply (lookup_kmap_None _). | ||
| 226 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0 Hxt /=. done. } | ||
| 227 | clear HE. revert ts2 Hts1 Hts2. generalize 0. | ||
| 228 | induction ts1 as [|t1 ts1 IH]; intros n [|t2 ts2] Hts1 Hts2; csimpl; [done..|]. | ||
| 229 | rewrite 4!subst_env_eq /= !(subst_env_eq (ELit _)) /=. rewrite /String.app. | ||
| 230 | rewrite (Hts1 0 t1) // (Hts2 0 t2) //=. | ||
| 231 | constructor; [repeat constructor| |done]. | ||
| 232 | apply IH; intros i t; rewrite Nat.add_succ_r; | ||
| 233 | [apply (Hts1 (S i))|apply (Hts2 (S i))]. | ||
| 234 | Qed. | ||
| 235 | |||
| 236 | Lemma interp_lt_list_correct ts1 ts2 : | ||
| 237 | thunk_to_expr (interp_lt_list ts1 ts2) =D=> | ||
| 238 | sem_lt_list (thunk_to_expr <$> ts1) (thunk_to_expr <$> ts2). | ||
| 239 | Proof. | ||
| 240 | simpl. | ||
| 241 | remember (kmap (λ n : nat, String "1" (pretty n)) ((ABS,.) <$> map_seq 0 ts1) ∪ | ||
| 242 | kmap (λ n : nat, String "2" (pretty n)) ((ABS,.) <$> map_seq 0 ts2)) | ||
| 243 | as E eqn:HE. | ||
| 244 | assert (∀ (i : nat) t, ts1 !! i = Some t → | ||
| 245 | E !! String "1" (pretty (i + 0)) = Some (ABS, t)) as Hts1. | ||
| 246 | { intros x t Hxt. rewrite Nat.add_0_r. | ||
| 247 | rewrite HE lookup_union (lookup_kmap _) lookup_fmap. | ||
| 248 | rewrite lookup_map_seq_0 Hxt /= union_Some_l. done. } | ||
| 249 | assert (∀ (i : nat) t, ts2 !! i = Some t → | ||
| 250 | E !! String "2" (pretty (i + 0)) = Some (ABS, t)) as Hts2. | ||
| 251 | { intros x t Hxt. rewrite Nat.add_0_r. | ||
| 252 | rewrite HE lookup_union_r; last by apply (lookup_kmap_None _). | ||
| 253 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0 Hxt /=. done. } | ||
| 254 | clear HE. revert ts2 Hts1 Hts2. generalize 0. | ||
| 255 | induction ts1 as [|t1 ts1 IH]; intros n [|t2 ts2] Hts1 Hts2; csimpl; [done..|]. | ||
| 256 | rewrite 4!subst_env_eq /= !(subst_env_eq (ELit _)) /=. rewrite /String.app. | ||
| 257 | rewrite (Hts1 0 t1) // (Hts2 0 t2) //=. | ||
| 258 | constructor; [repeat constructor..|]. | ||
| 259 | rewrite 4!subst_env_eq /= !(subst_env_eq (ELit _)) /=. | ||
| 260 | rewrite (Hts1 0 t1) // (Hts2 0 t2) //=. | ||
| 261 | constructor; [repeat constructor| |done]. | ||
| 262 | apply IH; intros i t; rewrite Nat.add_succ_r; | ||
| 263 | [apply (Hts1 (S i))|apply (Hts2 (S i))]. | ||
| 264 | Qed. | ||
| 265 | |||
| 266 | Lemma interp_eq_attr_correct ts1 ts2 : | ||
| 267 | dom ts1 = dom ts2 → | ||
| 268 | thunk_to_expr (interp_eq_attr ts1 ts2) =D=> | ||
| 269 | sem_eq_attr (AttrN ∘ thunk_to_expr <$> ts1) (AttrN ∘ thunk_to_expr <$> ts2). | ||
| 270 | Proof. | ||
| 271 | intros Hdom. simpl. | ||
| 272 | remember (kmap (String "1") ((ABS,.) <$> ts1) ∪ | ||
| 273 | kmap (String "2") ((ABS,.) <$> ts2)) as E eqn:HE. | ||
| 274 | assert (map_Forall (λ x t, E !! String "1" x = Some (ABS, t)) ts1) as Hts1. | ||
| 275 | { intros x t Hxt. | ||
| 276 | rewrite HE lookup_union (lookup_kmap (String "1")) lookup_fmap. | ||
| 277 | by rewrite Hxt /= union_Some_l. } | ||
| 278 | assert (map_Forall (λ x t, E !! String "2" x = Some (ABS, t)) ts2) as Hts2. | ||
| 279 | { intros x t Hxt. | ||
| 280 | rewrite HE lookup_union_r; last by apply (lookup_kmap_None _). | ||
| 281 | by rewrite (lookup_kmap (String "2")) lookup_fmap Hxt. } | ||
| 282 | clear HE. revert ts2 Hdom Hts1 Hts2. | ||
| 283 | induction ts1 as [|x t1 ts1 Hts1x IH] using (map_sorted_ind attr_le); | ||
| 284 | intros ts2 Hdom Hts1 Hts2. | ||
| 285 | { apply symmetry, dom_empty_inv_L in Hdom as ->. done. } | ||
| 286 | rewrite dom_insert_L in Hdom. | ||
| 287 | assert (is_Some (ts2 !! x)) as [t2 Hxt2] by (apply elem_of_dom; set_solver). | ||
| 288 | assert (dom ts1 = dom (delete x ts2)). | ||
| 289 | { rewrite dom_delete_L -Hdom. apply not_elem_of_dom in Hts1x. set_solver. } | ||
| 290 | rewrite -(insert_delete ts2 x t2) //. rewrite -(insert_delete ts2 x t2) // in Hts2. | ||
| 291 | apply map_Forall_insert in Hts1 as [Hx1 Hts1]; last done. | ||
| 292 | apply map_Forall_insert in Hts2 as [Hx2 Hts2]; last by rewrite lookup_delete. | ||
| 293 | rewrite /sem_eq_attr !fmap_insert /=. erewrite <-insert_merge by done. | ||
| 294 | rewrite sem_and_attr_insert; first last. | ||
| 295 | { intros y. rewrite lookup_merge !lookup_fmap /is_Some. | ||
| 296 | destruct (ts1 !! y) eqn:? , (delete x ts2 !! y); naive_solver. } | ||
| 297 | { rewrite lookup_merge !lookup_fmap lookup_delete /=. by destruct (ts1 !! x). } | ||
| 298 | rewrite map_imap_insert sem_and_attr_insert; first last. | ||
| 299 | { intros y. rewrite map_lookup_imap /is_Some. | ||
| 300 | destruct (ts1 !! y) eqn:?; naive_solver. } | ||
| 301 | { by rewrite map_lookup_imap Hts1x. } | ||
| 302 | rewrite 4!subst_env_eq /= !(subst_env_eq (ELit _)) /= Hx1 Hx2 /=. | ||
| 303 | constructor; [|apply IHts1; by auto|done]. by do 2 constructor. | ||
| 304 | Qed. | ||
| 305 | |||
| 306 | Lemma interp_bin_op_Some_Some_1 op v1 f Φ v2 t3 : | ||
| 307 | interp_bin_op op v1 = Some f → | ||
| 308 | sem_bin_op op (val_to_expr v1) Φ → | ||
| 309 | f v2 = Some t3 → | ||
| 310 | ∃ e3, Φ (val_to_expr v2) e3 ∧ thunk_to_expr t3 =D=> e3. | ||
| 311 | Proof. | ||
| 312 | unfold interp_bin_op, interp_eq, type_of_val, type_of_expr; | ||
| 313 | destruct v1, v2; inv 2; intros; | ||
| 314 | repeat match goal with | ||
| 315 | | _ => progress simplify_option_eq | ||
| 316 | | H : _ <$> _ = ∅ |- _ => apply fmap_empty_inv in H | ||
| 317 | | H : context [dom (_ <$> _)] |- _ => rewrite !dom_fmap_L in H | ||
| 318 | | H : context [length (_ <$> _)] |- _ => rewrite !length_fmap in H | ||
| 319 | | _ => case_match | ||
| 320 | | _ => eexists; split; [done|] | ||
| 321 | | _ => by apply interp_eq_list_correct | ||
| 322 | | _ => eexists; split; [|by apply: interp_lt_list_correct] | ||
| 323 | | _ => by apply: interp_eq_attr_correct | ||
| 324 | | _ => eexists; split; [|done] | ||
| 325 | | _ => split; [|done] | ||
| 326 | | _ => rewrite map_fmap_singleton | ||
| 327 | | _ => rewrite fmap_delete | ||
| 328 | | _ => rewrite subst_env_empty | ||
| 329 | | _ => rewrite fmap_app | ||
| 330 | | _ => rewrite lookup_fmap | ||
| 331 | | _ => by constructor | ||
| 332 | end; eauto using final. | ||
| 333 | - apply reflexive_eq. f_equal. apply map_eq=> x. | ||
| 334 | rewrite !lookup_fmap. by destruct (_ !! _) as [[]|]. | ||
| 335 | - by destruct (ts !! _). | ||
| 336 | - apply (map_minimal_key_Some _) in H as [[t Hx] ?]. | ||
| 337 | split; [done|]. right. eexists s, _; split_and!. | ||
| 338 | + by rewrite lookup_fmap Hx. | ||
| 339 | + intros y. rewrite lookup_fmap fmap_is_Some. auto. | ||
| 340 | + rewrite 3!fmap_insert map_fmap_singleton /=. | ||
| 341 | by rewrite lookup_total_alt Hx fmap_delete. | ||
| 342 | - apply map_minimal_key_None in H as ->. auto. | ||
| 343 | - split; [done|]. by rewrite map_fmap_union. | ||
| 344 | Qed. | ||
| 345 | |||
| 346 | Lemma interp_bin_op_Some_Some_2 op v1 f Φ v2 e3 : | ||
| 347 | interp_bin_op op v1 = Some f → | ||
| 348 | sem_bin_op op (val_to_expr v1) Φ → | ||
| 349 | Φ (val_to_expr v2) e3 → | ||
| 350 | ∃ t3, f v2 = Some t3 ∧ thunk_to_expr t3 =D=> e3. | ||
| 351 | Proof. | ||
| 352 | unfold interp_bin_op, interp_eq; destruct v1; inv 2; intros; | ||
| 353 | repeat match goal with | ||
| 354 | | H : ∃ _, _ |- _ => destruct H | ||
| 355 | | H : _ ∧ _ |- _ => destruct H | ||
| 356 | | H : _ <$> _ = ∅ |- _ => apply fmap_empty_inv in H | ||
| 357 | | H : context [(_ <$> _) !! _ = _] |- _ => rewrite lookup_fmap in H | ||
| 358 | | H : context [dom (_ <$> _)] |- _ => rewrite !dom_fmap_L in H | ||
| 359 | | H : context [length (_ <$> _)] |- _ => rewrite !length_fmap in H | ||
| 360 | | _ => progress simplify_option_eq | ||
| 361 | | H : val_to_expr ?v2 = _ |- _ => destruct v2 | ||
| 362 | | _ => case_match | ||
| 363 | | _ => eexists; split; [|by apply interp_eq_attr_correct] | ||
| 364 | | _ => eexists; split; [|by apply interp_eq_list_correct] | ||
| 365 | | _ => eexists; split; [|by apply interp_lt_list_correct] | ||
| 366 | | _ => eexists; split; [done|] | ||
| 367 | | _ => rewrite map_fmap_singleton | ||
| 368 | | _ => rewrite fmap_delete | ||
| 369 | | _ => rewrite subst_env_empty | ||
| 370 | | _ => rewrite fmap_app | ||
| 371 | | _ => rewrite map_fmap_union | ||
| 372 | end; eauto. | ||
| 373 | - rewrite (option_to_eq_Some_Some _ _ H1) /=. by eexists. | ||
| 374 | - apply reflexive_eq. f_equal. apply map_eq=> x. | ||
| 375 | rewrite !lookup_fmap. by destruct (_ !! _) as [[]|]. | ||
| 376 | - rewrite lookup_fmap. by destruct (_ !! _). | ||
| 377 | - destruct H1 as [[Hemp _]|(x & e' & Hx & Hleast & ->)]. | ||
| 378 | { by apply fmap_empty_inv in Hemp as ->. } | ||
| 379 | rewrite lookup_fmap fmap_Some in Hx. destruct Hx as (t & Hx & [= ->]). | ||
| 380 | setoid_rewrite lookup_fmap in Hleast. setoid_rewrite fmap_is_Some in Hleast. | ||
| 381 | apply (map_minimal_key_Some _) in H as [??]. | ||
| 382 | assert (s = x) as -> by (apply (anti_symm attr_le); naive_solver). | ||
| 383 | rewrite 3!fmap_insert map_fmap_singleton /= fmap_delete. | ||
| 384 | rewrite lookup_total_alt Hx. done. | ||
| 385 | - apply map_minimal_key_None in H as ->. naive_solver. | ||
| 386 | Qed. | ||
| 387 | |||
| 388 | Lemma interp_match_subst E ts ms strict : | ||
| 389 | interp_match ts (fmap (M:=option) (subst_env E) <$> ms) strict = | ||
| 390 | fmap (M:=gmap string) (sum_map (subst_env E) id) <$> interp_match ts ms strict. | ||
| 391 | Proof. | ||
| 392 | unfold interp_match. set (f mt mme := match mt with _ => _ end). | ||
| 393 | revert ts. induction ms as [|x mt ms Hmsx IH] using map_ind; intros ts. | ||
| 394 | { rewrite fmap_empty merge_empty_r. | ||
| 395 | induction ts as [|x t ts Hmsx IH] using map_ind; [done|]. | ||
| 396 | rewrite omap_insert /=. destruct strict; simplify_eq/=. | ||
| 397 | { rewrite map_mapM_insert_option //= lookup_omap Hmsx. done. } | ||
| 398 | rewrite -omap_delete delete_notin //. } | ||
| 399 | destruct (ts !! x) as [t|] eqn:Htsx. | ||
| 400 | { rewrite -(insert_delete ts x t) // fmap_insert. | ||
| 401 | rewrite -!(insert_merge _ _ _ _ (Some (inr t))) //. | ||
| 402 | rewrite !map_mapM_insert_option /=; | ||
| 403 | [|by rewrite lookup_merge lookup_delete ?lookup_fmap Hmsx..]. | ||
| 404 | rewrite IH. destruct (map_mapM id _); rewrite /= ?fmap_insert //. } | ||
| 405 | rewrite -(insert_merge_r _ _ _ _ (inl <$> mt)) /=; last first. | ||
| 406 | { rewrite Htsx /=. by destruct mt. } | ||
| 407 | rewrite fmap_insert. | ||
| 408 | rewrite -(insert_merge_r _ _ _ _ (inl <$> (subst_env E <$> mt))) /=; last first. | ||
| 409 | { rewrite Htsx /=. by destruct mt. } | ||
| 410 | rewrite !map_mapM_insert_option /=; | ||
| 411 | [|by rewrite lookup_merge ?lookup_fmap Htsx Hmsx..]. | ||
| 412 | rewrite IH. destruct mt, (map_mapM id _); rewrite /= ?fmap_insert //. | ||
| 413 | Qed. | ||
| 414 | |||
| 415 | Lemma interp_match_Some_1 ts ms strict tαs : | ||
| 416 | interp_match ts ms strict = Some tαs → | ||
| 417 | matches (thunk_to_expr <$> ts) ms strict (tattr_to_attr ∅ <$> tαs). | ||
| 418 | Proof. | ||
| 419 | unfold interp_match. set (f mt mme := match mt with _ => _ end). | ||
| 420 | revert ts tαs. | ||
| 421 | induction ms as [|x mt ms Hmsx IH] using map_ind; intros ts αs Hmatch. | ||
| 422 | { rewrite merge_empty_r in Hmatch. revert αs Hmatch. | ||
| 423 | induction ts as [|x t ts Hmsx IH] using map_ind; intros ts' Hmatch. | ||
| 424 | { rewrite omap_empty map_mapM_empty in Hmatch. injection Hmatch as <-. | ||
| 425 | rewrite !fmap_empty. constructor. } | ||
| 426 | rewrite omap_insert /= in Hmatch. destruct strict; simplify_eq/=. | ||
| 427 | { rewrite map_mapM_insert_option //= in Hmatch. by rewrite lookup_omap Hmsx. } | ||
| 428 | rewrite fmap_insert. | ||
| 429 | rewrite -omap_delete delete_notin // in Hmatch. apply IH in Hmatch. | ||
| 430 | apply matches_strict; rewrite ?lookup_fmap ?Hmsx; eauto. } | ||
| 431 | destruct (ts !! x) as [t|] eqn:Htsx. | ||
| 432 | { rewrite -(insert_delete ts x t) //. | ||
| 433 | rewrite -(insert_delete ts x t) // in Hmatch. | ||
| 434 | rewrite -(insert_merge _ _ _ _ (Some (inr t))) // in Hmatch. | ||
| 435 | rewrite map_mapM_insert_option /= in Hmatch; | ||
| 436 | last (by rewrite lookup_merge lookup_delete Hmsx). | ||
| 437 | destruct (map_mapM id _) as [E''|] eqn:?; simplify_eq/=. | ||
| 438 | injection Hmatch as <-. | ||
| 439 | rewrite !fmap_insert /=. constructor. | ||
| 440 | - by rewrite lookup_fmap lookup_delete. | ||
| 441 | - done. | ||
| 442 | - by apply IH. } | ||
| 443 | rewrite -(insert_merge_r _ _ _ _ (inl <$> mt)) /= in Hmatch; last first. | ||
| 444 | { rewrite Htsx /=. by destruct mt. } | ||
| 445 | rewrite map_mapM_insert_option /= in Hmatch; | ||
| 446 | last (by rewrite lookup_merge Htsx Hmsx). | ||
| 447 | destruct mt as [t|]; simplify_eq/=. | ||
| 448 | destruct (map_mapM id _) as [E''|] eqn:?; simplify_eq/=. | ||
| 449 | injection Hmatch as <-. rewrite !fmap_insert /= subst_env_empty. constructor. | ||
| 450 | - by rewrite lookup_fmap Htsx. | ||
| 451 | - done. | ||
| 452 | - by apply IH. | ||
| 453 | Qed. | ||
| 454 | |||
| 455 | Lemma interp_match_Some_2 es ms strict αs : | ||
| 456 | matches es ms strict αs → | ||
| 457 | interp_match (Thunk ∅ <$> es) ms strict = Some (attr_to_tattr ∅ <$> αs). | ||
| 458 | Proof. | ||
| 459 | unfold interp_match. set (f mt mme := match mt with _ => _ end). | ||
| 460 | induction 1; [done|..]. | ||
| 461 | - rewrite fmap_empty merge_empty_r. | ||
| 462 | induction es as [|x e es ? IH] using map_ind; [done|]. | ||
| 463 | rewrite fmap_insert omap_insert /= -omap_delete -fmap_delete delete_notin //. | ||
| 464 | - rewrite !fmap_insert /=. | ||
| 465 | rewrite -(insert_merge _ _ _ _ (Some (inr (Thunk ∅ e)))) //. | ||
| 466 | rewrite map_mapM_insert_option /=; last first. | ||
| 467 | { by rewrite lookup_merge !lookup_fmap H H0. } | ||
| 468 | by rewrite IHmatches. | ||
| 469 | - rewrite !fmap_insert /=. | ||
| 470 | rewrite -(insert_merge_r _ _ _ _ (Some (inl d))) /=; last first. | ||
| 471 | { by rewrite lookup_fmap H. } | ||
| 472 | rewrite map_mapM_insert_option /=; last first. | ||
| 473 | { by rewrite lookup_merge !lookup_fmap H H0. } | ||
| 474 | by rewrite IHmatches /=. | ||
| 475 | Qed. | ||
| 476 | |||
| 477 | Lemma force_deep_le {n1 n2 v mv} : | ||
| 478 | force_deep n1 v = Res mv → n1 ≤ n2 → force_deep n2 v = Res mv | ||
| 479 | with interp_le {n1 n2 E e mv} : | ||
| 480 | interp n1 E e = Res mv → n1 ≤ n2 → interp n2 E e = Res mv | ||
| 481 | with interp_thunk_le {n1 n2 t mvs} : | ||
| 482 | interp_thunk n1 t = Res mvs → n1 ≤ n2 → interp_thunk n2 t = Res mvs | ||
| 483 | with interp_app_le {n1 n2 v t mv} : | ||
| 484 | interp_app n1 v t = Res mv → n1 ≤ n2 → interp_app n2 v t = Res mv. | ||
| 485 | Proof. | ||
| 486 | - destruct n1 as [|n1], n2 as [|n2]; intros Ht ?; [done || lia..|]. | ||
| 487 | rewrite force_deep_S in Ht; rewrite force_deep_S; simpl in *. | ||
| 488 | destruct v as []; simplify_res; try done. | ||
| 489 | + destruct (mapM _ _) as [mvs|] eqn:Hmap; simplify_res. | ||
| 490 | erewrite mapM_Res_impl; [done..|]. intros t mw Hinterp; simpl in *. | ||
| 491 | destruct (interp_thunk n1 _) as [mw'|] eqn:Hinterp'; simplify_res. | ||
| 492 | rewrite (interp_thunk_le _ _ _ _ Hinterp') /=; last lia. | ||
| 493 | destruct mw'; simplify_res; eauto with lia. | ||
| 494 | + destruct (map_mapM_sorted _ _ _) eqn:?; simplify_res. | ||
| 495 | erewrite (map_mapM_sorted_Res_impl attr_le); [done..|]. | ||
| 496 | clear dependent ts. intros t mw Hinterp; simpl in *. | ||
| 497 | destruct (interp_thunk n1 _) as [mw'|] eqn:Hinterp'; simplify_res. | ||
| 498 | rewrite (interp_thunk_le _ _ _ _ Hinterp') /=; last lia. | ||
| 499 | destruct mw'; simplify_res; eauto with lia. | ||
| 500 | - destruct n1 as [|n1], n2 as [|n2]; intros He ?; [done || lia..|]. | ||
| 501 | rewrite interp_S in He; rewrite interp_S; destruct e; | ||
| 502 | repeat match goal with | ||
| 503 | | _ => case_match | ||
| 504 | | H : context [(_ <$> ?mx)] |- _ => destruct mx eqn:?; simplify_res | ||
| 505 | | H : ?r ≫= _ = _ |- _ => destruct r as [[]|] eqn:?; simplify_res | ||
| 506 | | H : _ <$> ?r = _ |- _ => destruct r as [[]|] eqn:?; simplify_res | ||
| 507 | | H : interp ?n' _ _ = Res ?mv |- interp ?n ?E ?e ≫= _ = _ => | ||
| 508 | rewrite (interp_le n' n E e mv); [|done || lia..] | ||
| 509 | | H : interp_app ?n' _ _ = Res ?mv |- interp_app ?n ?E ?e ≫= _ = _ => | ||
| 510 | rewrite (interp_app_le n' n E e mv); [|done || lia..] | ||
| 511 | | H : force_deep ?n' _ = Res ?mv |- force_deep ?n ?t ≫= _ = _ => | ||
| 512 | rewrite (force_deep_le n' n t mv); [|done || lia..] | ||
| 513 | | _ => progress simplify_res | ||
| 514 | | _ => progress simplify_option_eq | ||
| 515 | end; eauto with lia. | ||
| 516 | - destruct n1 as [|n1], n2 as [|n2]; intros He ?; [by (done || lia)..|]. | ||
| 517 | rewrite interp_thunk_S in He. rewrite interp_thunk_S. | ||
| 518 | destruct t; repeat (case_match || destruct (_ !! _) | ||
| 519 | || simplify_res); eauto with lia. | ||
| 520 | - destruct n1 as [|n1], n2 as [|n2]; intros He ?; [by (done || lia)..|]. | ||
| 521 | rewrite interp_app_S /= in He; rewrite interp_app_S /=. | ||
| 522 | destruct v; simplify_res; eauto with lia. | ||
| 523 | + destruct (interp_thunk n1 t) as [mw|] eqn:?; simplify_res. | ||
| 524 | erewrite interp_thunk_le by eauto with lia. simpl. | ||
| 525 | destruct mw as [w|]; simplify_res; [|done]. | ||
| 526 | destruct (maybe VAttr w) as [ts|]; simplify_res; [|done]. | ||
| 527 | destruct (interp_match _ _ _); simplify_res; eauto with lia. | ||
| 528 | + destruct (_ !! "__functor") as [tf|]; simplify_res; [|done]. | ||
| 529 | destruct (interp_thunk n1 tf) as [mw|] eqn:?; simplify_res. | ||
| 530 | erewrite interp_thunk_le by eauto with lia. simpl. | ||
| 531 | destruct mw as [w|]; simplify_res; [|done]. | ||
| 532 | destruct (interp_app n1 _ _) as [mw|] eqn:?; simplify_res. | ||
| 533 | erewrite interp_app_le by eauto with lia; simpl. | ||
| 534 | destruct mw; simplify_res; eauto with lia. | ||
| 535 | Qed. | ||
| 536 | |||
| 537 | Lemma mapM_interp_le {n1 n2 ts mvs} : | ||
| 538 | mapM (mbind (force_deep n1) ∘ interp_thunk n1) ts = Res mvs → | ||
| 539 | n1 ≤ n2 → | ||
| 540 | mapM (mbind (force_deep n2) ∘ interp_thunk n2) ts = Res mvs. | ||
| 541 | Proof. | ||
| 542 | intros. eapply mapM_Res_impl; [done|]; simpl; intros t mv ?. | ||
| 543 | destruct (interp_thunk _ _) as [mw|] eqn:Hthunk; simplify_res. | ||
| 544 | rewrite (interp_thunk_le Hthunk) //=. | ||
| 545 | destruct mw; simplify_res; eauto using force_deep_le. | ||
| 546 | Qed. | ||
| 547 | Lemma map_mapM_interp_le {n1 n2 ts mvs} : | ||
| 548 | map_mapM_sorted attr_le (mbind (force_deep n1) ∘ interp_thunk n1) ts = Res mvs → | ||
| 549 | n1 ≤ n2 → | ||
| 550 | map_mapM_sorted attr_le (mbind (force_deep n2) ∘ interp_thunk n2) ts = Res mvs. | ||
| 551 | Proof. | ||
| 552 | intros. eapply (map_mapM_sorted_Res_impl attr_le); [done|]; simpl. | ||
| 553 | intros t mv ?. destruct (interp_thunk _ _) as [mw|] eqn:Hthunk; simplify_res. | ||
| 554 | rewrite (interp_thunk_le Hthunk) //=. | ||
| 555 | destruct mw; simplify_res; eauto using force_deep_le. | ||
| 556 | Qed. | ||
| 557 | |||
| 558 | Lemma interp_agree {n1 n2 E e mv1 mv2} : | ||
| 559 | interp n1 E e = Res mv1 → interp n2 E e = Res mv2 → mv1 = mv2. | ||
| 560 | Proof. | ||
| 561 | intros He1 He2. apply (inj Res). destruct (total (≤) n1 n2). | ||
| 562 | - rewrite -He2. symmetry. eauto using interp_le. | ||
| 563 | - rewrite -He1. eauto using interp_le. | ||
| 564 | Qed. | ||
| 565 | |||
| 566 | Lemma subst_env_union E1 E2 e : | ||
| 567 | subst_env (union_kinded E1 E2) e = subst_env E1 (subst_env E2 e). | ||
| 568 | Proof. | ||
| 569 | rewrite !subst_env_alt -subst_union. f_equal. apply map_eq=> x. | ||
| 570 | rewrite lookup_union_with !lookup_fmap lookup_union_with. | ||
| 571 | by repeat destruct (_ !! _) as [[[]]|]. | ||
| 572 | Qed. | ||
| 573 | |||
| 574 | Lemma union_kinded_union {A} (E1 E2 : gmap string (kind * A)) : | ||
| 575 | map_Forall (λ _ ka, ka.1 = ABS) E1 → union_kinded E1 E2 = E1 ∪ E2. | ||
| 576 | Proof. | ||
| 577 | rewrite map_Forall_lookup; intros. | ||
| 578 | apply map_eq=> x. rewrite lookup_union_with lookup_union. | ||
| 579 | destruct (E1 !! x) as [[[] a]|] eqn:?; naive_solver. | ||
| 580 | Qed. | ||
| 581 | |||
| 582 | Lemma subst_abs_env_insert E x e t : | ||
| 583 | subst_env (<[x:=(ABS, t)]> E) e | ||
| 584 | = subst {[x:=(ABS, thunk_to_expr t)]} (subst_env E e). | ||
| 585 | Proof. | ||
| 586 | assert (<[x:=(ABS, t)]> E = | ||
| 587 | union_kinded {[x:=(ABS, t)]} E) as ->. | ||
| 588 | { apply map_eq=> y. rewrite lookup_union_with. | ||
| 589 | destruct (decide (x = y)) as [->|]. | ||
| 590 | - rewrite lookup_insert lookup_singleton /=. by destruct (_ !! _). | ||
| 591 | - rewrite lookup_insert_ne // lookup_singleton_ne //. by destruct (_ !! _). } | ||
| 592 | rewrite subst_env_union subst_env_alt. by rewrite map_fmap_singleton. | ||
| 593 | Qed. | ||
| 594 | |||
| 595 | Lemma subst_abs_as_subst_env x e1 e2 : | ||
| 596 | subst {[x:=(ABS, e2)]} e1 = subst_env (<[x:=(ABS, Thunk ∅ e2)]> ∅) e1. | ||
| 597 | Proof. rewrite subst_abs_env_insert //= !subst_env_empty //. Qed. | ||
| 598 | |||
| 599 | Lemma subst_env_insert_proper e1 e2 E1 E2 x t1 t2 : | ||
| 600 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 601 | thunk_to_expr t1 = thunk_to_expr t2 → | ||
| 602 | subst_env (<[x:=(ABS, t1)]> E1) e1 = subst_env (<[x:=(ABS, t2)]> E2) e2. | ||
| 603 | Proof. rewrite !subst_abs_env_insert //. auto with f_equal. Qed. | ||
| 604 | |||
| 605 | Lemma subst_env_insert_proper' e1 e2 E1 E2 x E1' E2' e1' e2' : | ||
| 606 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 607 | subst_env E1' e1' = subst_env E2' e2' → | ||
| 608 | subst_env (<[x:=(ABS, Thunk E1' e1')]> E1) e1 | ||
| 609 | = subst_env (<[x:=(ABS, Thunk E2' e2')]> E2) e2. | ||
| 610 | Proof. intros. by apply subst_env_insert_proper. Qed. | ||
| 611 | |||
| 612 | Lemma subst_env_union_fmap_proper k e1 e2 E1 E2 ts1 ts2 : | ||
| 613 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 614 | AttrN ∘ thunk_to_expr <$> ts1 = AttrN ∘ thunk_to_expr <$> ts2 → | ||
| 615 | subst_env (union_kinded ((k,.) <$> ts1) E1) e1 | ||
| 616 | = subst_env (union_kinded ((k,.) <$> ts2) E2) e2. | ||
| 617 | Proof. | ||
| 618 | intros He Hes. rewrite !subst_env_union; [|by apply env_unionable_with..]. | ||
| 619 | rewrite He !subst_env_alt /=. f_equal. | ||
| 620 | apply map_eq=> x. rewrite !lookup_fmap. | ||
| 621 | apply (f_equal (.!! x)) in Hes. rewrite !lookup_fmap in Hes. | ||
| 622 | destruct (ts1 !! x), (ts2 !! x); simplify_eq/=; auto with f_equal. | ||
| 623 | Qed. | ||
| 624 | |||
| 625 | Lemma subst_env_fmap_proper k e ts1 ts2 : | ||
| 626 | AttrN ∘ thunk_to_expr <$> ts1 = AttrN ∘ thunk_to_expr <$> ts2 → | ||
| 627 | subst_env ((k,.) <$> ts1) e = subst_env ((k,.) <$> ts2) e. | ||
| 628 | Proof. | ||
| 629 | intros. rewrite -(right_id_L ∅ (union_kinded) (_ <$> ts1)) | ||
| 630 | -(right_id_L ∅ (union_kinded) (_ <$> ts2)). | ||
| 631 | by apply subst_env_union_fmap_proper. | ||
| 632 | Qed. | ||
| 633 | |||
| 634 | Lemma tattr_to_attr_from_attr E (αs : gmap string attr) : | ||
| 635 | tattr_to_attr E <$> (attr_to_tattr E <$> αs) = attr_subst_env E <$> αs. | ||
| 636 | Proof. | ||
| 637 | apply map_eq=> x. rewrite /tattr_to_attr !lookup_fmap. | ||
| 638 | destruct (αs !! x) as [[[] ]|] eqn:?; auto. | ||
| 639 | Qed. | ||
| 640 | |||
| 641 | Lemma tattr_to_attr_from_attr_empty (αs : gmap string attr) : | ||
| 642 | tattr_to_attr ∅ <$> (attr_to_tattr ∅ <$> αs) = αs. | ||
| 643 | Proof. | ||
| 644 | rewrite tattr_to_attr_from_attr. apply map_eq=> x. rewrite !lookup_fmap. | ||
| 645 | destruct (αs !! x) as [[[] ]|] eqn:?; f_equal/=; by rewrite subst_env_empty. | ||
| 646 | Qed. | ||
| 647 | |||
| 648 | Lemma indirects_env_proper E1 E2 tαs1 tαs2 e1 e2 : | ||
| 649 | tattr_to_attr E1 <$> tαs1 = tattr_to_attr E2 <$> tαs2 → | ||
| 650 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 651 | subst_env (indirects_env E1 tαs1) e1 = subst_env (indirects_env E2 tαs2) e2. | ||
| 652 | Proof. | ||
| 653 | intros Htαs HE. rewrite /indirects_env -!union_kinded_union; | ||
| 654 | [|intros ??; rewrite map_lookup_imap=> ?; by simplify_option_eq..]. | ||
| 655 | rewrite !subst_env_union HE !subst_env_alt. f_equal. | ||
| 656 | apply map_eq=> x. rewrite !lookup_fmap !map_lookup_imap. | ||
| 657 | pose proof (f_equal (.!! x) Htαs) as Hx. rewrite !lookup_fmap in Hx. | ||
| 658 | repeat destruct (_ !! x) as [[]|]; simplify_eq/=; auto with f_equal. | ||
| 659 | Qed. | ||
| 660 | |||
| 661 | Lemma subst_env_indirects_env E tαs e : | ||
| 662 | subst_env (indirects_env E tαs) e | ||
| 663 | = subst_env (indirects_env ∅ (attr_to_tattr ∅ <$> (tattr_to_attr E <$> tαs))) | ||
| 664 | (subst_env E e). | ||
| 665 | Proof. | ||
| 666 | rewrite /indirects_env -!union_kinded_union; | ||
| 667 | [|intros ??; rewrite map_lookup_imap=> ?; by simplify_option_eq..]. | ||
| 668 | rewrite !subst_env_union subst_env_empty !subst_env_alt. | ||
| 669 | f_equal. apply map_eq=> x. rewrite !lookup_fmap !map_lookup_imap !lookup_fmap. | ||
| 670 | destruct (_ !! _) as [[]|]; | ||
| 671 | do 4 f_equal/=; auto using tattr_to_attr_from_attr_empty. | ||
| 672 | Qed. | ||
| 673 | |||
| 674 | Lemma subst_env_indirects_env_attr_to_tattr E αs e : | ||
| 675 | subst_env (indirects_env E (attr_to_tattr E <$> αs)) e | ||
| 676 | = subst (indirects (attr_subst_env E <$> αs)) (subst_env E e). | ||
| 677 | Proof. | ||
| 678 | rewrite /indirects_env -!union_kinded_union; | ||
| 679 | [|intros ??; rewrite map_lookup_imap=> ?; by simplify_option_eq..]. | ||
| 680 | rewrite subst_env_union !subst_env_alt. f_equal. | ||
| 681 | apply map_eq=> x. rewrite !lookup_fmap !map_lookup_imap !lookup_fmap. | ||
| 682 | repeat destruct (_ !! x) as [[]|]; simplify_eq/=; do 4 f_equal/=. | ||
| 683 | by rewrite tattr_to_attr_from_attr. | ||
| 684 | Qed. | ||
| 685 | |||
| 686 | Lemma subst_env_indirects_env_attr_to_tattr_empty αs e : | ||
| 687 | subst_env (indirects_env ∅ (attr_to_tattr ∅ <$> αs)) e = | ||
| 688 | subst (indirects αs) e. | ||
| 689 | Proof. | ||
| 690 | rewrite subst_env_indirects_env_attr_to_tattr subst_env_empty. do 3 f_equal. | ||
| 691 | apply map_eq=> x. rewrite !lookup_fmap. | ||
| 692 | destruct (_ !! x) as [[]|]; do 2 f_equal/=; auto using subst_env_empty. | ||
| 693 | Qed. | ||
| 694 | |||
| 695 | Lemma interp_val_to_expr E e v : | ||
| 696 | subst_env E e = val_to_expr v → | ||
| 697 | ∃ w m, interp m E e = mret w ∧ val_to_expr v = val_to_expr w. | ||
| 698 | Proof. | ||
| 699 | revert v. induction e; intros []; | ||
| 700 | rewrite subst_env_eq; intros; simplify_eq/=. | ||
| 701 | - eexists (VLit _ ltac:(done)), 1. split; [|done]. by rewrite interp_lit. | ||
| 702 | - eexists (VClo _ _ _), 1. rewrite interp_S /=. auto with f_equal. | ||
| 703 | - eexists (VCloMatch _ _ _ _), 1. rewrite interp_S /=. auto with f_equal. | ||
| 704 | - eexists (VList _), 1. rewrite interp_S /=. split; [done|]. | ||
| 705 | f_equal. rewrite -H0. clear. | ||
| 706 | induction es; f_equal/=; auto. | ||
| 707 | - eexists (VAttr _), 1. rewrite interp_S /=. split; [done|]. | ||
| 708 | assert (no_recs αs) as Hrecs. | ||
| 709 | { intros y α Hy. | ||
| 710 | apply (f_equal (.!! y)) in H0. rewrite !lookup_fmap Hy /= in H0. | ||
| 711 | destruct (ts !! y), α; by simplify_eq/=. } | ||
| 712 | rewrite from_attr_no_recs // -H0. | ||
| 713 | f_equal. apply map_eq=> y. | ||
| 714 | rewrite !lookup_fmap. destruct (αs !! y) as [[]|] eqn:?; do 2 f_equal/=. | ||
| 715 | eauto using no_recs_lookup. | ||
| 716 | Qed. | ||
| 717 | |||
| 718 | Lemma interp_val_to_expr_Res m E e v mw : | ||
| 719 | subst_env E e = val_to_expr v → | ||
| 720 | interp m E e = Res mw → | ||
| 721 | Some (val_to_expr v) = val_to_expr <$> mw. | ||
| 722 | Proof. | ||
| 723 | intros (mw' & m' & Hinterp' & ->)%interp_val_to_expr Hinterp. | ||
| 724 | by assert (mw = Some mw') as -> by eauto using interp_agree. | ||
| 725 | Qed. | ||
| 726 | |||
| 727 | Lemma interp_empty_val_to_expr v : | ||
| 728 | ∃ w m, interp m ∅ (val_to_expr v) = mret w ∧ val_to_expr v = val_to_expr w. | ||
| 729 | Proof. apply interp_val_to_expr. by rewrite subst_env_empty. Qed. | ||
| 730 | |||
| 731 | Lemma interp_empty_val_to_expr_Res m v mw : | ||
| 732 | interp m ∅ (val_to_expr v) = Res mw → | ||
| 733 | Some (val_to_expr v) = val_to_expr <$> mw. | ||
| 734 | Proof. apply interp_val_to_expr_Res. by rewrite subst_env_empty. Qed. | ||
| 735 | |||
| 736 | Lemma interp_eq_list_proper ts1 ts2 ts1' ts2' : | ||
| 737 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts1' → | ||
| 738 | thunk_to_expr <$> ts2 = thunk_to_expr <$> ts2' → | ||
| 739 | thunk_to_expr (interp_eq_list ts1 ts2) | ||
| 740 | = thunk_to_expr (interp_eq_list ts1' ts2'). | ||
| 741 | Proof. | ||
| 742 | intros Hts1 Hts2. rewrite /= !subst_env_alt. | ||
| 743 | f_equal; last first. | ||
| 744 | { revert ts1' ts2 ts2' Hts1 Hts2. generalize 0. | ||
| 745 | induction ts1; intros ? [] [] [] ??; simplify_eq/=; auto with f_equal. } | ||
| 746 | rewrite !map_fmap_union. f_equal; apply map_eq=> y; rewrite !lookup_fmap. | ||
| 747 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 748 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 749 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0. | ||
| 750 | rewrite lookup_fmap lookup_map_seq_0 in Hy. | ||
| 751 | apply (f_equal (.!! y')) in Hts1. rewrite !list_lookup_fmap in Hts1. | ||
| 752 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 753 | + rewrite lookup_kmap_None in Hy. | ||
| 754 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 755 | intros y' ->. generalize (Hy _ eq_refl). | ||
| 756 | rewrite !lookup_fmap !lookup_map_seq_0. | ||
| 757 | apply (f_equal (.!! y')) in Hts1. rewrite !list_lookup_fmap in Hts1. | ||
| 758 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 759 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 760 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 761 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0. | ||
| 762 | rewrite lookup_fmap lookup_map_seq_0 in Hy. | ||
| 763 | apply (f_equal (.!! y')) in Hts2. rewrite !list_lookup_fmap in Hts2. | ||
| 764 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 765 | + rewrite lookup_kmap_None in Hy. | ||
| 766 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 767 | intros y' ->. generalize (Hy _ eq_refl). | ||
| 768 | rewrite !lookup_fmap !lookup_map_seq_0. | ||
| 769 | apply (f_equal (.!! y')) in Hts2. rewrite !list_lookup_fmap in Hts2. | ||
| 770 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 771 | Qed. | ||
| 772 | |||
| 773 | Lemma interp_lt_list_proper ts1 ts2 ts1' ts2' : | ||
| 774 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts1' → | ||
| 775 | thunk_to_expr <$> ts2 = thunk_to_expr <$> ts2' → | ||
| 776 | thunk_to_expr (interp_lt_list ts1 ts2) | ||
| 777 | = thunk_to_expr (interp_lt_list ts1' ts2'). | ||
| 778 | Proof. | ||
| 779 | intros Hts1 Hts2. rewrite /= !subst_env_alt. | ||
| 780 | f_equal; last first. | ||
| 781 | { revert ts1' ts2 ts2' Hts1 Hts2. generalize 0. | ||
| 782 | induction ts1; intros ? [] [] [] ??; simplify_eq/=; auto with f_equal. } | ||
| 783 | rewrite !map_fmap_union. f_equal; apply map_eq=> y; rewrite !lookup_fmap. | ||
| 784 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 785 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 786 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0. | ||
| 787 | rewrite lookup_fmap lookup_map_seq_0 in Hy. | ||
| 788 | apply (f_equal (.!! y')) in Hts1. rewrite !list_lookup_fmap in Hts1. | ||
| 789 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 790 | + rewrite lookup_kmap_None in Hy. | ||
| 791 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 792 | intros y' ->. generalize (Hy _ eq_refl). | ||
| 793 | rewrite !lookup_fmap !lookup_map_seq_0. | ||
| 794 | apply (f_equal (.!! y')) in Hts1. rewrite !list_lookup_fmap in Hts1. | ||
| 795 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 796 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 797 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 798 | rewrite (lookup_kmap _) lookup_fmap lookup_map_seq_0. | ||
| 799 | rewrite lookup_fmap lookup_map_seq_0 in Hy. | ||
| 800 | apply (f_equal (.!! y')) in Hts2. rewrite !list_lookup_fmap in Hts2. | ||
| 801 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 802 | + rewrite lookup_kmap_None in Hy. | ||
| 803 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 804 | intros y' ->. generalize (Hy _ eq_refl). | ||
| 805 | rewrite !lookup_fmap !lookup_map_seq_0. | ||
| 806 | apply (f_equal (.!! y')) in Hts2. rewrite !list_lookup_fmap in Hts2. | ||
| 807 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 808 | Qed. | ||
| 809 | |||
| 810 | Lemma interp_eq_attr_proper ts1 ts2 ts1' ts2' : | ||
| 811 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts1' → | ||
| 812 | thunk_to_expr <$> ts2 = thunk_to_expr <$> ts2' → | ||
| 813 | thunk_to_expr (interp_eq_attr ts1 ts2) | ||
| 814 | = thunk_to_expr (interp_eq_attr ts1' ts2'). | ||
| 815 | Proof. | ||
| 816 | intros Hts1 Hts2. rewrite /= !subst_env_alt. | ||
| 817 | f_equal; last first. | ||
| 818 | { clear Hts2. f_equal. apply map_eq=> y. | ||
| 819 | rewrite !map_lookup_imap. apply (f_equal (.!! y)) in Hts1. | ||
| 820 | rewrite !lookup_fmap in Hts1. by repeat destruct (_ !! _). } | ||
| 821 | rewrite !map_fmap_union. f_equal; apply map_eq=> y; rewrite !lookup_fmap. | ||
| 822 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 823 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 824 | rewrite (lookup_kmap (String "1")) lookup_fmap. | ||
| 825 | rewrite lookup_fmap in Hy. | ||
| 826 | apply (f_equal (.!! y')) in Hts1. rewrite !lookup_fmap in Hts1. | ||
| 827 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 828 | + rewrite lookup_kmap_None in Hy. | ||
| 829 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 830 | intros y' ->. generalize (Hy _ eq_refl). rewrite !lookup_fmap. | ||
| 831 | apply (f_equal (.!! y')) in Hts1. rewrite !lookup_fmap in Hts1. | ||
| 832 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 833 | - destruct (kmap _ _ !! y) as [[k e]|] eqn:Hy; simplify_eq/=. | ||
| 834 | + apply (lookup_kmap_Some _) in Hy as (y' & -> & Hy). | ||
| 835 | rewrite (lookup_kmap (String "2")) lookup_fmap. | ||
| 836 | rewrite lookup_fmap in Hy. | ||
| 837 | apply (f_equal (.!! y')) in Hts2. rewrite !lookup_fmap in Hts2. | ||
| 838 | repeat destruct (_ !! _); simplify_eq/=; auto with f_equal. | ||
| 839 | + rewrite lookup_kmap_None in Hy. | ||
| 840 | apply symmetry, fmap_None, (lookup_kmap_None _). | ||
| 841 | intros y' ->. generalize (Hy _ eq_refl). rewrite !lookup_fmap. | ||
| 842 | apply (f_equal (.!! y')) in Hts2. rewrite !lookup_fmap in Hts2. | ||
| 843 | intros. repeat destruct (_ !! _); by simplify_eq/=. | ||
| 844 | Qed. | ||
| 845 | |||
| 846 | Opaque interp_eq_list interp_lt_list interp_eq_attr. | ||
| 847 | |||
| 848 | Lemma interp_bin_op_proper op v1 v2 : | ||
| 849 | val_to_expr v1 = val_to_expr v2 → | ||
| 850 | match interp_bin_op op v1, interp_bin_op op v2 with | ||
| 851 | | None, None => True | ||
| 852 | | Some f1, Some f2 => ∀ v1' v2', | ||
| 853 | val_to_expr v1' = val_to_expr v2' → | ||
| 854 | thunk_to_expr <$> f1 v1' = thunk_to_expr <$> f2 v2' | ||
| 855 | | _, _ => False | ||
| 856 | end. | ||
| 857 | Proof. | ||
| 858 | intros. unfold interp_bin_op, interp_eq; | ||
| 859 | repeat (done || case_match || simplify_eq/= || | ||
| 860 | destruct (option_to_eq_Some _) as [[]|]); | ||
| 861 | intros [] [] ?; simplify_eq/=; | ||
| 862 | repeat match goal with | ||
| 863 | | _ => done | ||
| 864 | | _ => progress simplify_option_eq | ||
| 865 | | _ => rewrite map_fmap_singleton | ||
| 866 | | _ => rewrite map_fmap_union | ||
| 867 | | _ => case_match | ||
| 868 | | |- context[ maybe _ ?x ] => is_var x; destruct x | ||
| 869 | end; auto with congruence. | ||
| 870 | - f_equal. by apply interp_eq_list_proper. | ||
| 871 | - apply (f_equal length) in H, H0. rewrite !length_fmap in H H0. congruence. | ||
| 872 | - apply (f_equal length) in H, H0. rewrite !length_fmap in H H0. congruence. | ||
| 873 | - f_equal. apply interp_eq_attr_proper. | ||
| 874 | + rewrite 2!map_fmap_compose in H. by simplify_eq. | ||
| 875 | + rewrite 2!map_fmap_compose in H0. by simplify_eq. | ||
| 876 | - apply (f_equal dom) in H, H0. rewrite !dom_fmap_L in H H0. congruence. | ||
| 877 | - apply (f_equal dom) in H, H0. rewrite !dom_fmap_L in H H0. congruence. | ||
| 878 | - destruct v1, v2; by simplify_eq/=. | ||
| 879 | - repeat destruct (option_to_eq_Some _) as [[]|]; simplify_eq/=; auto. | ||
| 880 | - do 3 f_equal. apply map_eq=> y. rewrite !lookup_fmap. | ||
| 881 | apply (f_equal (.!! y)) in H. rewrite !lookup_fmap in H. | ||
| 882 | repeat destruct (_ !! _) as [[]|]; naive_solver. | ||
| 883 | - f_equal. by apply interp_lt_list_proper. | ||
| 884 | - rewrite !fmap_insert /=. auto 10 with f_equal. | ||
| 885 | - by rewrite !fmap_app H0 H. | ||
| 886 | - apply (f_equal (.!! s)) in H. rewrite !lookup_fmap in H. | ||
| 887 | repeat destruct (_ !! _); simplify_eq/=; by f_equal. | ||
| 888 | - apply (f_equal (.!! s)) in H. rewrite !lookup_fmap in H. | ||
| 889 | repeat destruct (_ !! _); simplify_eq/=; by f_equal. | ||
| 890 | - rewrite !fmap_delete. congruence. | ||
| 891 | - assert (∀ y, is_Some (ts !! y) ↔ is_Some (ts0 !! y)) as Hx. | ||
| 892 | { intros y. rewrite -!(fmap_is_Some (AttrN ∘ thunk_to_expr)) -!lookup_fmap. | ||
| 893 | by rewrite H. } | ||
| 894 | apply (map_minimal_key_Some _) in H5 as [[t1 Hx1] ?], H8 as [[t2 Hx2] ?]. | ||
| 895 | assert (s0 = s) as -> by (apply (anti_symm attr_le); naive_solver). | ||
| 896 | pose proof (f_equal (.!! s) H) as Hs. rewrite !lookup_fmap in Hs. | ||
| 897 | rewrite !fmap_insert !fmap_empty /= !lookup_total_alt Hx1 Hx2 /=. | ||
| 898 | rewrite Hx1 Hx2 /= in Hs. simplify_eq/=. rewrite Hs !fmap_delete H. done. | ||
| 899 | - apply map_minimal_key_None in H8 as ->. | ||
| 900 | rewrite fmap_empty in H. by apply fmap_empty_inv in H as ->. | ||
| 901 | - apply map_minimal_key_None in H5 as ->. | ||
| 902 | rewrite fmap_empty in H. by apply symmetry, fmap_empty_inv in H as ->. | ||
| 903 | Qed. | ||
| 904 | |||
| 905 | Lemma interp_match_proper E1 E2 ts1 ts2 ms1 ms2 strict : | ||
| 906 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts2 → | ||
| 907 | fmap (M:=option) (subst_env E1) <$> ms1 = fmap (subst_env E2) <$> ms2 → | ||
| 908 | fmap (M:=gmap string) (tattr_to_attr E1) <$> interp_match ts1 ms1 strict | ||
| 909 | = fmap (tattr_to_attr E2) <$> interp_match ts2 ms2 strict. | ||
| 910 | Proof. | ||
| 911 | revert ms2 ts1 ts2. | ||
| 912 | induction ms1 as [|x m1 ms1 Hmsx IH] using map_ind; intros ms2 ts1 ts2 Hts Hms. | ||
| 913 | { rewrite fmap_empty in Hms. apply symmetry, fmap_empty_inv in Hms as ->. | ||
| 914 | rewrite /interp_match !merge_empty_r. revert ts2 Hts. | ||
| 915 | induction ts1 as [|x t1 ts1 Htsx IH] using map_ind; intros ts2 Hts. | ||
| 916 | { rewrite fmap_empty in Hts. by apply symmetry, fmap_empty_inv in Hts as ->. } | ||
| 917 | rewrite fmap_insert in Hts. | ||
| 918 | apply symmetry, fmap_insert_inv in Hts as (t2&ts2'&?&Htsx2&->&Hts); | ||
| 919 | last by rewrite lookup_fmap Htsx. | ||
| 920 | rewrite !omap_insert /=. destruct strict; simpl; | ||
| 921 | rewrite ?map_mapM_insert_option ?delete_notin //= ?lookup_omap ?Htsx ?Htsx2; | ||
| 922 | auto. } | ||
| 923 | rewrite fmap_insert in Hms. | ||
| 924 | apply symmetry, fmap_insert_inv in Hms as (m2&ms2'&?&Hmsx2&->&Hms); | ||
| 925 | last by rewrite lookup_fmap Hmsx. | ||
| 926 | pose proof (f_equal (.!! x) Hts) as Hx. rewrite !lookup_fmap in Hx. | ||
| 927 | destruct (ts1 !! x) as [t1|] eqn:Hts1x, (ts2 !! x) as [t2|] eqn:Hts2x; simplify_eq/=. | ||
| 928 | - rewrite -(insert_delete ts1 x t1) // -(insert_delete ts2 x t2) //. | ||
| 929 | rewrite /interp_match. erewrite <-!insert_merge by done. | ||
| 930 | rewrite !map_mapM_insert_option ?lookup_merge ?lookup_delete ?Hmsx ?Hmsx2 //=. | ||
| 931 | apply (f_equal (delete x)) in Hts. rewrite -!fmap_delete in Hts. | ||
| 932 | eapply IH in Hms; [|done]. rewrite /interp_match in Hms. | ||
| 933 | repeat destruct (map_mapM id _); simplify_eq/=; [|done..]. | ||
| 934 | rewrite !fmap_insert /=. auto with f_equal. | ||
| 935 | - rewrite /interp_match. | ||
| 936 | rewrite -!(insert_merge_r _ ts1 _ _ (inl <$> m1)); | ||
| 937 | last (rewrite Hts1x; by destruct m1). | ||
| 938 | rewrite -!(insert_merge_r _ ts2 _ _ (inl <$> m2)); | ||
| 939 | last (rewrite Hts2x; by destruct m2). | ||
| 940 | rewrite !map_mapM_insert_option ?lookup_merge ?Hts1x ?Hts2x ?Hmsx ?Hmsx2 //. | ||
| 941 | eapply IH in Hms; [|done]. rewrite /interp_match in Hms. | ||
| 942 | destruct m1, m2; simplify_eq/=; auto. | ||
| 943 | repeat destruct (map_mapM id _); simplify_eq/=; [|done..]. | ||
| 944 | rewrite !fmap_insert /=. auto with f_equal. | ||
| 945 | Qed. | ||
| 946 | |||
| 947 | Lemma mapM_interp_proper' n ts1 ts2 mvs : | ||
| 948 | (∀ v1 v2 mv, | ||
| 949 | val_to_expr v1 = val_to_expr v2 → | ||
| 950 | force_deep n v1 = Res mv → | ||
| 951 | ∃ mw m, force_deep m v2 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw) → | ||
| 952 | (∀ t1 t2 mv, | ||
| 953 | thunk_to_expr t1 = thunk_to_expr t2 → | ||
| 954 | interp_thunk n t1 = Res mv → | ||
| 955 | ∃ mw m, interp_thunk m t2 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw) → | ||
| 956 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts2 → | ||
| 957 | mapM (mbind (force_deep n) ∘ interp_thunk n) ts1 = Res mvs → | ||
| 958 | ∃ mws m, mapM (mbind (force_deep m) ∘ interp_thunk m) ts2 = Res mws ∧ | ||
| 959 | fmap (M:=list) val_to_expr <$> mvs = fmap (M:=list) val_to_expr <$> mws. | ||
| 960 | Proof. | ||
| 961 | intros force_deep_proper interp_thunk_proper Hts. | ||
| 962 | revert mvs. rewrite list_eq_Forall2 Forall2_fmap in Hts. | ||
| 963 | induction Hts as [|t1 t2 ts1 ts2 ?? IH]; intros mvs ?; simplify_res. | ||
| 964 | { by exists (Some []), 0. } | ||
| 965 | destruct (interp_thunk _ _) as [mv|] eqn:Hinterp'; simplify_res. | ||
| 966 | eapply interp_thunk_proper in Hinterp' | ||
| 967 | as (mw & m1 & Hinterp1 & Hw); [|by eauto..]. | ||
| 968 | destruct mv as [v|], mw as [w|]; simplify_res; last first. | ||
| 969 | { exists None, m1. by rewrite /= Hinterp1. } | ||
| 970 | destruct (force_deep _ _) as [mv'|] eqn:Hforce; simplify_res. | ||
| 971 | eapply force_deep_proper in Hforce as (mw' & m2 & Hforce2 & Hw'); last done. | ||
| 972 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 973 | { exists None, (m1 `max` m2). | ||
| 974 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 975 | rewrite (force_deep_le Hforce2) /=; last lia. done. } | ||
| 976 | destruct (mapM _ _) as [mvs'|] eqn:?; simplify_res. | ||
| 977 | destruct (IH _ eq_refl) as (mws & m3 & Hmap3 & ?). | ||
| 978 | exists ((w' ::.) <$> mws), (S (m1 `max` m2 `max` m3)). | ||
| 979 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 980 | rewrite (force_deep_le Hforce2) /=; last lia. | ||
| 981 | rewrite (mapM_interp_le Hmap3) /=; last lia. split; [by destruct mws|]. | ||
| 982 | destruct mvs', mws; simplify_res; auto 10 with f_equal. | ||
| 983 | Qed. | ||
| 984 | |||
| 985 | Lemma force_deep_proper n v1 v2 mv : | ||
| 986 | val_to_expr v1 = val_to_expr v2 → | ||
| 987 | force_deep n v1 = Res mv → | ||
| 988 | ∃ mw m, force_deep m v2 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw | ||
| 989 | with interp_proper n E1 E2 e1 e2 mv : | ||
| 990 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 991 | interp n E1 e1 = Res mv → | ||
| 992 | ∃ mw m, interp m E2 e2 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw | ||
| 993 | with interp_thunk_proper n t1 t2 mv : | ||
| 994 | thunk_to_expr t1 = thunk_to_expr t2 → | ||
| 995 | interp_thunk n t1 = Res mv → | ||
| 996 | ∃ mw m, interp_thunk m t2 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw | ||
| 997 | with interp_app_proper n v1 v2 t1' t2' mv : | ||
| 998 | val_to_expr v1 = val_to_expr v2 → | ||
| 999 | thunk_to_expr t1' = thunk_to_expr t2' → | ||
| 1000 | interp_app n v1 t1' = Res mv → | ||
| 1001 | ∃ mw m, interp_app m v2 t2' = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1002 | Proof. | ||
| 1003 | (* force_deep_proper *) | ||
| 1004 | - destruct n as [|n]; [done|]. | ||
| 1005 | intros Hv Hinterp. rewrite force_deep_S /force_deep1 in Hinterp. | ||
| 1006 | destruct v1 as [| | |ts1|ts1], v2 as [| | |ts2|ts2]; simplify_res. | ||
| 1007 | { eexists _, 1; split; [by rewrite force_deep_S|]. done. } | ||
| 1008 | { eexists _, 1; split; [by rewrite force_deep_S|]. simpl. auto with f_equal. } | ||
| 1009 | { eexists _, 1; split; [by rewrite force_deep_S|]. simpl. auto with f_equal. } | ||
| 1010 | { destruct (mapM _ _) as [mvs|] eqn:Hmap; simplify_res. | ||
| 1011 | eapply mapM_interp_proper' in Hmap as (mws & m & Hmap & Hmvs); [|by eauto..]. | ||
| 1012 | exists (VList ∘ fmap Forced <$> mws), (S m). | ||
| 1013 | rewrite force_deep_S /= Hmap. split; [done|]. | ||
| 1014 | destruct mvs, mws; simplify_eq/=; do 2 f_equal. | ||
| 1015 | rewrite list_eq_Forall2 Forall2_fmap in Hmvs. | ||
| 1016 | by rewrite list_eq_Forall2 !Forall2_fmap. } | ||
| 1017 | destruct (map_mapM_sorted _ _ _) as [mvs|] eqn:Hmap; simplify_res. | ||
| 1018 | assert (∃ mws m, | ||
| 1019 | map_mapM_sorted attr_le | ||
| 1020 | (mbind (force_deep m) ∘ interp_thunk m) ts2 = Res mws ∧ | ||
| 1021 | fmap (M:=gmap _) val_to_expr <$> mvs = fmap (M:=gmap _) val_to_expr <$> mws) | ||
| 1022 | as (mws & m & Hmap' & Hmvs); last first. | ||
| 1023 | { exists (VAttr ∘ fmap Forced <$> mws), (S m). | ||
| 1024 | rewrite force_deep_S /= Hmap'. split; [done|]. | ||
| 1025 | destruct mvs, mws; simplify_eq/=; do 2 f_equal. | ||
| 1026 | apply map_eq=> x. rewrite !lookup_fmap. | ||
| 1027 | apply (f_equal (.!! x)) in Hmvs. rewrite !lookup_fmap in Hmvs. | ||
| 1028 | repeat destruct (_ !! x); simplify_res; auto with f_equal. } | ||
| 1029 | revert ts2 mvs Hmap Hv. induction ts1 as [|x t1 ts1 Hx1 ? IH] | ||
| 1030 | using (map_sorted_ind attr_le); intros ts2' mvs Hmap Hts. | ||
| 1031 | { exists (Some ∅), 0. rewrite fmap_empty in Hts. | ||
| 1032 | apply symmetry, fmap_empty_inv in Hts as ->. | ||
| 1033 | rewrite map_mapM_sorted_empty in Hmap; simplify_res. | ||
| 1034 | by rewrite map_mapM_sorted_empty. } | ||
| 1035 | rewrite map_mapM_sorted_insert //= in Hmap. rewrite fmap_insert in Hts. | ||
| 1036 | apply symmetry, fmap_insert_inv in Hts as (t2 & ts2 & Ht & ? & -> & Hts); | ||
| 1037 | simplify_eq/=; last by rewrite lookup_fmap Hx1. | ||
| 1038 | assert (∀ j, is_Some (ts2 !! j) → attr_le x j). | ||
| 1039 | { intros j. rewrite -(fmap_is_Some (AttrN ∘ thunk_to_expr)). | ||
| 1040 | rewrite -lookup_fmap -Hts lookup_fmap fmap_is_Some. auto. } | ||
| 1041 | destruct (interp_thunk _ _) as [mv|] eqn:Hinterp'; simplify_res. | ||
| 1042 | eapply interp_thunk_proper in Hinterp' | ||
| 1043 | as (mw & m1 & Hinterp1 & Hw); [|by eauto..]. | ||
| 1044 | destruct mv as [v|], mw as [w|]; simplify_res; last first. | ||
| 1045 | { exists None, m1. by rewrite map_mapM_sorted_insert //= Hinterp1. } | ||
| 1046 | destruct (force_deep _ _) as [mv'|] eqn:Hforce; simplify_res. | ||
| 1047 | eapply force_deep_proper in Hforce as (mw' & m2 & Hforce2 & Hw'); last done. | ||
| 1048 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1049 | { exists None, (m1 `max` m2). rewrite map_mapM_sorted_insert //=. | ||
| 1050 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1051 | rewrite (force_deep_le Hforce2) /=; last lia. done. } | ||
| 1052 | destruct (map_mapM_sorted _ _ _) as [mvs'|] eqn:?; simplify_res. | ||
| 1053 | eapply IH in Hts as (mws & m3 & Hmap3 & ?); last done. | ||
| 1054 | exists (<[x:=w']> <$> mws), (S (m1 `max` m2 `max` m3)). | ||
| 1055 | rewrite map_mapM_sorted_insert //=. | ||
| 1056 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1057 | rewrite (force_deep_le Hforce2) /=; last lia. | ||
| 1058 | rewrite (map_mapM_interp_le Hmap3) /=; last lia. | ||
| 1059 | destruct mvs' as [vs'|], mws as [ws'|]; simplify_res; last done. | ||
| 1060 | split; [done|]. rewrite !fmap_insert. auto 10 with f_equal. | ||
| 1061 | (* interp_proper *) | ||
| 1062 | - destruct n as [|n]; [done|]. intros Hsubst Hinterp. | ||
| 1063 | rewrite 2!subst_env_eq in Hsubst. | ||
| 1064 | rewrite interp_S in Hinterp. destruct e1, e2; simplify_res; try done. | ||
| 1065 | + (* ELit *) | ||
| 1066 | case_guard; simplify_res. | ||
| 1067 | * eexists (Some (VLit _ ltac:(done))), 1. by rewrite interp_lit. | ||
| 1068 | * exists None, 1. split; [|done]. rewrite interp_S /=. by case_guard. | ||
| 1069 | + (* EId *) | ||
| 1070 | assert (∀ (mke : option (kind * expr)) (E : env) x, | ||
| 1071 | prod_map id thunk_to_expr <$> | ||
| 1072 | union_kinded (E !! x) (prod_map id (Thunk ∅) <$> mke) | ||
| 1073 | = union_kinded (prod_map id thunk_to_expr <$> E !! x) mke) as HE. | ||
| 1074 | { intros mke' E x. destruct (E !! _) as [[[] ?]|], mke' as [[[] ?]|]; | ||
| 1075 | simplify_eq/=; rewrite ?subst_env_empty //. } | ||
| 1076 | rewrite -!HE {HE} in H. | ||
| 1077 | destruct (union_kinded (E1 !! _) _) as [[k1 t1]|], | ||
| 1078 | (union_kinded (E2 !! _) _) as [[k2 t2]|] eqn:HE2; simplify_res; last first. | ||
| 1079 | { exists None, (S n). rewrite interp_S /=. by rewrite HE2. } | ||
| 1080 | eapply interp_thunk_proper | ||
| 1081 | in Hinterp as (mw & m & Hinterp & ?); [|by eauto..]. | ||
| 1082 | exists mw, (S (n `max` m)). split; [|done]. rewrite interp_S /= HE2 /=. | ||
| 1083 | eauto using interp_thunk_le with lia. | ||
| 1084 | + (* EAbs *) eexists (Some (VClo _ _ _)), 1; split; [by rewrite interp_S|]. | ||
| 1085 | simpl. auto with f_equal. | ||
| 1086 | + (* EAbsMatch *) | ||
| 1087 | eexists (Some (VCloMatch _ _ _ _)), 1; split; [by rewrite interp_S|]. | ||
| 1088 | simpl. auto with f_equal. | ||
| 1089 | + (* EApp *) | ||
| 1090 | destruct (interp n _ e1_1) as [mv1|] eqn:Hinterp'; simplify_eq/=. | ||
| 1091 | eapply interp_proper in Hinterp' as (mw1 & m1 & Hinterp1 & ?); last done. | ||
| 1092 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1093 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1094 | destruct (interp_app n _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 1095 | eapply (interp_app_proper _ _ _ _ (Thunk _ _)) in Hinterp' | ||
| 1096 | as (mw & m2 & Hinterp2 & ?); [|done..]. | ||
| 1097 | exists mw, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1098 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1099 | rewrite (interp_app_le Hinterp2) /=; last lia. done. | ||
| 1100 | + (* ESeq *) | ||
| 1101 | destruct (interp n _ e1_1) as [mv1|] eqn:Hinterp'; simplify_eq/=. | ||
| 1102 | eapply interp_proper in Hinterp' as (mw1 & m1 & Hinterp1 & ?); last done. | ||
| 1103 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1104 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1105 | destruct μ0; simplify_res. | ||
| 1106 | { eapply interp_proper in Hinterp as (w2 & m2 & Hinterp2 & ?); last done. | ||
| 1107 | exists w2, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1108 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1109 | rewrite (interp_le Hinterp2) /=; last lia. done. } | ||
| 1110 | destruct (force_deep _ _) as [mv'|] eqn:Hforce; simplify_res. | ||
| 1111 | eapply force_deep_proper in Hforce as (mw' & m2 & Hforce & ?); last done. | ||
| 1112 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1113 | { exists None, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1114 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1115 | rewrite (force_deep_le Hforce) /=; last lia. done. } | ||
| 1116 | eapply interp_proper in Hinterp as (w2 & m3 & Hinterp3 & ?); last done. | ||
| 1117 | exists w2, (S (m1 `max` m2 `max` m3)). rewrite interp_S /=. | ||
| 1118 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1119 | rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1120 | rewrite (interp_le Hinterp3) /=; last lia. done. | ||
| 1121 | + (* EList *) | ||
| 1122 | eexists (Some (VList _)), 1; rewrite interp_S /=; split; [done|]. | ||
| 1123 | do 2 f_equal. revert es0 Hsubst. | ||
| 1124 | induction es; intros [] ?; simplify_eq/=; f_equal/=; auto. | ||
| 1125 | + (* EAttr *) | ||
| 1126 | eexists (Some (VAttr _)), 1; rewrite interp_S /=; split; [done|]. | ||
| 1127 | do 2 f_equal. apply map_eq=> x. rewrite !lookup_fmap. | ||
| 1128 | pose proof (f_equal (.!! x) Hsubst) as Hx. rewrite !lookup_fmap in Hx. | ||
| 1129 | destruct (αs !! x) as [[[]]|], (αs0 !! x) as [[[]]|]; | ||
| 1130 | simplify_eq/=; do 2 f_equal; auto. | ||
| 1131 | apply indirects_env_proper, Hx. by rewrite !tattr_to_attr_from_attr. | ||
| 1132 | + (* ELetAttr *) | ||
| 1133 | destruct (interp n _ _) as [mv'|] eqn:Hinterp'; simplify_eq/=. | ||
| 1134 | eapply interp_proper in Hinterp' as (mw' & m1 & Hinterp1 & ?); last done. | ||
| 1135 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1136 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1137 | destruct (maybe VAttr _) eqn:Hattr; simplify_res; last first. | ||
| 1138 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 1139 | by assert (maybe VAttr w' = None) as -> by (by destruct v', w'). } | ||
| 1140 | destruct v', w'; simplify_res. | ||
| 1141 | eapply interp_proper in Hinterp as (mw & m2 & Hinterp2 & ?); | ||
| 1142 | [|by apply subst_env_union_fmap_proper]. | ||
| 1143 | exists mw, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1144 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1145 | rewrite (interp_le Hinterp2) /=; last lia. done. | ||
| 1146 | + (* EBinOp *) | ||
| 1147 | destruct (interp n _ e1_1) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 1148 | eapply interp_proper in Hinterp1 as (mw1 & m1 & Hinterp1 & Hw1); last done. | ||
| 1149 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1150 | { exists None. exists (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1151 | apply (interp_bin_op_proper op0) in Hw1. | ||
| 1152 | destruct (interp_bin_op _ v1) as [f|] eqn:Hop1, | ||
| 1153 | (interp_bin_op _ w1) as [g|] eqn:Hop2; simplify_res; try done; last first. | ||
| 1154 | { exists None. exists (S m1). by rewrite interp_S /= Hinterp1 /= Hop2. } | ||
| 1155 | destruct (interp n _ e1_2) as [mv2|] eqn:Hinterp2; simplify_res. | ||
| 1156 | eapply interp_proper in Hinterp2 as (mw2 & m2 & Hinterp2 & Hw2); last done. | ||
| 1157 | destruct mv2 as [v2|], mw2 as [w2|]; simplify_res; last first. | ||
| 1158 | { exists None. exists (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1159 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1160 | rewrite Hop2 /= (interp_le Hinterp2) /=; last lia. done. } | ||
| 1161 | apply Hw1 in Hw2. | ||
| 1162 | destruct (f v2) as [t|] eqn:Hf, | ||
| 1163 | (g w2) as [t'|] eqn:Hg; simplify_res; last first. | ||
| 1164 | { exists None. exists (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1165 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1166 | rewrite Hop2 /= (interp_le Hinterp2) /=; last lia. by rewrite Hg. } | ||
| 1167 | eapply interp_thunk_proper in Hinterp | ||
| 1168 | as (mw & m3 & Hforce3 & Hw); [|by eauto..]. | ||
| 1169 | exists mw, (S (m1 `max` m2 `max` m3)). rewrite interp_S /=. split; [|done]. | ||
| 1170 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1171 | rewrite Hop2 /= (interp_le Hinterp2) /=; last lia. | ||
| 1172 | rewrite Hg /=. eauto using interp_thunk_le with lia. | ||
| 1173 | + (* EIf *) | ||
| 1174 | destruct (interp n _ e1_1) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 1175 | eapply interp_proper in Hinterp1 as (mw1 & m1 & Hinterp1 & Hw1); last done. | ||
| 1176 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1177 | { exists None. exists (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1178 | destruct (maybe_VLit _ ≫= maybe LitBool) as [b|] eqn:Hbool; | ||
| 1179 | simplify_res; last first. | ||
| 1180 | { exists None. exists (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 1181 | destruct v1, w1; repeat destruct select base_lit; naive_solver. } | ||
| 1182 | eapply (interp_proper _ _ _ _ (if b then _ else _)) in Hinterp | ||
| 1183 | as (mw & m2 & Hinterp & Hw); last by destruct b. | ||
| 1184 | exists mw, (S (m1 `max` m2)). split; [|done]. rewrite interp_S /=. | ||
| 1185 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1186 | assert (maybe_VLit w1 ≫= maybe LitBool = Some b) as ->. | ||
| 1187 | { destruct v1, w1; repeat destruct select base_lit; naive_solver. } | ||
| 1188 | rewrite /=. eauto using interp_le with lia. | ||
| 1189 | (* interp_thunk_proper *) | ||
| 1190 | - destruct n as [|n]; [done|]. | ||
| 1191 | intros Ht Hinterp. rewrite interp_thunk_S in Hinterp. | ||
| 1192 | destruct t1 as [v1|E1 e1|x1 E1 tαs1], t2 as [v2|E2 e2|x2 E2 tαs2]; simplify_res. | ||
| 1193 | + exists (Some v2), 1. rewrite interp_thunk_S /=. auto with f_equal. | ||
| 1194 | + destruct (interp_val_to_expr E2 e2 v1) as (w & m & ? & ?); first done. | ||
| 1195 | exists (Some w), (S m); simpl; auto with f_equal. | ||
| 1196 | + by destruct v1. | ||
| 1197 | + exists (Some v2), 1; split; [done|]; simpl. | ||
| 1198 | symmetry. eauto using interp_val_to_expr_Res. | ||
| 1199 | + eapply interp_proper in Hinterp as (mw & m & ? & ?); eauto. | ||
| 1200 | exists mw, (S m). eauto. | ||
| 1201 | + assert (∃ αs1, e1 = ESelect (EAttr αs1) x2 ∧ | ||
| 1202 | attr_subst_env E1 <$> αs1 = tattr_to_attr E2 <$> tαs2) as (αs1 & -> & Hαs). | ||
| 1203 | { repeat match goal with | ||
| 1204 | | H : subst_env _ ?e = _ |- _ => | ||
| 1205 | rewrite subst_env_eq in H; destruct e; simplify_eq; [] | ||
| 1206 | end; eauto. } | ||
| 1207 | clear Ht. destruct n as [|n]; [done|]. | ||
| 1208 | rewrite !interp_S /= in Hinterp. | ||
| 1209 | (* without [in Hinterp at 2 3] the termination checker loops *) | ||
| 1210 | destruct n as [|n'] in Hinterp at 2 3; [done|]. | ||
| 1211 | rewrite !interp_S /= lookup_fmap in Hinterp. | ||
| 1212 | pose proof (f_equal (.!! x2) Hαs) as Hx. rewrite !lookup_fmap in Hx. | ||
| 1213 | destruct (αs1 !! x2) as [[[] e1]|], | ||
| 1214 | (tαs2 !! x2) as [[e2|t2]|] eqn:Hx2; simplify_res. | ||
| 1215 | * rewrite -tattr_to_attr_from_attr in Hαs. | ||
| 1216 | destruct n as [|n]; [done|]. rewrite interp_thunk_S in Hinterp. | ||
| 1217 | eapply interp_proper in Hinterp as (mw & m & Hinterp & ?); | ||
| 1218 | last by apply indirects_env_proper. | ||
| 1219 | exists mw, (S m). by rewrite interp_thunk_S /= Hx2. | ||
| 1220 | * eapply interp_thunk_proper in Hinterp | ||
| 1221 | as (mw & m & Hinterp & ?); last done. | ||
| 1222 | exists mw, (S m). rewrite interp_thunk_S /= Hx2. done. | ||
| 1223 | * exists None, (S n). by rewrite interp_thunk_S /= Hx2. | ||
| 1224 | + by destruct v2. | ||
| 1225 | + assert (∃ αs2, e2 = ESelect (EAttr αs2) x1 ∧ | ||
| 1226 | attr_subst_env E2 <$> αs2 = tattr_to_attr E1 <$> tαs1) as (αs2 & -> & Hαs). | ||
| 1227 | { repeat match goal with | ||
| 1228 | | H : _ = subst_env _ ?e |- _ => | ||
| 1229 | rewrite subst_env_eq in H; destruct e; simplify_eq; [] | ||
| 1230 | end; eauto. } | ||
| 1231 | clear Ht. | ||
| 1232 | pose proof (f_equal (.!! x1) Hαs) as Hx. rewrite !lookup_fmap in Hx. | ||
| 1233 | destruct (tαs1 !! x1) as [[e1|t1]|], | ||
| 1234 | (αs2 !! x1) as [[[] e2]|] eqn:Hx2; simplify_res. | ||
| 1235 | * rewrite -tattr_to_attr_from_attr in Hαs. | ||
| 1236 | eapply interp_proper in Hinterp as (mw & m & Hinterp & ?); | ||
| 1237 | last by apply indirects_env_proper. | ||
| 1238 | exists mw, (S (S (S m))). rewrite interp_thunk_S /= !interp_S /=. | ||
| 1239 | rewrite lookup_fmap Hx2 /= interp_thunk_S /=. done. | ||
| 1240 | * apply (interp_thunk_proper _ _ (Thunk E2 e2)) | ||
| 1241 | in Hinterp as (mw & m & Hinterp & ?); last done. | ||
| 1242 | destruct m as [|m]; [done|]. | ||
| 1243 | exists mw, (S (S (S m))). rewrite interp_thunk_S /= !interp_S /=. | ||
| 1244 | rewrite lookup_fmap Hx2 /= interp_thunk_S /=. done. | ||
| 1245 | * exists None, (S (S (S n))). rewrite interp_thunk_S /= !interp_S /=. | ||
| 1246 | rewrite lookup_fmap Hx2 /=. done. | ||
| 1247 | + pose proof (f_equal (.!! x2) Ht) as Hx. rewrite !lookup_fmap in Hx. | ||
| 1248 | destruct (tαs1 !! x2) as [[e1|t1]|] eqn:Hx1, | ||
| 1249 | (tαs2 !! _) as [[e2|t2]|] eqn:Hx2; simplify_res. | ||
| 1250 | * eapply interp_proper in Hinterp | ||
| 1251 | as (mw & m & Hinterp & ?); [|by apply indirects_env_proper]. | ||
| 1252 | exists mw, (S m). rewrite interp_thunk_S /= Hx2. done. | ||
| 1253 | * eapply interp_thunk_proper in Hinterp as (mw & m & Hinterp & ?); [|done]. | ||
| 1254 | exists mw, (S m). rewrite interp_thunk_S /= Hx2. done. | ||
| 1255 | * exists None, 1. by rewrite interp_thunk_S /= Hx2. | ||
| 1256 | (* interp_app_proper *) | ||
| 1257 | - destruct n as [|n]; [done|]. | ||
| 1258 | intros Hv Ht Hinterp. rewrite interp_app_S /= in Hinterp. | ||
| 1259 | destruct v1, v2; simplify_res. | ||
| 1260 | + (* VLit *) by eexists None, 1. | ||
| 1261 | + (* VClo *) | ||
| 1262 | eapply interp_proper in Hinterp as (mw & m & Hinterp' & ?); | ||
| 1263 | last by eapply subst_env_insert_proper. | ||
| 1264 | eexists _, (S m). rewrite interp_app_S /= Hinterp'. done. | ||
| 1265 | + (* VCloMatch *) | ||
| 1266 | destruct (interp_thunk n t1') as [mv1|] eqn:Hthunk; simplify_res. | ||
| 1267 | eapply interp_thunk_proper in Hthunk as (mw1 & m1 & Hthunk & Hw); [|by eauto..]. | ||
| 1268 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1269 | { exists None, (S m1). split; [|done]. | ||
| 1270 | rewrite interp_app_S /= Hthunk /=. done. } | ||
| 1271 | destruct (maybe VAttr v1) as [ts1|] eqn:?; simplify_res; last first. | ||
| 1272 | { exists None, (S m1). split; [|done]. | ||
| 1273 | rewrite interp_app_S /= Hthunk /=. destruct v1, w1; naive_solver. } | ||
| 1274 | destruct v1, w1; simplify_eq/=. | ||
| 1275 | rewrite 2!map_fmap_compose in Hw. apply (inj _) in Hw. | ||
| 1276 | eapply (interp_match_proper _ _ _ _ _ _ strict0) in Hw; last done. | ||
| 1277 | destruct (interp_match ts1 _ _) as [tαs1|] eqn:Hmatch1, | ||
| 1278 | (interp_match ts0 _ _) as [tαs2|] eqn:Hmatch2; | ||
| 1279 | simplify_res; try done; last first. | ||
| 1280 | { exists None, (S m1). split; [|done]. | ||
| 1281 | rewrite interp_app_S /= Hthunk /= Hmatch2. done. } | ||
| 1282 | eapply interp_proper in Hinterp as (mw & m2 & Hinterp & ?); last first. | ||
| 1283 | { by apply indirects_env_proper. } | ||
| 1284 | exists mw, (S (m1 `max` m2)). split; [|done]. | ||
| 1285 | rewrite interp_app_S /=. | ||
| 1286 | rewrite (interp_thunk_le Hthunk) /=; last lia. | ||
| 1287 | rewrite Hmatch2 /=. eauto using interp_le with lia. | ||
| 1288 | + (* VList *) by eexists None, 1. | ||
| 1289 | + (* VAttr *) | ||
| 1290 | pose proof (f_equal (.!! "__functor") Hv) as Htf. | ||
| 1291 | rewrite !lookup_fmap /= in Htf. | ||
| 1292 | destruct (ts !! _) as [e|] eqn:Hfunc, (ts0 !! _) as [e'|] eqn:Hfunc'; | ||
| 1293 | simplify_res; last first. | ||
| 1294 | { exists None, 1. by rewrite interp_app_S /= Hfunc'. } | ||
| 1295 | destruct (interp_thunk _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 1296 | eapply interp_thunk_proper in Hinterp' | ||
| 1297 | as (mw' & m1 & Hinterp1 & Hw'); [|by eauto..]. | ||
| 1298 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1299 | { exists None, (S m1). by rewrite interp_app_S /= Hfunc' /= Hinterp1. } | ||
| 1300 | destruct (interp_app _ _ _) as [mv'|] eqn:Happ; simplify_res. | ||
| 1301 | eapply (interp_app_proper _ _ _ _ (Forced (VAttr _))) in Happ | ||
| 1302 | as (mw' & m2 & Happ2 & ?); [|done|by rewrite /= Hv]. | ||
| 1303 | destruct mv', mw'; simplify_res; last first. | ||
| 1304 | { exists None, (S (m1 `max` m2)). rewrite interp_app_S /= Hfunc' /=. | ||
| 1305 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1306 | rewrite (interp_app_le Happ2) /=; last lia. done. } | ||
| 1307 | eapply interp_app_proper in Hinterp as (mw' & m3 & Happ3 & ?); [|done..]. | ||
| 1308 | exists mw', (S (m1 `max` m2 `max` m3)). rewrite interp_app_S /= Hfunc' /=. | ||
| 1309 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1310 | rewrite (interp_app_le Happ2) /=; last lia. | ||
| 1311 | rewrite (interp_app_le Happ3) /=; last lia. done. | ||
| 1312 | Qed. | ||
| 1313 | |||
| 1314 | Lemma mapM_interp_proper n ts1 ts2 mvs : | ||
| 1315 | thunk_to_expr <$> ts1 = thunk_to_expr <$> ts2 → | ||
| 1316 | mapM (mbind (force_deep n) ∘ interp_thunk n) ts1 = Res mvs → | ||
| 1317 | ∃ mws m, mapM (mbind (force_deep m) ∘ interp_thunk m) ts2 = Res mws ∧ | ||
| 1318 | fmap (M:=list) val_to_expr <$> mvs = fmap (M:=list) val_to_expr <$> mws. | ||
| 1319 | Proof. eauto using mapM_interp_proper', force_deep_proper, interp_thunk_proper. Qed. | ||
| 1320 | |||
| 1321 | Lemma interp_thunk_as_interp n t mv : | ||
| 1322 | interp_thunk n t = Res mv → | ||
| 1323 | ∃ mw m, interp m ∅ (thunk_to_expr t) = Res mw ∧ | ||
| 1324 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1325 | Proof. | ||
| 1326 | revert t mv. induction n as [|n IH]; intros t mv Hinterp; [done|]. | ||
| 1327 | rewrite interp_thunk_S in Hinterp. destruct t as [v|E e|x E tαs]; simplify_res. | ||
| 1328 | { destruct (interp_empty_val_to_expr v) as (w & m & Hinterp & ?). | ||
| 1329 | exists (Some w), m; simpl; auto using f_equal. } | ||
| 1330 | { eapply interp_proper, Hinterp. by rewrite subst_env_empty. } | ||
| 1331 | destruct (tαs !! x) as [[e|t]|] eqn:Hx; simplify_res. | ||
| 1332 | - eapply interp_proper in Hinterp as (mw & m & Hinterp & ?); | ||
| 1333 | last apply subst_env_indirects_env. | ||
| 1334 | exists mw, (S (S m)). rewrite !interp_S /=. | ||
| 1335 | rewrite !lookup_fmap Hx /= interp_thunk_S /=. done. | ||
| 1336 | - apply IH in Hinterp as (mw & m & Hinterp & ?). | ||
| 1337 | exists mw, (S (S m)). rewrite !interp_S /=. | ||
| 1338 | rewrite !lookup_fmap Hx /= interp_thunk_S //=. | ||
| 1339 | - exists None, (S (S n)). rewrite !interp_S /=. | ||
| 1340 | by rewrite !lookup_fmap Hx /=. | ||
| 1341 | Qed. | ||
| 1342 | |||
| 1343 | Lemma interp_as_interp_thunk n t mv : | ||
| 1344 | interp n ∅ (thunk_to_expr t) = Res mv → | ||
| 1345 | ∃ mw m, interp_thunk m t = Res mw ∧ | ||
| 1346 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1347 | Proof. | ||
| 1348 | revert t mv. induction (lt_wf n) as [[|n] _ IH]; intros t mv Hinterp; [done|]. | ||
| 1349 | destruct t as [v|E e|x E tαs]; simplify_res. | ||
| 1350 | { apply interp_empty_val_to_expr_Res in Hinterp. by exists (Some v), 1. } | ||
| 1351 | { eapply (interp_proper _ _ E) in Hinterp as (mw & m & Hinterp & ?); | ||
| 1352 | [|by rewrite subst_env_empty]. | ||
| 1353 | exists mw, (S m). by rewrite interp_thunk_S /=. } | ||
| 1354 | destruct n as [|n]; [done|]. rewrite !interp_S /= in Hinterp. | ||
| 1355 | rewrite !lookup_fmap in Hinterp. | ||
| 1356 | destruct (tαs !! x) as [[e|t]|] eqn:Hx; simplify_res. | ||
| 1357 | - rewrite interp_thunk_S /= in Hinterp. | ||
| 1358 | eapply interp_proper in Hinterp as (mw & m & Hinterp & ?); | ||
| 1359 | last apply symmetry, subst_env_indirects_env. | ||
| 1360 | exists mw, (S m). rewrite interp_thunk_S /= Hx. done. | ||
| 1361 | - rewrite interp_thunk_S /= in Hinterp. | ||
| 1362 | eapply IH in Hinterp as (mw & m & Hinterp & ?); last lia. | ||
| 1363 | exists mw, (S m). rewrite !interp_thunk_S /= Hx. done. | ||
| 1364 | - exists None, (S n). rewrite !interp_thunk_S /= Hx. done. | ||
| 1365 | Qed. | ||
| 1366 | |||
| 1367 | Lemma delayed_interp n e e' mv : | ||
| 1368 | interp n ∅ e' = Res mv → | ||
| 1369 | e =D=> e' → | ||
| 1370 | ∃ m, interp m ∅ e = Res mv. | ||
| 1371 | Proof. | ||
| 1372 | intros Hinterp Hdel. revert n mv Hinterp. induction Hdel; intros n mv Hinterp. | ||
| 1373 | - by eauto. | ||
| 1374 | - apply IHHdel in Hinterp as [m Hinterp]. | ||
| 1375 | exists (S (S m)). rewrite interp_S /= lookup_empty left_id_L /=. | ||
| 1376 | by rewrite interp_thunk_S /=. | ||
| 1377 | - destruct n as [|n]; [done|]. rewrite interp_S /= in Hinterp. | ||
| 1378 | destruct (interp _ _ e1') as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 1379 | apply IHHdel1 in Hinterp1 as [m1 Hinterp1]. | ||
| 1380 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 1381 | { exists (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1382 | destruct (interp_bin_op op v1) as [f|] eqn:Hf; simplify_res; last first. | ||
| 1383 | { exists (S m1). by rewrite interp_S /= Hinterp1 /= Hf. } | ||
| 1384 | destruct (interp _ _ e2') as [mv2|] eqn:Hinterp2; simplify_res. | ||
| 1385 | apply IHHdel2 in Hinterp2 as [m2 Hinterp2]. exists (S (n `max` m1 `max` m2)). | ||
| 1386 | rewrite interp_S /= (interp_le Hinterp1); last lia. | ||
| 1387 | rewrite /= Hf /= (interp_le Hinterp2); last lia. | ||
| 1388 | destruct mv2; simplify_res; [|done]. | ||
| 1389 | destruct (f _); simplify_res; [|done]. | ||
| 1390 | eauto using interp_thunk_le with lia. | ||
| 1391 | - destruct n as [|n]; [done|]. rewrite interp_S /= in Hinterp. | ||
| 1392 | destruct (interp _ _ e1') as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 1393 | apply IHHdel1 in Hinterp1 as [m1 Hinterp1]. | ||
| 1394 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 1395 | { exists (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1396 | destruct (maybe_VLit v1 ≫= maybe LitBool) as [[]|] eqn: Hbool; simplify_res. | ||
| 1397 | + apply IHHdel2 in Hinterp as [m2 Hinterp2]. exists (S (m1 `max` m2)). | ||
| 1398 | rewrite interp_S /= (interp_le Hinterp1); last lia. | ||
| 1399 | rewrite /= Hbool /=. eauto using interp_le with lia. | ||
| 1400 | + apply IHHdel3 in Hinterp as [m3 Hinterp3]. exists (S (m1 `max` m3)). | ||
| 1401 | rewrite interp_S /= (interp_le Hinterp1); last lia. | ||
| 1402 | rewrite /= Hbool /=. eauto using interp_le with lia. | ||
| 1403 | + exists (S m1). rewrite interp_S /= Hinterp1 /= Hbool. done. | ||
| 1404 | Qed. | ||
| 1405 | |||
| 1406 | Lemma interp_subst_abs n x e1 e2 mv : | ||
| 1407 | interp n ∅ (subst {[x:=(ABS, e2)]} e1) = Res mv → | ||
| 1408 | ∃ mw m, interp m (<[x:=(ABS, Thunk ∅ e2)]> ∅) e1 = Res mw ∧ | ||
| 1409 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1410 | Proof. | ||
| 1411 | apply interp_proper. by rewrite subst_env_empty subst_abs_as_subst_env. | ||
| 1412 | Qed. | ||
| 1413 | |||
| 1414 | Lemma interp_subst_indirects n e αs mv : | ||
| 1415 | interp n ∅ (subst (indirects αs) e) = Res mv → | ||
| 1416 | ∃ mw m, | ||
| 1417 | interp m (indirects_env ∅ (attr_to_tattr ∅ <$> αs)) e = Res mw ∧ | ||
| 1418 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1419 | Proof. | ||
| 1420 | apply interp_proper. rewrite subst_env_empty. rewrite subst_env_alt. | ||
| 1421 | f_equal. apply map_eq=> x. rewrite !lookup_fmap. | ||
| 1422 | destruct (αs !! x) eqn:?; do 2 f_equal/=; | ||
| 1423 | rewrite /indirects /indirects_env right_id_L !map_lookup_imap | ||
| 1424 | !lookup_fmap !Heqo //=. | ||
| 1425 | rewrite tattr_to_attr_from_attr_empty //. | ||
| 1426 | Qed. | ||
| 1427 | |||
| 1428 | Lemma interp_subst_fmap k n e es mv : | ||
| 1429 | interp n ∅ (subst ((k,.) <$> es) e) = Res mv → | ||
| 1430 | ∃ mw m, interp m ((k,.) ∘ Thunk ∅ <$> es) e = Res mw ∧ | ||
| 1431 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 1432 | Proof. | ||
| 1433 | apply interp_proper. rewrite subst_env_empty. rewrite subst_env_alt. | ||
| 1434 | f_equal. apply map_eq=> x. rewrite !lookup_fmap. | ||
| 1435 | destruct (es !! x) as [d|]; do 2 f_equal/=. by rewrite subst_env_empty. | ||
| 1436 | Qed. | ||
| 1437 | |||
| 1438 | Lemma final_interp μ e : | ||
| 1439 | final μ e → | ||
| 1440 | ∃ w m, interp m ∅ e = mret w ∧ e = val_to_expr w. | ||
| 1441 | Proof. | ||
| 1442 | revert μ. induction e; intros μ'; intros Hfinal; try by inv Hfinal. | ||
| 1443 | - inv Hfinal. eexists (VLit _ _), 1. by rewrite interp_lit /=. | ||
| 1444 | - eexists (VClo _ _ _), 1. rewrite interp_S /=. split; [done|]. | ||
| 1445 | by rewrite /= subst_env_empty. | ||
| 1446 | - eexists (VCloMatch _ _ _ _), 1. rewrite interp_S /=. split; [done|]. | ||
| 1447 | rewrite /= subst_env_empty. f_equal. | ||
| 1448 | apply map_eq=> x. rewrite lookup_fmap. | ||
| 1449 | destruct (ms !! x) as [[]|]; do 2 f_equal/=. by rewrite subst_env_empty. | ||
| 1450 | - eexists (VList _), 1. rewrite interp_S /=. split; [done|]. f_equal. clear. | ||
| 1451 | induction es; f_equal/=; [|done]. | ||
| 1452 | by rewrite /= subst_env_empty. | ||
| 1453 | - eexists (VAttr _), 1. rewrite interp_S /=. split; [done|]. | ||
| 1454 | f_equal. apply map_eq=> x. | ||
| 1455 | assert (no_recs αs) by (by inv Hfinal). | ||
| 1456 | rewrite from_attr_no_recs // !lookup_fmap. | ||
| 1457 | destruct (_ !! _) as [[]|] eqn:?; f_equal/=. | ||
| 1458 | f_equal; eauto using no_recs_lookup, subst_env_empty. | ||
| 1459 | Qed. | ||
| 1460 | |||
| 1461 | Lemma final_force_deep' v : | ||
| 1462 | final DEEP (val_to_expr v) → | ||
| 1463 | ∃ w m, force_deep m v = mret w ∧ val_to_expr v = val_to_expr w. | ||
| 1464 | Proof. | ||
| 1465 | intros Hfinal. remember (val_to_expr v) as e eqn:He. | ||
| 1466 | revert v He. induction e; intros [] ?; simplify_eq/=; inv Hfinal. | ||
| 1467 | - (* VLit *) eexists (VLit _ _), 1. by rewrite force_deep_S. | ||
| 1468 | - (* VClo *) | ||
| 1469 | eexists (VClo _ _ _), 1. by rewrite force_deep_S. | ||
| 1470 | - (* VCloMatch *) | ||
| 1471 | eexists (VCloMatch _ _ _ _), 1. by rewrite force_deep_S. | ||
| 1472 | - (* VList *) | ||
| 1473 | assert (∃ vs m, mapM (mbind (force_deep m) ∘ interp_thunk m) ts = mret vs ∧ | ||
| 1474 | val_to_expr <$> vs = thunk_to_expr <$> ts) | ||
| 1475 | as (vs & m & Hmap & Hvs); last first. | ||
| 1476 | { exists (VList (Forced <$> vs)), (S m). rewrite force_deep_S /= Hmap /=. | ||
| 1477 | split; [done|]. f_equal. rewrite -Hvs. | ||
| 1478 | clear. by induction vs; by f_equal/=. } | ||
| 1479 | rewrite Forall_fmap in H1. induction H1 as [|t ts Ht ? IH]; simplify_eq/=. | ||
| 1480 | { by exists [], 0. } | ||
| 1481 | apply Forall_cons in H as [IHt IHts]. | ||
| 1482 | destruct IH as (ws & m1 & Hinterp1 & ?); simplify_eq/=; [done|]. clear IHts. | ||
| 1483 | destruct (final_interp DEEP (thunk_to_expr t)) | ||
| 1484 | as (v' & m & Hinterp & ?); [done|]. | ||
| 1485 | apply interp_as_interp_thunk in Hinterp | ||
| 1486 | as ([v''|] & m' & Hinterp & ?); simplify_res. | ||
| 1487 | destruct (IHt Ht v'') as (w & m'' & Hforce & ?); [congruence|]. | ||
| 1488 | exists (w :: ws), (m1 `max` m' `max` m''); csimpl. | ||
| 1489 | rewrite (interp_thunk_le Hinterp) /=; last lia. | ||
| 1490 | rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1491 | rewrite (mapM_interp_le Hinterp1) /=; last lia. auto with f_equal. | ||
| 1492 | - (* VAttr *) clear H1. assert (∃ vs m, | ||
| 1493 | map_mapM_sorted attr_le | ||
| 1494 | (mbind (force_deep m) ∘ interp_thunk m) ts = mret vs ∧ | ||
| 1495 | val_to_expr <$> vs = thunk_to_expr <$> ts) | ||
| 1496 | as (vs & m & Hmap & Hvs); last first. | ||
| 1497 | { exists (VAttr (Forced <$> vs)), (S m). rewrite force_deep_S /= Hmap /=. | ||
| 1498 | split; [done|]. rewrite 2!map_fmap_compose -Hvs. f_equal. | ||
| 1499 | apply map_eq=> x. rewrite !lookup_fmap. by destruct (vs !! x). } | ||
| 1500 | induction ts as [|x t ts Hx ? IH] using (map_sorted_ind attr_le). | ||
| 1501 | { exists ∅, 0. by rewrite map_mapM_sorted_empty. } | ||
| 1502 | rewrite fmap_insert /= in H, H2. | ||
| 1503 | apply map_Forall_insert in H as [IHt IHts]; last by rewrite lookup_fmap Hx. | ||
| 1504 | apply map_Forall_insert in H2 as [Ht Hts]; last by rewrite lookup_fmap Hx. | ||
| 1505 | apply IH in IHts as (ws & m1 & Hinterp1 & ?); clear IH; simplify_eq/=; last done. | ||
| 1506 | destruct (final_interp DEEP (thunk_to_expr t)) | ||
| 1507 | as (v' & m & Hinterp & ?); [done|]. | ||
| 1508 | apply interp_as_interp_thunk in Hinterp | ||
| 1509 | as ([v''|] & m' & Hinterp & ?); simplify_res. | ||
| 1510 | destruct (IHt Ht v'') as (w & m'' & Hforce & ?); [congruence|]. | ||
| 1511 | exists (<[x:=w]> ws), (m1 `max` m' `max` m''). | ||
| 1512 | rewrite fmap_insert map_mapM_sorted_insert //=. | ||
| 1513 | rewrite (interp_thunk_le Hinterp) /=; last lia. | ||
| 1514 | rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1515 | rewrite (map_mapM_interp_le Hinterp1) /=; last lia. | ||
| 1516 | rewrite fmap_insert. auto with f_equal. | ||
| 1517 | Qed. | ||
| 1518 | |||
| 1519 | Lemma interp_step μ e1 e2 : | ||
| 1520 | e1 -{μ}-> e2 → | ||
| 1521 | (∀ n mv, | ||
| 1522 | ¬final SHALLOW e1 → | ||
| 1523 | interp n ∅ e2 = Res mv → | ||
| 1524 | exists mw m, interp m ∅ e1 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw) ∧ | ||
| 1525 | (∀ n v1 v2 mv, | ||
| 1526 | μ = DEEP → | ||
| 1527 | e1 = val_to_expr v1 → | ||
| 1528 | e2 = val_to_expr v2 → | ||
| 1529 | force_deep n v2 = Res mv → | ||
| 1530 | exists mw m, force_deep m v1 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw). | ||
| 1531 | Proof. | ||
| 1532 | intros Hstep. induction Hstep; inv_step. | ||
| 1533 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1534 | apply interp_subst_abs in Hinterp as (mw & [|m] & Hinterp & Hv); simplify_eq/=. | ||
| 1535 | exists mw, (S (S (S m))). split; [|done]. | ||
| 1536 | rewrite interp_S /= interp_app_S /= /= !interp_S /=. | ||
| 1537 | rewrite -!interp_S /=. rewrite (interp_le Hinterp); last lia. done. | ||
| 1538 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1539 | destruct n as [|n]; simplify_eq/=. apply interp_match_Some_2 in H0. | ||
| 1540 | apply interp_subst_indirects in Hinterp as (mw & m & Hinterp & ?). | ||
| 1541 | exists mw, (S (S (S (S m)))); split; [|done]. | ||
| 1542 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /= (interp_S m) /=. | ||
| 1543 | rewrite from_attr_no_recs // map_fmap_compose H0 /=. | ||
| 1544 | eauto using interp_le with lia. | ||
| 1545 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1546 | destruct n as [|[|[|n]]]; simplify_eq/=. | ||
| 1547 | rewrite !interp_S /= -!interp_S in Hinterp. | ||
| 1548 | destruct (interp _ _ e1) as [mw|] eqn:Hinterp'; simplify_res. | ||
| 1549 | destruct mw as [w|]; simplify_res; last first. | ||
| 1550 | { exists None, (S (S (S (S n)))). split; [|done]. | ||
| 1551 | rewrite 2!interp_S /= interp_app_S /=. | ||
| 1552 | rewrite from_attr_no_recs // lookup_fmap H0 /=. | ||
| 1553 | rewrite interp_thunk_S /= Hinterp'. done. } | ||
| 1554 | destruct (interp_app _ _ _) as [mv'|] eqn:Happ; simplify_res. | ||
| 1555 | eapply (interp_app_proper _ _ _ _ | ||
| 1556 | (Forced (VAttr (Thunk ∅ ∘ attr_expr <$> αs)))) | ||
| 1557 | in Happ as (mw' & m1 & Happ1 & Hw); [|done|]; last first. | ||
| 1558 | { rewrite /= subst_env_eq /=. f_equal. | ||
| 1559 | apply map_eq=> y. rewrite !lookup_fmap. | ||
| 1560 | destruct (αs !! y) as [[]|] eqn:?; do 2 f_equal/=; eauto using no_recs_lookup. } | ||
| 1561 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1562 | { exists None, (S (S (S (S (n `max` m1))))). split; [|done]. | ||
| 1563 | rewrite 2!interp_S /= interp_app_S /=. | ||
| 1564 | rewrite from_attr_no_recs // lookup_fmap H0 /=. | ||
| 1565 | rewrite interp_thunk_S /= (interp_le Hinterp') /=; last lia. | ||
| 1566 | rewrite (interp_app_le Happ1); last lia. done. } | ||
| 1567 | eapply interp_app_proper in Hinterp as (mw & m2 & ? & Hinterp); [|done..]. | ||
| 1568 | exists mw, (S (S (S (S (n `max` m1 `max` m2))))). split; [|done]. | ||
| 1569 | rewrite !interp_S /= interp_app_S /=. | ||
| 1570 | rewrite from_attr_no_recs // lookup_fmap H0 /=. | ||
| 1571 | rewrite interp_thunk_S /= (interp_le Hinterp') /=; last lia. | ||
| 1572 | rewrite (interp_app_le Happ1) /=; last lia. | ||
| 1573 | eauto using interp_app_le with lia. | ||
| 1574 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1575 | destruct (final_interp μ' e1) as (v & m & Hinterp' & ->); first done. | ||
| 1576 | destruct μ'. | ||
| 1577 | { exists mv, (S (n `max` m)). rewrite interp_S /=. | ||
| 1578 | rewrite (interp_le Hinterp) /=; last lia. | ||
| 1579 | by rewrite (interp_le Hinterp') /=; last lia. } | ||
| 1580 | destruct (final_force_deep' v) as (w & m' & Hforce & ?); first done. | ||
| 1581 | exists mv, (S (n `max` m `max` m')). rewrite interp_S /=. | ||
| 1582 | rewrite (interp_le Hinterp) /=; last lia. | ||
| 1583 | rewrite (interp_le Hinterp') /=; last lia. | ||
| 1584 | by rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1585 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1586 | rewrite map_fmap_compose in Hinterp. | ||
| 1587 | apply interp_subst_fmap in Hinterp as (mw & [|m] & Hinterp & Hv); simplify_eq/=. | ||
| 1588 | rewrite map_fmap_compose in Hinterp. | ||
| 1589 | exists mw, (S (S m)). rewrite !interp_S /= -interp_S. | ||
| 1590 | rewrite from_attr_no_recs // right_id_L map_fmap_compose. done. | ||
| 1591 | - split; last first. | ||
| 1592 | { intros n [] v2 mv _ Hαs; simplify_eq/=. by destruct H. } | ||
| 1593 | intros n mv _ Hinterp. destruct n as [|n]; [done|]. | ||
| 1594 | rewrite interp_S /= in Hinterp; simplify_res. | ||
| 1595 | eexists _, 1; split; [by rewrite interp_S|]. | ||
| 1596 | do 2 f_equal/=. apply map_eq=> x /=. rewrite !lookup_fmap. | ||
| 1597 | destruct (αs !! x) as [[[] ?]|]; do 2 f_equal/=. | ||
| 1598 | by rewrite subst_env_indirects_env_attr_to_tattr_empty subst_env_empty. | ||
| 1599 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1600 | apply final_interp in H as (v1 & m1 & Hinterp1 & ->). | ||
| 1601 | pose proof H1 as Hsem. apply interp_bin_op_Some_2 in H1 as [f Hf]. | ||
| 1602 | eapply final_interp in H0 as (v2 & m2 & Hinterp2 & ->). | ||
| 1603 | eapply interp_bin_op_Some_Some_2 in H2 as (t3 & Hfv & Hdel); [|done..]. | ||
| 1604 | eapply delayed_interp in Hinterp as (m3 & Hinterp); [|done]. | ||
| 1605 | apply interp_as_interp_thunk in Hinterp as (mw & m & Hinterp3 & ?). | ||
| 1606 | exists mw, (S (m `max` m1 `max` m2)). split; [|done]. rewrite interp_S /=. | ||
| 1607 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1608 | rewrite Hf /= (interp_le Hinterp2) /=; last lia. | ||
| 1609 | rewrite Hfv /= (interp_thunk_le Hinterp3); last lia. done. | ||
| 1610 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1611 | exists mv, (S (S n)). rewrite !interp_S /= -interp_S. | ||
| 1612 | eauto using interp_le with lia. | ||
| 1613 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1614 | exists mv, (S (S n)). rewrite !interp_S /= lookup_empty /=. done. | ||
| 1615 | - split; [intros ?? []; constructor; by eauto|]. | ||
| 1616 | intros n [] [] mv _ Hts Hts' Hforce; simplify_eq. | ||
| 1617 | destruct n as [|n]; [done|rewrite force_deep_S /= in Hforce]. | ||
| 1618 | destruct (mapM _ _) as [mvs|] eqn:Hmap; simplify_eq/=. | ||
| 1619 | destruct IHHstep as [IH1 IH2]. | ||
| 1620 | apply symmetry, fmap_app_inv in Hts | ||
| 1621 | as (ts1 & [|t1 ts1'] & ? & ? & ?); simplify_eq/=. | ||
| 1622 | apply symmetry, fmap_app_inv in Hts' | ||
| 1623 | as (ts2 & [|t2 ts2'] & Hts & ? & ?); simplify_eq/=. | ||
| 1624 | assert (∃ mws m, | ||
| 1625 | mapM (mbind (force_deep m) ∘ interp_thunk m) (ts1 ++ t1 :: ts1') = Res mws ∧ | ||
| 1626 | fmap (M:=list) val_to_expr <$> mvs = fmap (M:=list) val_to_expr <$> mws) | ||
| 1627 | as (mws & m & Hmap' & Hmvs); last first. | ||
| 1628 | { exists (VList ∘ fmap Forced <$> mws), (S m). rewrite force_deep_S /= Hmap'. | ||
| 1629 | split; [done|]. | ||
| 1630 | destruct mvs as [vs|], mws as [ws|]; simplify_eq/=; do 2 f_equal. | ||
| 1631 | rewrite list_eq_Forall2 Forall2_fmap in Hmvs. | ||
| 1632 | by rewrite list_eq_Forall2 !Forall2_fmap. } | ||
| 1633 | rewrite mapM_res_app in Hmap. | ||
| 1634 | destruct (mapM _ ts2) as [mvs1|] eqn:Hmap1; simplify_res. | ||
| 1635 | eapply mapM_interp_proper in Hmap1 as (mws1 & m1 & Hmap1 & ?); [|done]. | ||
| 1636 | destruct mvs1 as [vs1|], mws1 as [ws1|]; simplify_res; last first. | ||
| 1637 | { exists None, m1. by rewrite mapM_res_app Hmap1. } | ||
| 1638 | destruct (interp_thunk n t2) as [mw|] eqn:Hinterp; simplify_res. | ||
| 1639 | apply interp_thunk_as_interp in Hinterp as (mw' & m & Hinterp & Hmw'). | ||
| 1640 | destruct (default mfail (force_deep n <$> mw)) | ||
| 1641 | as [mu|] eqn:Hforce; simplify_res. | ||
| 1642 | destruct (step_any_shallow _ _ _ Hstep) as [|Hfinal1]. | ||
| 1643 | + (* SHALLOW *) | ||
| 1644 | apply IH1 in Hinterp as (mw'' & m' & Hinterp & Hmw''); | ||
| 1645 | [|by eauto using step_not_final]. | ||
| 1646 | apply interp_as_interp_thunk in Hinterp as (mw''' & m2 & Hinterp & ?). | ||
| 1647 | destruct mw as [w|], mw', mw'', mw''' as [w'''|]; simplify_res; last first. | ||
| 1648 | { exists None, (m1 `max` m2). rewrite mapM_res_app. | ||
| 1649 | rewrite (mapM_interp_le Hmap1) /=; last lia. | ||
| 1650 | rewrite (interp_thunk_le Hinterp) /=; last lia. done. } | ||
| 1651 | eapply (force_deep_proper _ _ w''') | ||
| 1652 | in Hforce as (mu' & m3 & Hforce & ?); last congruence. | ||
| 1653 | destruct mu as [u|], mu' as [u'|]; simplify_res; last first. | ||
| 1654 | { exists None, (m1 `max` m2 `max` m3). rewrite mapM_res_app. | ||
| 1655 | rewrite (mapM_interp_le Hmap1) /=; last lia. | ||
| 1656 | rewrite (interp_thunk_le Hinterp) /=; last lia. | ||
| 1657 | rewrite (force_deep_le Hforce) /=; last lia. done. } | ||
| 1658 | destruct (mapM _ ts2') as [mvs2|] eqn:Hmap2; simplify_res. | ||
| 1659 | eapply mapM_interp_proper in Hmap2 as (mws2 & m4 & Hmap2 & ?); [|done]. | ||
| 1660 | exists ((ws1 ++.) ∘ (u' ::.) <$> mws2), (m1 `max` m2 `max` m3 `max` m4). | ||
| 1661 | rewrite mapM_res_app. | ||
| 1662 | rewrite (mapM_interp_le Hmap1) /=; last lia. | ||
| 1663 | rewrite (interp_thunk_le Hinterp) /=; last lia. | ||
| 1664 | rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1665 | rewrite (mapM_interp_le Hmap2) /=; last lia. split; [by destruct mws2|]. | ||
| 1666 | destruct mvs2, mws2; simplify_res; f_equal. rewrite !fmap_app !fmap_cons. | ||
| 1667 | congruence. | ||
| 1668 | + (* DEEP *) | ||
| 1669 | apply step_final_shallow in Hstep as Hfinal2; last done. | ||
| 1670 | apply final_interp in Hfinal1 as (w1 & m2 & Hinterpt1 & ?). | ||
| 1671 | apply interp_as_interp_thunk in Hinterpt1 as (mw'' & m3 & Hinterpt1 & ?). | ||
| 1672 | apply final_interp in Hfinal2 as (w2' & m4 & Hinterpt2 & ?). | ||
| 1673 | eapply interp_agree in Hinterp; last apply Hinterpt2. | ||
| 1674 | destruct mw as [w2|], mw'' as [w2''|]; simplify_res. | ||
| 1675 | eapply IH2 in Hforce as (mu' & m5 & Hforce & ?); [|by auto with congruence..]. | ||
| 1676 | eapply (force_deep_proper _ _ w2'') | ||
| 1677 | in Hforce as (mu'' & m6 & Hforce & ?); last congruence. | ||
| 1678 | destruct mu as [u|], mu' as [u'|], mu'' as [u''|]; simplify_res; last first. | ||
| 1679 | { exists None, (m1 `max` m3 `max` m6). rewrite mapM_res_app. | ||
| 1680 | rewrite (mapM_interp_le Hmap1) /=; last lia. | ||
| 1681 | rewrite (interp_thunk_le Hinterpt1) /=; last lia. | ||
| 1682 | rewrite (force_deep_le Hforce) /=; last lia. done. } | ||
| 1683 | destruct (mapM _ ts2') as [mvs2|] eqn:Hmap2; simplify_res. | ||
| 1684 | eapply mapM_interp_proper in Hmap2 as (mws2 & m7 & Hmap2 & ?); [|done]. | ||
| 1685 | exists ((ws1 ++.) ∘ (u'' ::.) <$> mws2), (m1 `max` m3 `max` m6 `max` m7). | ||
| 1686 | rewrite mapM_res_app. | ||
| 1687 | rewrite (mapM_interp_le Hmap1) /=; last lia. | ||
| 1688 | rewrite (interp_thunk_le Hinterpt1) /=; last lia. | ||
| 1689 | rewrite (force_deep_le Hforce) /=; last lia. | ||
| 1690 | rewrite (mapM_interp_le Hmap2) /=; last lia. split; [by destruct mws2|]. | ||
| 1691 | destruct mvs2, mws2; simplify_res; f_equal. rewrite !fmap_app !fmap_cons. | ||
| 1692 | congruence. | ||
| 1693 | - split; [intros ?? []; constructor; by eauto using no_recs_insert|]. | ||
| 1694 | intros n [] [] mv _ Hts Hts' Hforce; simplify_eq. | ||
| 1695 | destruct n as [|n]; [done|rewrite force_deep_S /= in Hforce]. | ||
| 1696 | destruct (map_mapM_sorted _ _ _) as [mvs|] eqn:Hmap; simplify_eq/=. | ||
| 1697 | destruct IHHstep as [IH1 IH2]. | ||
| 1698 | apply symmetry, fmap_insert_inv in Hts | ||
| 1699 | as (t1 & ts1 & ? & Hx1 & ? & ?); simplify_eq/=; last done. | ||
| 1700 | apply symmetry, fmap_insert_inv in Hts' as (t2 & ts2 & ? & Hx2 & ? & Hts); | ||
| 1701 | simplify_eq/=; last by rewrite lookup_fmap Hx1. | ||
| 1702 | assert (∃ mws m, | ||
| 1703 | map_mapM_sorted attr_le (mbind (force_deep m) ∘ interp_thunk m) | ||
| 1704 | (<[x:=t1]> ts1) = Res mws ∧ | ||
| 1705 | fmap (M:=gmap _) val_to_expr <$> mvs = fmap (M:=gmap _) val_to_expr <$> mws) | ||
| 1706 | as (mws & m & Hmap' & Hmvs); last first. | ||
| 1707 | { exists (VAttr ∘ fmap Forced <$> mws), (S m). rewrite force_deep_S /= Hmap'. | ||
| 1708 | split; [done|]. | ||
| 1709 | destruct mvs as [vs|], mws as [ws|]; simplify_eq/=; do 2 f_equal. | ||
| 1710 | apply map_eq=> y. rewrite !lookup_fmap. | ||
| 1711 | apply (f_equal (.!! y)) in Hmvs. rewrite !lookup_fmap in Hmvs. | ||
| 1712 | destruct (vs !! _), (ws !! _); simplify_eq/=; auto with f_equal. } | ||
| 1713 | destruct (step_any_shallow _ _ _ Hstep) as [|Hfinal]. | ||
| 1714 | + (* SHALLOW *) assert (map_Forall2 (λ _ t1 t2, ∀ n mv, | ||
| 1715 | interp n ∅ (thunk_to_expr t2) = Res mv → | ||
| 1716 | ∃ mw m, interp m ∅ (thunk_to_expr t1) = Res mw ∧ | ||
| 1717 | val_to_expr <$> mv = val_to_expr <$> mw) | ||
| 1718 | (<[x:=t1]> ts1) (<[x:=t2]> ts2)) as IHts. | ||
| 1719 | { apply map_Forall2_insert_2; first by eauto using step_not_final. | ||
| 1720 | intros y. apply (f_equal (.!! y)) in Hts. rewrite !lookup_fmap in Hts. | ||
| 1721 | destruct (ts1 !! y), (ts2 !! y); simplify_eq/=; constructor. | ||
| 1722 | rewrite -Hts; eauto. } | ||
| 1723 | revert IHts Hmap. generalize (<[x:=t1]> ts1) (<[x:=t2]> ts2). clear. | ||
| 1724 | intros ts1. revert n mvs. | ||
| 1725 | induction ts1 as [|x t1 ts1 ?? IH] using (map_sorted_ind attr_le); | ||
| 1726 | intros n mvs ts2' IHts Hmap. | ||
| 1727 | { apply map_Forall2_empty_inv_l in IHts as ->. | ||
| 1728 | rewrite map_mapM_sorted_empty in Hmap; simplify_res. | ||
| 1729 | by exists (Some ∅), 1. } | ||
| 1730 | apply map_Forall2_insert_inv_l in IHts | ||
| 1731 | as (t2 & ts2 & -> & ? & IHt & IHts); simplify_eq/=; last done. | ||
| 1732 | assert (∀ j, is_Some (ts2 !! j) → attr_le x j). | ||
| 1733 | { apply map_Forall2_dom_L in IHts. intros j. | ||
| 1734 | rewrite -elem_of_dom -IHts elem_of_dom. auto. } | ||
| 1735 | rewrite map_mapM_sorted_insert //= in Hmap. | ||
| 1736 | destruct (interp_thunk _ _) as [mv|] eqn:Hinterp; simplify_res. | ||
| 1737 | assert (∃ mw m, interp_thunk m t1 = Res mw ∧ | ||
| 1738 | val_to_expr <$> mv = val_to_expr <$> mw) as (mw & m1 & Hinterp1 & ?). | ||
| 1739 | { apply interp_thunk_as_interp in Hinterp as (mw & m & Hinterp & ?). | ||
| 1740 | apply IHt in Hinterp as (mw' & m' & Hinterp & ?). | ||
| 1741 | eapply interp_as_interp_thunk in Hinterp as (mw'' & m'' & Hinterp & ?). | ||
| 1742 | exists mw'', m''. eauto with congruence. } | ||
| 1743 | destruct mv as [v|], mw as [w|]; simplify_res; last first. | ||
| 1744 | { exists None, m1. split; [|done]. rewrite map_mapM_sorted_insert //=. | ||
| 1745 | by rewrite Hinterp1. } | ||
| 1746 | destruct (force_deep _ _) as [mv|] eqn:Hforce; simplify_res. | ||
| 1747 | eapply force_deep_proper in Hforce as (mw & m2 & Hforce' & ?); last done. | ||
| 1748 | destruct mv as [v'|], mw as [w'|]; simplify_res; last first. | ||
| 1749 | { exists None, (m1 `max` m2). split; [|done]. | ||
| 1750 | rewrite map_mapM_sorted_insert //=. | ||
| 1751 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1752 | rewrite (force_deep_le Hforce') /=; last lia. done. } | ||
| 1753 | destruct (map_mapM_sorted _ _ _) as [mvs'|] eqn:Hmap'; simplify_res. | ||
| 1754 | apply IH in Hmap' as (mws & m3 & Hmap3 & ?); last done. | ||
| 1755 | exists (fmap <[x:=w']> mws), (m1 `max` m2 `max` m3). | ||
| 1756 | rewrite map_mapM_sorted_insert //=. | ||
| 1757 | rewrite (interp_thunk_le Hinterp1) /=; last lia. | ||
| 1758 | rewrite (force_deep_le Hforce') /=; last lia. | ||
| 1759 | rewrite (map_mapM_interp_le Hmap3) /=; last lia. | ||
| 1760 | destruct mvs', mws; simplify_res; last done. | ||
| 1761 | rewrite !fmap_insert. auto with f_equal. | ||
| 1762 | + (* DEEP *) | ||
| 1763 | assert (map_Forall2 (λ _ t1 t2, | ||
| 1764 | thunk_to_expr t1 = thunk_to_expr t2 ∨ | ||
| 1765 | ∃ v1 v2, | ||
| 1766 | thunk_to_expr t1 = val_to_expr v1 ∧ | ||
| 1767 | thunk_to_expr t2 = val_to_expr v2 ∧ | ||
| 1768 | ∀ n mv, | ||
| 1769 | force_deep n v2 = Res mv → | ||
| 1770 | ∃ mw m, force_deep m v1 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw) | ||
| 1771 | (<[x:=t1]> ts1) (<[x:=t2]> ts2)) as IHts. | ||
| 1772 | { apply map_Forall2_insert_2; last first. | ||
| 1773 | { intros y. apply (f_equal (.!! y)) in Hts. rewrite !lookup_fmap in Hts. | ||
| 1774 | destruct (ts1 !! y), (ts2 !! y); simplify_eq/=; constructor; eauto. } | ||
| 1775 | assert (final SHALLOW (thunk_to_expr t2)) | ||
| 1776 | as (v2 & m2 & Hinterp2 & Ht2)%final_interp | ||
| 1777 | by eauto using step_final_shallow. | ||
| 1778 | apply final_interp in Hfinal as (v1 & m1 & Hinterp1 & Ht1); eauto 10. } | ||
| 1779 | revert IHts Hmap. generalize (<[x:=t1]> ts1) (<[x:=t2]> ts2). clear. | ||
| 1780 | intros ts1. revert n mvs. | ||
| 1781 | induction ts1 as [|x t1 ts1 ?? IH] using (map_sorted_ind attr_le); | ||
| 1782 | intros n mvs ts2' IHts Hmap. | ||
| 1783 | { apply map_Forall2_empty_inv_l in IHts as ->. | ||
| 1784 | rewrite map_mapM_sorted_empty in Hmap; simplify_res. | ||
| 1785 | by exists (Some ∅), 1. } | ||
| 1786 | apply map_Forall2_insert_inv_l in IHts | ||
| 1787 | as (t2 & ts2 & -> & ? & IHt & IHts); simplify_eq/=; last done. | ||
| 1788 | assert (∀ j, is_Some (ts2 !! j) → attr_le x j). | ||
| 1789 | { apply map_Forall2_dom_L in IHts. intros j. | ||
| 1790 | rewrite -elem_of_dom -IHts elem_of_dom. auto. } | ||
| 1791 | rewrite map_mapM_sorted_insert //= in Hmap. | ||
| 1792 | destruct (interp_thunk n t2 ≫= force_deep n) | ||
| 1793 | as [mv|] eqn:Hinterp; simplify_res. | ||
| 1794 | assert (∃ mw m, interp_thunk m t1 ≫= force_deep m = Res mw ∧ | ||
| 1795 | val_to_expr <$> mv = val_to_expr <$> mw) as (mw & m1 & Hinterp1 & ?). | ||
| 1796 | { destruct (interp_thunk _ _) as [mv'|] eqn:Hthunk; simplify_res. | ||
| 1797 | destruct IHt as [|(v1 & v2 & Ht1 & Ht2 & IHt)]. | ||
| 1798 | * eapply interp_thunk_proper in Hthunk | ||
| 1799 | as (mw' & m1 & Hthunk1 & ?); last done. | ||
| 1800 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1801 | { exists None, m1. by rewrite Hthunk1. } | ||
| 1802 | eapply force_deep_proper in Hinterp | ||
| 1803 | as (mw & m2 & Hforce2 & ?); last done. | ||
| 1804 | exists mw, (m1 `max` m2). split; [|done]. | ||
| 1805 | rewrite (interp_thunk_le Hthunk1) /=; last lia. | ||
| 1806 | eauto using force_deep_le with lia. | ||
| 1807 | * destruct (interp_empty_val_to_expr v1) as (v1' & m1 & Hinterp1 & ?). | ||
| 1808 | rewrite -Ht1 in Hinterp1. | ||
| 1809 | eapply interp_as_interp_thunk in Hinterp1 | ||
| 1810 | as ([v1''|] & m1' & Hthunk1 & ?); simplify_res. | ||
| 1811 | eapply (interp_thunk_proper _ _ (Forced v2)) in Hthunk | ||
| 1812 | as (mw2 & m2 & Hthunk2 & ?); simplify_res; [|done]. | ||
| 1813 | destruct m2 as [|m2]; [done|]. | ||
| 1814 | rewrite interp_thunk_S in Hthunk2; simplify_res. | ||
| 1815 | destruct mv' as [v2'|]; simplify_res. | ||
| 1816 | eapply force_deep_proper in Hinterp | ||
| 1817 | as (mv' & m2' & Hforce2 & ?); last done. | ||
| 1818 | eapply IHt in Hforce2 as (mw' & m2'' & Hforce2 & ?). | ||
| 1819 | eapply (force_deep_proper _ _ v1'') in Hforce2 | ||
| 1820 | as (mw'' & m2''' & Hforce2 & ?); last congruence. | ||
| 1821 | exists mw'', (m1' `max` m2'''). | ||
| 1822 | rewrite (interp_thunk_le Hthunk1) /=; last lia. | ||
| 1823 | rewrite (force_deep_le Hforce2) /=; last lia. auto with congruence. } | ||
| 1824 | destruct mv as [v|], mw as [w|]; simplify_res; last first. | ||
| 1825 | { exists None, m1. split; [|done]. rewrite map_mapM_sorted_insert //=. | ||
| 1826 | by rewrite Hinterp1. } | ||
| 1827 | destruct (map_mapM_sorted _ _ _) as [mvs'|] eqn:Hmap'; simplify_res. | ||
| 1828 | apply IH in Hmap' as (mws & m2 & Hmap2 & ?); last done. | ||
| 1829 | exists (fmap <[x:=w]> mws), (m1 `max` m2). | ||
| 1830 | rewrite map_mapM_sorted_insert //=. | ||
| 1831 | destruct (interp_thunk m1 t1) as [[]|] eqn:Hinterp'; simplify_res. | ||
| 1832 | rewrite (interp_thunk_le Hinterp') /=; last lia. | ||
| 1833 | rewrite (force_deep_le Hinterp1) /=; last lia. | ||
| 1834 | rewrite (map_mapM_interp_le Hmap2) /=; last lia. | ||
| 1835 | destruct mvs', mws; simplify_res; last done. | ||
| 1836 | rewrite !fmap_insert. auto with f_equal. | ||
| 1837 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1838 | destruct n as [|n]; simplify_eq/=. | ||
| 1839 | rewrite interp_S /= in Hinterp. | ||
| 1840 | destruct (interp n ∅ e') as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 1841 | apply IHHstep in Hinterp' | ||
| 1842 | as (mw' & m1 & Hinterp1 & ?); last by eauto using step_not_final. | ||
| 1843 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1844 | { exists None, (S m1). split; [|done]. by rewrite interp_S /= Hinterp1. } | ||
| 1845 | eapply interp_app_proper in Hinterp as (mw & m2 & Happ2 & ?); [|done..]. | ||
| 1846 | exists mw, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 1847 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1848 | rewrite (interp_app_le Happ2) /=; last lia. done. | ||
| 1849 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1850 | destruct n as [|[|[|n]]]; simplify_eq/=. | ||
| 1851 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /= in Hinterp. | ||
| 1852 | destruct (interp n ∅ e') as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 1853 | apply IHHstep in Hinterp' | ||
| 1854 | as (mw' & m1 & Hinterp1 & Hw'); last by eauto using step_not_final. | ||
| 1855 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1856 | { exists None, (S (S (S m1))). split; [|done]. | ||
| 1857 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 1858 | by rewrite Hinterp1. } | ||
| 1859 | destruct (maybe VAttr v') as [ts|] eqn:?; simplify_res; last first. | ||
| 1860 | { exists None, (S (S (S m1))). split; [|done]. | ||
| 1861 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /= Hinterp1 /=. | ||
| 1862 | assert (maybe VAttr w' = None) as ->; [|done]. | ||
| 1863 | destruct v', w'; naive_solver. } | ||
| 1864 | destruct v', w'; simplify_eq/=. | ||
| 1865 | rewrite 2!map_fmap_compose in Hw'. apply (inj _) in Hw'. | ||
| 1866 | eapply (interp_match_proper ∅ ∅ _ _ ms ms strict) in Hw'; [|done]. | ||
| 1867 | destruct (interp_match ts _ strict) as [tαs1|] eqn:Hmatch1, | ||
| 1868 | (interp_match ts1 _ strict) as [tαs2|] eqn:Hmatch2; | ||
| 1869 | simplify_res; try done; last first. | ||
| 1870 | { exists None, (S (S (S m1))). split; [|done]. | ||
| 1871 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 1872 | rewrite Hinterp1 /= Hmatch2. done. } | ||
| 1873 | eapply interp_proper in Hinterp | ||
| 1874 | as (mw & m2 & Hinterp & ?); last first. | ||
| 1875 | { by apply indirects_env_proper. } | ||
| 1876 | exists mw, (S (S (S (m1 `max` m2)))). split; [|done]. | ||
| 1877 | rewrite !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 1878 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1879 | rewrite Hmatch2 /=. eauto using interp_le with lia. | ||
| 1880 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1881 | destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 1882 | destruct (interp n _ e') as [mv'|] eqn:Hinterp'; simplify_eq/=. | ||
| 1883 | destruct (step_any_shallow μ e e') as [|Hfinal]; first done. | ||
| 1884 | + apply IHHstep in Hinterp' | ||
| 1885 | as (mw' & m & Hinterp' & Hw); last by eauto using step_not_final. | ||
| 1886 | destruct mv' as [v|], mw' as [w'|]; simplify_res; last first. | ||
| 1887 | { exists None, (S m). by rewrite interp_S /= Hinterp'. } | ||
| 1888 | destruct μ; simplify_res. | ||
| 1889 | { exists mv, (S (n `max` m)). rewrite interp_S /=. | ||
| 1890 | rewrite (interp_le Hinterp') /=; last lia. | ||
| 1891 | rewrite (interp_le Hinterp) /=; last lia. done. } | ||
| 1892 | destruct (force_deep n v) as [mv'|] eqn:Hforce; simplify_res. | ||
| 1893 | eapply force_deep_proper | ||
| 1894 | in Hforce as (mw' & m2 & Hforce2 & ?); last done. | ||
| 1895 | exists mv, (S (n `max` m `max` m2)). split; [|done]. rewrite interp_S /=. | ||
| 1896 | rewrite (interp_le Hinterp') /=; last lia. | ||
| 1897 | rewrite (force_deep_le Hforce2) /=; last lia. | ||
| 1898 | destruct mv', mw'; simplify_res; eauto using interp_le with lia. | ||
| 1899 | + destruct μ; [by odestruct step_not_final|]. | ||
| 1900 | assert (final SHALLOW e') as (w & m & Hinterp'' & ->)%final_interp | ||
| 1901 | by eauto using step_final_shallow. | ||
| 1902 | apply interp_empty_val_to_expr_Res in Hinterp'. | ||
| 1903 | destruct mv' as [v|]; simplify_res. | ||
| 1904 | destruct (force_deep n v) as [mv'|] eqn:Hforce; simplify_res. | ||
| 1905 | apply final_interp in Hfinal as (w' & m' & Hinterp''' & ->). | ||
| 1906 | eapply IHHstep in Hforce as (mw' & m'' & Hforce' & ?); [|done..]. | ||
| 1907 | exists mv, (S (n `max` m' `max` m'')). rewrite interp_S /=. | ||
| 1908 | rewrite (interp_le Hinterp''') /=; last lia. | ||
| 1909 | rewrite (force_deep_le Hforce') /=; last lia. | ||
| 1910 | destruct mv', mw'; simplify_res; eauto using interp_le with lia. | ||
| 1911 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1912 | destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 1913 | destruct (interp n _ _) as [mv'|] eqn:Hinterp'; simplify_eq/=. | ||
| 1914 | apply IHHstep in Hinterp' | ||
| 1915 | as (mw' & m1 & Hinterp1 & Hw); last by eauto using step_not_final. | ||
| 1916 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 1917 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 1918 | destruct (maybe VAttr _) eqn:Hattr; simplify_res; last first. | ||
| 1919 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 1920 | by assert (maybe VAttr w' = None) as -> by (by destruct v', w'). } | ||
| 1921 | destruct v', w'; simplify_res. | ||
| 1922 | rewrite right_id_L in Hinterp. | ||
| 1923 | eapply interp_proper in Hinterp as (mw & m2 & Hinterp2 & ?); | ||
| 1924 | last by apply subst_env_fmap_proper. | ||
| 1925 | exists mw, (S (m1 `max` m2)). rewrite !interp_S /=. | ||
| 1926 | rewrite (interp_le Hinterp1) /=; last lia. rewrite right_id_L. | ||
| 1927 | by rewrite (interp_le Hinterp2) /=; last lia. | ||
| 1928 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1929 | destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 1930 | destruct (interp n _ e') as [mv1|] eqn:Hinterp1; simplify_eq/=. | ||
| 1931 | apply IHHstep in Hinterp1 as (mw1 & m & Hinterp1 & Hw1); | ||
| 1932 | last by eauto using step_not_final. | ||
| 1933 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1934 | { exists None, (S m). by rewrite interp_S /= Hinterp1. } | ||
| 1935 | apply (interp_bin_op_proper op) in Hw1. | ||
| 1936 | destruct (interp_bin_op _ v1) as [f|] eqn:Hopf; simplify_res; last first. | ||
| 1937 | { exists None, (S m). rewrite interp_S /= Hinterp1 /=. | ||
| 1938 | by destruct (interp_bin_op _ w1). } | ||
| 1939 | destruct (interp_bin_op _ w1) as [g|] eqn:Hopg; simplify_res; [|done]. | ||
| 1940 | destruct (interp n _ e2) as [mv2|] eqn:Hinterp2; simplify_res. | ||
| 1941 | destruct mv2 as [v2|]; simplify_res; last first. | ||
| 1942 | { exists None, (S (n `max` m)). split; [|done]. rewrite interp_S /=. | ||
| 1943 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1944 | rewrite (interp_le Hinterp2) /=; last lia. by rewrite Hopg. } | ||
| 1945 | specialize (Hw1 v2 _ eq_refl). | ||
| 1946 | destruct (f v2) as [t2|], (g v2) as [t2'|] eqn:Hg; simplify_res; last first. | ||
| 1947 | { exists None, (S (n `max` m)). split; [|done]. rewrite interp_S /=. | ||
| 1948 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1949 | rewrite (interp_le Hinterp2) /=; last lia. by rewrite Hopg /= Hg. } | ||
| 1950 | eapply interp_thunk_proper in Hinterp as (mw & m' & Hthunk & ?); last done. | ||
| 1951 | exists mw, (S (n `max` m `max` m')). split; [|done]. rewrite interp_S /=. | ||
| 1952 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1953 | rewrite (interp_le Hinterp2) /=; last lia. rewrite Hopg /= Hg /=. | ||
| 1954 | rewrite (interp_thunk_le Hthunk) /=; last lia. done. | ||
| 1955 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1956 | destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 1957 | destruct (interp n ∅ e1) as [mw1|] eqn:Hinterp1; simplify_res. | ||
| 1958 | apply final_interp in H0 as (v1 & m1 & Hinterp1' & ->). | ||
| 1959 | apply interp_bin_op_Some_2 in H1 as [f Hop]. | ||
| 1960 | assert (mw1 = Some v1) as -> by eauto using interp_agree. | ||
| 1961 | rewrite /= Hop /= in Hinterp. | ||
| 1962 | destruct (interp _ _ e') as [mv2|] eqn:Hinterp2; simplify_res; last first. | ||
| 1963 | apply IHHstep in Hinterp2 as (mw2 & m & Hinterp2 & Hw); | ||
| 1964 | last by eauto using step_not_final. | ||
| 1965 | destruct mv2 as [v2|], mw2 as [w2|]; simplify_res; last first. | ||
| 1966 | { exists None, (S (n `max` m)). split; [|done]. rewrite interp_S /=. | ||
| 1967 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1968 | rewrite (interp_le Hinterp2) /=; last lia. by rewrite Hop. } | ||
| 1969 | pose proof @eq_refl as Hf%(interp_bin_op_proper op v1). rewrite !Hop in Hf. | ||
| 1970 | apply Hf in Hw; clear Hf. | ||
| 1971 | destruct (f v2) as [t|] eqn:Hf, | ||
| 1972 | (f w2) as [t'|] eqn:Hf'; simplify_res; last first. | ||
| 1973 | { exists None, (S (n `max` m)). split; [|done]. rewrite interp_S /=. | ||
| 1974 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1975 | rewrite (interp_le Hinterp2) /=; last lia. by rewrite Hop /= Hf'. } | ||
| 1976 | eapply interp_thunk_proper in Hinterp as (mw & m' & Hthunk & ?); last done. | ||
| 1977 | exists mw, (S (n `max` m `max` m')). split; [|done]. rewrite interp_S /=. | ||
| 1978 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 1979 | rewrite (interp_le Hinterp2) /=; last lia. rewrite Hop /= Hf' /=. | ||
| 1980 | eauto using interp_thunk_le with lia. | ||
| 1981 | - split; [|by intros ? []]. intros n mv _ Hinterp. | ||
| 1982 | destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 1983 | destruct (interp n _ e') as [mv1|] eqn:Hinterp1; simplify_eq/=. | ||
| 1984 | apply IHHstep in Hinterp1 as (mw1 & m & Hinterp1 & Hw1); | ||
| 1985 | last by eauto using step_not_final. | ||
| 1986 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 1987 | { exists None, (S m). by rewrite interp_S /= Hinterp1. } | ||
| 1988 | exists mv, (S (n `max` m)). split; [|done]. | ||
| 1989 | rewrite interp_S /= (interp_le Hinterp1) /=; last lia. | ||
| 1990 | assert (maybe_VLit w1 ≫= maybe LitBool = maybe_VLit v1 ≫= maybe LitBool) as ->. | ||
| 1991 | { destruct v1, w1; repeat destruct select base_lit; naive_solver. } | ||
| 1992 | destruct (maybe_VLit v1 ≫= maybe LitBool); simplify_res; [|done]. | ||
| 1993 | eauto using interp_le with lia. | ||
| 1994 | Qed. | ||
| 1995 | |||
| 1996 | Lemma final_interp' μ e : | ||
| 1997 | final μ e → | ||
| 1998 | ∃ w m, interp' m μ ∅ e = mret w ∧ e = val_to_expr w. | ||
| 1999 | Proof. | ||
| 2000 | intros Hfinal. destruct (final_interp _ _ Hfinal) as (w & m & Hinterp & ->). | ||
| 2001 | destruct μ. | ||
| 2002 | { exists w, m. by rewrite interp_shallow'. } | ||
| 2003 | apply final_force_deep' in Hfinal as (w' & m' & Hforce & ?). | ||
| 2004 | exists w', (m `max` m'); split; [|done]. rewrite /interp'. | ||
| 2005 | rewrite (interp_le Hinterp) /=; last lia. eauto using force_deep_le with lia. | ||
| 2006 | Qed. | ||
| 2007 | |||
| 2008 | Lemma force_deep_le' {n1 n2 μ v mv} : | ||
| 2009 | force_deep' n1 μ v = Res mv → n1 ≤ n2 → force_deep' n2 μ v = Res mv. | ||
| 2010 | Proof. destruct μ; eauto using force_deep_le. Qed. | ||
| 2011 | |||
| 2012 | Lemma interp_le' {n1 n2 μ E e mv} : | ||
| 2013 | interp' n1 μ E e = Res mv → n1 ≤ n2 → interp' n2 μ E e = Res mv. | ||
| 2014 | Proof. | ||
| 2015 | rewrite /interp'. intros. | ||
| 2016 | destruct (interp n1 _ _) as [mw|] eqn:Hinterp; simplify_res. | ||
| 2017 | rewrite (interp_le Hinterp); last lia. | ||
| 2018 | destruct mw; simplify_res; eauto using force_deep_le'. | ||
| 2019 | Qed. | ||
| 2020 | |||
| 2021 | Lemma interp_agree' {n1 n2 μ E e mv1 mv2} : | ||
| 2022 | interp' n1 μ E e = Res mv1 → interp' n2 μ E e = Res mv2 → mv1 = mv2. | ||
| 2023 | Proof. | ||
| 2024 | intros He1 He2. apply (inj Res). destruct (total (≤) n1 n2). | ||
| 2025 | - rewrite -He2. symmetry. eauto using interp_le'. | ||
| 2026 | - rewrite -He1. eauto using interp_le'. | ||
| 2027 | Qed. | ||
| 2028 | |||
| 2029 | Lemma interp_step' n μ e1 e2 mv : | ||
| 2030 | e1 -{μ}-> e2 → | ||
| 2031 | interp' n μ ∅ e2 = Res mv → | ||
| 2032 | ∃ mw m, interp' m μ ∅ e1 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 2033 | Proof. | ||
| 2034 | intros Hstep. destruct μ. | ||
| 2035 | { setoid_rewrite interp_shallow'. | ||
| 2036 | eapply interp_step; eauto using step_not_final. } | ||
| 2037 | intros Hinterp. rewrite /interp' in Hinterp. | ||
| 2038 | destruct (interp n ∅ e2) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 2039 | destruct (step_any_shallow _ _ _ Hstep) as [|Hfinal]. | ||
| 2040 | - eapply interp_step in Hinterp' as (mw' & m & Hinterp' & ?); | ||
| 2041 | [|by eauto using step_not_final..]. | ||
| 2042 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 2043 | { exists None, m. by rewrite /interp' Hinterp'. } | ||
| 2044 | eapply force_deep_proper in Hinterp as (mw' & m' & Hforce & ?); last done. | ||
| 2045 | exists mw', (m `max` m'). rewrite /interp'. | ||
| 2046 | rewrite (interp_le Hinterp') /=; last lia. | ||
| 2047 | eauto using force_deep_le with lia. | ||
| 2048 | - assert (final SHALLOW e2) | ||
| 2049 | as (w2 & m2 & Hinterpw2 & ->)%final_interp by eauto using step_final_shallow. | ||
| 2050 | apply final_interp in Hfinal as (w1 & m1 & Hinterpw1 & ->). | ||
| 2051 | apply interp_empty_val_to_expr_Res in Hinterp'; destruct mv'; simplify_res. | ||
| 2052 | eapply interp_step in Hstep as [_ Hstep]. | ||
| 2053 | eapply Hstep in Hinterp as (mw & m & Hforce & ?); [|done..]. | ||
| 2054 | exists mw, (m `max` m1). split; [|done]. rewrite /interp'. | ||
| 2055 | rewrite (interp_le Hinterpw1) /=; last lia. | ||
| 2056 | eauto using force_deep_le with lia. | ||
| 2057 | Qed. | ||
| 2058 | |||
| 2059 | Lemma final_val_to_expr' n μ E e v : | ||
| 2060 | interp' n μ E e = mret v → final μ (val_to_expr v). | ||
| 2061 | Proof. | ||
| 2062 | rewrite /interp'. intros Hinterp. | ||
| 2063 | destruct (interp _ _ e) as [[w|]|] eqn:Hinterp'; simplify_res. | ||
| 2064 | destruct μ; simplify_res; eauto using final_force_deep. | ||
| 2065 | Qed. | ||
| 2066 | |||
| 2067 | Lemma red_final_interp μ e : | ||
| 2068 | red (step μ) e ∨ final μ e ∨ ∃ m, interp' m μ ∅ e = mfail. | ||
| 2069 | Proof. | ||
| 2070 | revert μ. induction e; intros μ'. | ||
| 2071 | - (* ELit *) | ||
| 2072 | destruct (decide (base_lit_ok b)). | ||
| 2073 | + right; left. by constructor. | ||
| 2074 | + do 2 right. exists 1. rewrite /interp' interp_S /=. by case_guard. | ||
| 2075 | - (* EId *) destruct mkd as [[k d]|]. | ||
| 2076 | + left. eexists; constructor. | ||
| 2077 | + do 2 right. by exists 1. | ||
| 2078 | - (* EAbs *) right; left. constructor. | ||
| 2079 | - (* EAbsMatch *) right; left. constructor. | ||
| 2080 | - (* EApp *) destruct (IHe1 SHALLOW) as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 2081 | + left. eexists. by eapply SAppL. | ||
| 2082 | + apply final_interp in Hfinal as ([] & m & _ & ->); simplify_res. | ||
| 2083 | { do 2 right. exists 3. rewrite /interp' interp_S /= interp_lit //. } | ||
| 2084 | { left. by repeat econstructor. } | ||
| 2085 | { destruct (IHe2 SHALLOW) as [[??]|[Hfinal|[m2 Hinterp2]]]. | ||
| 2086 | * left. by repeat econstructor. | ||
| 2087 | * apply final_interp in Hfinal as (w2 & m2 & Hinterp2 & ->). | ||
| 2088 | destruct (maybe VAttr w2) as [ts|] eqn:Hw2; last first. | ||
| 2089 | { do 2 right. exists (S (S (S m2))). | ||
| 2090 | rewrite /interp' !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 2091 | rewrite Hinterp2 /= Hw2. done. } | ||
| 2092 | destruct w2; simplify_eq/=. | ||
| 2093 | destruct (interp_match ts (fmap (M:=option) (subst_env E) <$> ms) strict) | ||
| 2094 | as [E'|] eqn:Hmatch; last first. | ||
| 2095 | { do 2 right. exists (S (S (S m2))). | ||
| 2096 | rewrite /interp' !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 2097 | rewrite Hinterp2 /= Hmatch. done. } | ||
| 2098 | apply interp_match_Some_1 in Hmatch. | ||
| 2099 | left. repeat econstructor; [done|]. | ||
| 2100 | by rewrite map_fmap_compose fmap_attr_expr_Attr. | ||
| 2101 | * rewrite interp_shallow' in Hinterp2. | ||
| 2102 | do 2 right. exists (S (S (S m2))). | ||
| 2103 | rewrite /interp' !interp_S /= interp_app_S /= interp_thunk_S /=. | ||
| 2104 | by rewrite Hinterp2. } | ||
| 2105 | { do 2 right. by exists 3. } | ||
| 2106 | destruct (ts !! "__functor") as [e|] eqn:Hfunc. | ||
| 2107 | { left. repeat econstructor; by simplify_map_eq. } | ||
| 2108 | do 2 right. exists (S (S m)). rewrite /interp' !interp_S /=. | ||
| 2109 | rewrite interp_app_S /= !lookup_fmap Hfunc. done. | ||
| 2110 | + rewrite interp_shallow' in Hinterp. | ||
| 2111 | do 2 right. exists (S m). by rewrite /interp' interp_S /= Hinterp. | ||
| 2112 | - (* ESeq *) destruct (IHe1 μ) as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 2113 | + left. eexists. by eapply SSeq. | ||
| 2114 | + left. by repeat econstructor. | ||
| 2115 | + do 2 right. exists (S m). rewrite /interp' interp_S /=. | ||
| 2116 | rewrite /interp' in Hinterp. | ||
| 2117 | destruct (interp _ _ e1) as [[]|], μ; simplify_res; [|done..]. | ||
| 2118 | by rewrite Hinterp. | ||
| 2119 | - (* EList *) | ||
| 2120 | destruct μ'. | ||
| 2121 | { right; left. by constructor. } | ||
| 2122 | assert (red (step DEEP) (EList es) ∨ Forall (final DEEP) es ∨ | ||
| 2123 | ∃ m, mapM (mbind (force_deep m) ∘ interp_thunk m) | ||
| 2124 | (Thunk ∅ <$> es) = mfail) as Hhelp; last first. | ||
| 2125 | { destruct Hhelp as [?|[?|[m Hinterp]]]; [by auto using final..|]. | ||
| 2126 | do 2 right. exists (S m). rewrite /interp' interp_S /=. | ||
| 2127 | rewrite force_deep_S /=. by rewrite Hinterp. } | ||
| 2128 | induction H as [|e es He Hes IH]; [by right; left|]. | ||
| 2129 | destruct (He DEEP) as [[??]|[Hfinal|[m Hinterp]]]; simplify_eq/=. | ||
| 2130 | + left. eexists. by eapply (SList []). | ||
| 2131 | + destruct IH as [[??]|[?|[m2 Hinterp2]]]; [|by eauto|]. | ||
| 2132 | * left. inv_step. eexists. eapply (SList (_ :: _)); by eauto. | ||
| 2133 | * apply final_interp' in Hfinal as (w & m1 & Hinterp1 & _). | ||
| 2134 | do 2 right. exists (S (m1 `max` m2)). | ||
| 2135 | rewrite /interp' /force_deep' in Hinterp1. | ||
| 2136 | destruct (interp m1 _ _) as [[]|] eqn:Hinterp1'; simplify_res. | ||
| 2137 | rewrite interp_thunk_S /= (interp_le Hinterp1') /=; last lia. | ||
| 2138 | rewrite (force_deep_le Hinterp1) /=; last lia. | ||
| 2139 | rewrite (mapM_interp_le Hinterp2) /=; last lia. done. | ||
| 2140 | + do 2 right. exists (S m). | ||
| 2141 | rewrite /interp' /force_deep' in Hinterp. | ||
| 2142 | destruct (interp m _ _) as [mw|] eqn:Hinterp1'; simplify_res. | ||
| 2143 | rewrite interp_thunk_S /= Hinterp1' /=. | ||
| 2144 | destruct mw as [w|]; simplify_res; [|done]. | ||
| 2145 | rewrite (force_deep_le Hinterp) /=; last lia. done. | ||
| 2146 | - (* EAttr *) destruct (decide (no_recs αs)) as [Hrecs|]; last first. | ||
| 2147 | { left. by repeat econstructor. } | ||
| 2148 | destruct μ'. | ||
| 2149 | { right; left. by constructor. } | ||
| 2150 | assert (red (step DEEP) (EAttr αs) ∨ | ||
| 2151 | map_Forall (λ _, final DEEP ∘ attr_expr) αs ∨ | ||
| 2152 | ∃ m, map_mapM_sorted attr_le (mbind (force_deep m) ∘ interp_thunk m) | ||
| 2153 | (Thunk ∅ ∘ attr_expr <$> αs) = mfail) as Hhelp; last first. | ||
| 2154 | { destruct Hhelp as [?|[?|[m Hinterp]]]; [by auto using final..|]. | ||
| 2155 | do 2 right. exists (S m). rewrite /interp' interp_S /=. | ||
| 2156 | rewrite from_attr_no_recs //. rewrite force_deep_S /=. by rewrite Hinterp. } | ||
| 2157 | induction αs as [|x [τ e] es Hx ? IH] | ||
| 2158 | using (map_sorted_ind attr_le); [by right; left|]. | ||
| 2159 | rewrite !map_Forall_insert //. | ||
| 2160 | apply map_Forall_insert in H as [He Hes%IH]; clear IH; | ||
| 2161 | [|by eauto using no_recs_insert_inv..]. | ||
| 2162 | assert (τ = NONREC) as -> by (by eapply no_recs_lookup, lookup_insert). | ||
| 2163 | assert (∀ y, is_Some ((Thunk ∅ ∘ attr_expr <$> es) !! y) → attr_le x y). | ||
| 2164 | { intros y. rewrite lookup_fmap fmap_is_Some. eauto. } | ||
| 2165 | destruct (He DEEP) as [[??]|[Hfinal|[m Hinterp]]]; simplify_eq/=. | ||
| 2166 | + left. eexists; eapply SAttr; naive_solver eauto using no_recs_insert_inv. | ||
| 2167 | + destruct Hes as [[??]|[?|[m2 Hinterp2]]]; [|by eauto|]. | ||
| 2168 | * left. inv_step; first by naive_solver eauto using no_recs_insert_inv. | ||
| 2169 | apply lookup_insert_None in Hx as [??]. | ||
| 2170 | rewrite insert_commute // in Hrecs. rewrite insert_commute //. | ||
| 2171 | eexists; eapply SAttr; [|by rewrite lookup_insert_ne| |done]. | ||
| 2172 | { eapply no_recs_insert_inv; [|done]. by rewrite lookup_insert_ne. } | ||
| 2173 | intros ?? [[<- <-]|[??]]%lookup_insert_Some; eauto. | ||
| 2174 | * apply final_interp' in Hfinal as (w & m1 & Hinterp1 & _). | ||
| 2175 | do 2 right. exists (S (m1 `max` m2)). rewrite fmap_insert /=. | ||
| 2176 | rewrite map_mapM_sorted_insert //=; last by rewrite lookup_fmap Hx. | ||
| 2177 | rewrite /interp' /force_deep' in Hinterp1. | ||
| 2178 | destruct (interp m1 _ _) as [[]|] eqn:Hinterp1'; simplify_res. | ||
| 2179 | rewrite interp_thunk_S /= (interp_le Hinterp1') /=; last lia. | ||
| 2180 | rewrite (force_deep_le Hinterp1) /=; last lia. | ||
| 2181 | rewrite (map_mapM_interp_le Hinterp2) /=; last lia. done. | ||
| 2182 | + do 2 right. exists (S m). rewrite fmap_insert /=. | ||
| 2183 | rewrite map_mapM_sorted_insert //=; last by rewrite lookup_fmap Hx. | ||
| 2184 | rewrite /interp' /force_deep' in Hinterp. | ||
| 2185 | destruct (interp m _ _) as [mw|] eqn:Hinterp'; simplify_res. | ||
| 2186 | rewrite interp_thunk_S /= (interp_le Hinterp') /=; last lia. | ||
| 2187 | destruct mw as [w|]; simplify_res; [|done]. | ||
| 2188 | rewrite (force_deep_le Hinterp) /=; last lia. done. | ||
| 2189 | - (* ELetAttr *) destruct (IHe1 SHALLOW) as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 2190 | + left. eexists. by eapply SLetAttr. | ||
| 2191 | + apply final_interp in Hfinal as (w & m & Hinterp & ->). | ||
| 2192 | destruct (maybe VAttr w) eqn:Hw. | ||
| 2193 | { destruct w; simplify_eq/=. left. by repeat econstructor. } | ||
| 2194 | do 2 right. exists (S m). by rewrite /interp' interp_S /= Hinterp /= Hw. | ||
| 2195 | + do 2 right. exists (S m). rewrite interp_shallow' in Hinterp. | ||
| 2196 | by rewrite /interp' interp_S /= Hinterp /=. | ||
| 2197 | - (* EBinOp *) | ||
| 2198 | destruct (IHe1 SHALLOW) as [[??]|[Hfinal1|[m Hinterp]]]. | ||
| 2199 | + left. eexists. by eapply SBinOpL. | ||
| 2200 | + apply final_interp in Hfinal1 as (w1 & m1 & Hinterp1 & ->). | ||
| 2201 | destruct (interp_bin_op op w1) as [f|] eqn:Hop; last first. | ||
| 2202 | { do 2 right. exists (S m1). rewrite /interp' interp_S /=. | ||
| 2203 | by rewrite Hinterp1 /= Hop. } | ||
| 2204 | pose proof Hop as [Φ ?]%interp_bin_op_Some_1. | ||
| 2205 | destruct (IHe2 SHALLOW) as [[??]|[Hfinal2|[m Hinterp2]]]. | ||
| 2206 | * left. by repeat econstructor. | ||
| 2207 | * apply final_interp in Hfinal2 as (w2 & m2 & Hinterp2 & ->). | ||
| 2208 | destruct (f w2) as [w|] eqn:Hf; last first. | ||
| 2209 | ** do 2 right. exists (S (m1 `max` m2)). rewrite /interp' interp_S /=. | ||
| 2210 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 2211 | rewrite Hop /= (interp_le Hinterp2) /=; last lia. by rewrite Hf. | ||
| 2212 | ** eapply interp_bin_op_Some_Some_1 in Hf as (?&?&?); [|done..]. | ||
| 2213 | left. by repeat econstructor. | ||
| 2214 | * rewrite interp_shallow' in Hinterp2. | ||
| 2215 | do 2 right. exists (S (m `max` m1)). rewrite /interp' interp_S /=. | ||
| 2216 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 2217 | rewrite Hop /= (interp_le Hinterp2) /=; last lia. done. | ||
| 2218 | + rewrite interp_shallow' in Hinterp. | ||
| 2219 | do 2 right. exists (S m). by rewrite /interp' interp_S /= Hinterp. | ||
| 2220 | - (* EIf *) | ||
| 2221 | destruct (IHe1 SHALLOW) as [[??]|[Hfinal1|[m Hinterp]]]. | ||
| 2222 | + left. eexists. by eapply SIf. | ||
| 2223 | + apply final_interp in Hfinal1 as (w1 & m1 & Hinterp1 & ->). | ||
| 2224 | destruct (maybe_VLit w1 ≫= maybe LitBool) as [b|] eqn:Hbool; last first. | ||
| 2225 | { do 2 right. exists (S m1). | ||
| 2226 | rewrite /interp' interp_S /= Hinterp1 /= Hbool. done. } | ||
| 2227 | left. destruct w1; repeat destruct select base_lit; simplify_eq/=. | ||
| 2228 | eexists; constructor. | ||
| 2229 | + rewrite interp_shallow' in Hinterp. | ||
| 2230 | do 2 right. exists (S m). by rewrite /interp' interp_S /= Hinterp. | ||
| 2231 | Qed. | ||
| 2232 | |||
| 2233 | Lemma interp_complete μ e1 e2 : | ||
| 2234 | e1 -{μ}->* e2 → nf (step μ) e2 → | ||
| 2235 | ∃ mw m, interp' m μ ∅ e1 = Res mw ∧ | ||
| 2236 | if mw is Some w then e2 = val_to_expr w else ¬final μ e2. | ||
| 2237 | Proof. | ||
| 2238 | intros Hsteps Hnf. induction Hsteps as [e|e1 e2 e3 Hstep _ IH]. | ||
| 2239 | { destruct (red_final_interp μ e) as [?|[Hfinal|[m Hinterp]]]; [done|..]. | ||
| 2240 | - apply final_interp' in Hfinal as (w & m & ? & ?). | ||
| 2241 | by exists (Some w), m. | ||
| 2242 | - exists None, m. split; [done|]. intros Hfinal. | ||
| 2243 | apply final_interp' in Hfinal as (w & m' & Hinterp' & _). | ||
| 2244 | rewrite /interp' in Hinterp, Hinterp'. | ||
| 2245 | by assert (mfail = mret w) by eauto using interp_agree'. } | ||
| 2246 | destruct IH as (mw & m & Hinterp & ?); first done. | ||
| 2247 | eapply interp_step' in Hstep as (mw' & m' & ? & ?); last done. | ||
| 2248 | destruct mw, mw'; naive_solver. | ||
| 2249 | Qed. | ||
| 2250 | |||
| 2251 | Lemma interp_complete_ret μ e1 e2 : | ||
| 2252 | e1 -{μ}->* e2 → final μ e2 → | ||
| 2253 | ∃ w m, interp' m μ ∅ e1 = mret w ∧ e2 = val_to_expr w. | ||
| 2254 | Proof. | ||
| 2255 | intros Hsteps Hfinal. apply interp_complete in Hsteps | ||
| 2256 | as ([w|] & m & ? & ?); naive_solver eauto using final_nf. | ||
| 2257 | Qed. | ||
| 2258 | Lemma interp_complete_fail μ e1 e2 : | ||
| 2259 | e1 -{μ}->* e2 → nf (step μ) e2 → ¬final μ e2 → | ||
| 2260 | ∃ m, interp' m μ ∅ e1 = mfail. | ||
| 2261 | Proof. | ||
| 2262 | intros Hsteps Hnf Hfinal. | ||
| 2263 | apply interp_complete in Hsteps as ([w|] & m & ? & ?); | ||
| 2264 | naive_solver eauto using final_val_to_expr'. | ||
| 2265 | Qed. | ||
| 2266 | |||
| 2267 | Lemma interp_sound_open n E e mv : | ||
| 2268 | interp n E e = Res mv → | ||
| 2269 | ∃ e', subst_env E e -{SHALLOW}->* e' ∧ | ||
| 2270 | if mv is Some v' then e' = val_to_expr v' else stuck SHALLOW e' | ||
| 2271 | with interp_thunk_sound n t mv : | ||
| 2272 | interp_thunk n t = Res mv → | ||
| 2273 | ∃ e', thunk_to_expr t -{SHALLOW}->* e' ∧ | ||
| 2274 | if mv is Some v' then e' = val_to_expr v' else stuck SHALLOW e' | ||
| 2275 | with interp_app_sound n v1 t2 mv : | ||
| 2276 | interp_app n v1 t2 = Res mv → | ||
| 2277 | ∃ e', EApp (val_to_expr v1) (thunk_to_expr t2) -{SHALLOW}->* e' ∧ | ||
| 2278 | if mv is Some v' then e' = val_to_expr v' else stuck SHALLOW e' | ||
| 2279 | with force_deep_sound n v mv : | ||
| 2280 | force_deep n v = Res mv → | ||
| 2281 | ∃ e', val_to_expr v -{DEEP}->* e' ∧ | ||
| 2282 | if mv is Some v' then e' = val_to_expr v' else stuck DEEP e'. | ||
| 2283 | Proof. | ||
| 2284 | - destruct n as [|n]; [done|]. | ||
| 2285 | rewrite subst_env_eq interp_S. intros Hinterp. | ||
| 2286 | destruct e; simplify_res. | ||
| 2287 | + (* ELit *) case_guard; simplify_res. | ||
| 2288 | * by eexists. | ||
| 2289 | * eexists; split; [done|]. split; [|by inv 1]. intros [??]; inv_step. | ||
| 2290 | + (* EId *) | ||
| 2291 | assert (union_kinded (prod_map id thunk_to_expr <$> E !! x) mke | ||
| 2292 | = prod_map id thunk_to_expr <$> (union_kinded (E !! x) | ||
| 2293 | (prod_map id (Thunk ∅) <$> mke))) as ->. | ||
| 2294 | { destruct (_ !! _) as [[[]]|], mke as [[[]]|]; | ||
| 2295 | by rewrite /= ?thunk_to_expr_eq /= ?subst_env_empty. } | ||
| 2296 | destruct (union_kinded _ _) as [[k t]|]; simplify_res. | ||
| 2297 | * apply interp_thunk_sound in Hinterp as (e' & Hsteps & He'). | ||
| 2298 | exists e'; split; [|done]. eapply rtc_l; [constructor|done]. | ||
| 2299 | * eexists; split; [done|]. split; [|inv 1]. intros [? Hstep]. inv_step. | ||
| 2300 | + (* EAbs *) by eexists. | ||
| 2301 | + (* EAbsMatch *) by eexists. | ||
| 2302 | + (* EApp *) | ||
| 2303 | destruct (interp _ _ _) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 2304 | apply interp_sound_open in Hinterp1 as (e1' & Hsteps1 & He1'). | ||
| 2305 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 2306 | { eexists; split; [by eapply SAppL_rtc|]. split; [|inv 1]. | ||
| 2307 | intros [??]. destruct He1' as [Hnf []]. | ||
| 2308 | inv_step; eauto using final. destruct Hnf; eauto. } | ||
| 2309 | apply interp_app_sound in Hinterp as (e' & Hsteps2 & He'). | ||
| 2310 | eexists e'; split; [|done]. etrans; [|done]. by eapply SAppL_rtc. | ||
| 2311 | + (* ESeq *) destruct (interp _ _ e1) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 2312 | apply interp_sound_open in Hinterp' as (e' & Hsteps & He'). | ||
| 2313 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 2314 | { eexists; repeat split; [by apply SSeq_rtc, steps_shallow_any| |inv 1]. | ||
| 2315 | intros [e'' Hstep]. destruct He' as [Hnf Hfinal]. | ||
| 2316 | destruct Hfinal. inv_step; eauto using final_any_shallow. | ||
| 2317 | apply step_any_shallow in H2 as []; [|done]. destruct Hnf; eauto. } | ||
| 2318 | destruct μ; simplify_res. | ||
| 2319 | { apply interp_sound_open in Hinterp as (e'' & Hsteps' & He''). | ||
| 2320 | eexists; split; [|done]. etrans; first by apply SSeq_rtc. | ||
| 2321 | eapply rtc_l; first by apply SSeqFinal. done. } | ||
| 2322 | destruct (force_deep _ _) as [mw|] eqn:Hforce; simplify_res. | ||
| 2323 | pose proof Hforce as Hforce'. | ||
| 2324 | apply force_deep_sound in Hforce' as (e'' & Hsteps' & He''). | ||
| 2325 | destruct mw as [w|]; simplify_res; last first. | ||
| 2326 | { eexists. split. | ||
| 2327 | { etrans; [by eapply SSeq_rtc, steps_shallow_any|]. | ||
| 2328 | etrans; [by eapply SSeq_rtc|]. done. } | ||
| 2329 | split; [|inv 1]. destruct He''. intros [e''' Hstep]. | ||
| 2330 | inv_step; eauto using step_not_final. } | ||
| 2331 | apply interp_sound_open in Hinterp as (e''' & Hsteps'' & He'''). | ||
| 2332 | exists e'''. split; [|done]. | ||
| 2333 | etrans; [by eapply SSeq_rtc, steps_shallow_any|]. | ||
| 2334 | etrans; [by eapply SSeq_rtc|]. | ||
| 2335 | eapply rtc_l; first by eapply SSeqFinal, final_force_deep. done. | ||
| 2336 | + (* EList *) | ||
| 2337 | eexists; split; [done|]. f_equal. | ||
| 2338 | induction es; f_equal/=; auto. | ||
| 2339 | + (* EAttr *) | ||
| 2340 | eexists; split; [apply SAttr_rec_rtc|]. | ||
| 2341 | f_equal. apply map_eq=> x. rewrite !lookup_fmap. | ||
| 2342 | destruct (αs !! x) as [[[] e]|] eqn:?; do 2 f_equal/=. | ||
| 2343 | by rewrite subst_env_indirects_env_attr_to_tattr. | ||
| 2344 | + (* ELetAttr *) destruct (interp _ _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 2345 | apply interp_sound_open in Hinterp' as (e' & Hsteps & He'). | ||
| 2346 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 2347 | { eexists; repeat split; [by apply SLetAttr_rtc| |inv 1]. | ||
| 2348 | intros [e'' Hstep]. destruct He' as [Hnf Hfinal]. | ||
| 2349 | inv_step; [by destruct Hfinal; constructor|]. destruct Hnf; eauto. } | ||
| 2350 | destruct (maybe VAttr v') eqn:?; simplify_res; last first. | ||
| 2351 | { eexists; repeat split; [by apply SLetAttr_rtc| |inv 1]. | ||
| 2352 | intros [e'' Hstep]. destruct v'; inv_step; simplify_eq/=. } | ||
| 2353 | destruct v'; simplify_res. | ||
| 2354 | apply interp_sound_open in Hinterp as (e'' & Hsteps' & He''). | ||
| 2355 | eexists; split; [|done]. etrans; [by apply SLetAttr_rtc|]. | ||
| 2356 | eapply rtc_l; [by econstructor|]. | ||
| 2357 | rewrite subst_env_union in Hsteps'. | ||
| 2358 | rewrite subst_env_alt -!map_fmap_compose in Hsteps'. | ||
| 2359 | by rewrite -map_fmap_compose. | ||
| 2360 | + (* EBinOp *) | ||
| 2361 | destruct (interp _ _ e1) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 2362 | apply interp_sound_open in Hinterp1 as (e1' & Hsteps1 & He1'). | ||
| 2363 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 2364 | { eexists; split; [by eapply SBinOpL_rtc|]. split; [|inv 1]. | ||
| 2365 | intros [? Hstep]. destruct He1'. inv_step; naive_solver. } | ||
| 2366 | destruct (interp_bin_op _ v1) as [f|] eqn:Hop; simplify_res; last first. | ||
| 2367 | { assert (¬∃ Φ, sem_bin_op op (val_to_expr v1) Φ). | ||
| 2368 | { by intros [? ?%interp_bin_op_Some_2%not_eq_None_Some]. } | ||
| 2369 | eexists; split; [by eapply SBinOpL_rtc|]. split; [|inv 1]. | ||
| 2370 | intros [? Hstep]. inv_step; eauto using step_not_val_to_expr. } | ||
| 2371 | pose proof Hop as [Φ ?]%interp_bin_op_Some_1. | ||
| 2372 | destruct (interp _ _ e2) as [mv2|] eqn:Hinterp2; simplify_res. | ||
| 2373 | apply interp_sound_open in Hinterp2 as (e2' & Hsteps2 & He2'). | ||
| 2374 | destruct mv2 as [v2|]; simplify_res; last first. | ||
| 2375 | { eexists; split. | ||
| 2376 | { etrans; [by eapply SBinOpL_rtc|]. | ||
| 2377 | eapply SBinOpR_rtc; eauto using interp_bin_op_Some_1. } | ||
| 2378 | split; [|inv 1]. destruct He2'. | ||
| 2379 | intros [? Hstep]. inv_step; eauto using step_not_val_to_expr. } | ||
| 2380 | destruct (f v2) eqn:Hf; simplify_res; last first. | ||
| 2381 | { eexists; split. | ||
| 2382 | { etrans; [by eapply SBinOpL_rtc|]. | ||
| 2383 | eapply SBinOpR_rtc; eauto using interp_bin_op_Some_1. } | ||
| 2384 | split; [|inv 1]. pose proof @interp_bin_op_Some_Some_2. | ||
| 2385 | intros [? Hstep]. inv_step; naive_solver eauto using step_not_val_to_expr. } | ||
| 2386 | apply interp_thunk_sound in Hinterp as (e' & Hsteps3 & He'). | ||
| 2387 | eapply interp_bin_op_Some_Some_1 in Hf as (e3 & ? & ?); [|done..]. | ||
| 2388 | eapply delayed_steps_l in Hsteps3 | ||
| 2389 | as (e'' & Hsteps3 & Hdel); last done. | ||
| 2390 | eexists e''; split. | ||
| 2391 | { etrans; [by eapply SBinOpL_rtc|]. | ||
| 2392 | etrans; [eapply SBinOpR_rtc; eauto using interp_bin_op_Some_1|]. | ||
| 2393 | eapply rtc_l; [by econstructor|]. done. } | ||
| 2394 | destruct mv. | ||
| 2395 | { subst e'. eapply delayed_final_l in Hdel as <-; done. } | ||
| 2396 | destruct He' as [Hnf Hfinal]. split. | ||
| 2397 | { intros [e4 Hsteps4]. destruct Hnf. | ||
| 2398 | eapply delayed_step_r in Hsteps4 as (e4' & Hstep4' & ?); [|done]. | ||
| 2399 | destruct Hstep4'; eauto. } | ||
| 2400 | intros Hfinal'. eapply Hnf. | ||
| 2401 | eapply delayed_final_r in Hfinal' as Hsteps; [|done]. | ||
| 2402 | destruct Hsteps; by eauto. | ||
| 2403 | + (* EIf *) | ||
| 2404 | destruct (interp _ _ e1) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 2405 | apply interp_sound_open in Hinterp1 as (e1' & Hsteps1 & He1'). | ||
| 2406 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 2407 | { eexists; repeat split; [by apply SIf_rtc| |inv 1]. | ||
| 2408 | intros [e'' Hstep]. destruct He1' as [Hnf Hfinal]. | ||
| 2409 | destruct Hfinal. inv_step; eauto using final. destruct Hnf; eauto. } | ||
| 2410 | destruct (maybe_VLit v1 ≫= maybe LitBool) as [b|] eqn:Hbool; | ||
| 2411 | simplify_res; last first. | ||
| 2412 | { eexists; repeat split; [by apply SIf_rtc| |inv 1]. | ||
| 2413 | intros [e'' ?]. destruct v1; inv_step; eauto using final. } | ||
| 2414 | apply interp_sound_open in Hinterp as (e' & Hsteps & He'). | ||
| 2415 | exists e'; split; [|done]. etrans; [by apply SIf_rtc|]. | ||
| 2416 | assert (val_to_expr v1 = ELit (LitBool b)) as ->. | ||
| 2417 | { destruct v1; repeat destruct select base_lit; naive_solver. } | ||
| 2418 | eapply rtc_l; [constructor|]. by destruct b. | ||
| 2419 | - destruct n as [|n]; [done|]. rewrite interp_thunk_S /=. | ||
| 2420 | intros Hthunk. destruct t; simplify_res; [by eauto using rtc..|]. | ||
| 2421 | destruct (tαs !! x) as [[e|t]|] eqn:Hx; simplify_res. | ||
| 2422 | + apply interp_sound_open in Hthunk as (e' & Hsteps & ?). | ||
| 2423 | exists e'; split; [|done]. etrans; [eapply SBinOpL_rtc, SAttr_rec_rtc|]. | ||
| 2424 | eapply rtc_l; [eapply SBinOp; repeat constructor|]; try done; simpl. | ||
| 2425 | eexists; split; [done|]. rewrite !lookup_fmap Hx /=. | ||
| 2426 | rewrite -subst_env_indirects_env_attr_to_tattr_empty. | ||
| 2427 | by rewrite -subst_env_indirects_env. | ||
| 2428 | + apply interp_thunk_sound in Hthunk as (e' & Hsteps & ?). | ||
| 2429 | exists e'; split; [|done]. etrans; [eapply SBinOpL_rtc, SAttr_rec_rtc|]. | ||
| 2430 | eapply rtc_l; [eapply SBinOp; repeat constructor|]; try done; simpl. | ||
| 2431 | eexists; split; [done|]. by rewrite !lookup_fmap Hx /=. | ||
| 2432 | + eexists. split; [eapply SBinOpL_rtc, SAttr_rec_rtc|]. split; [|inv 1]. | ||
| 2433 | intros [??]. inv_step. inv H7. destruct H8 as (? & ? & Hx'); simplify_eq/=. | ||
| 2434 | by rewrite !lookup_fmap Hx in Hx'. | ||
| 2435 | - destruct n as [|n]; [done|]. rewrite interp_app_S /=. intros Happ. | ||
| 2436 | destruct v1; simplify_res. | ||
| 2437 | + eexists; split; [done|]. split; [|inv 1]. intros [??]; inv_step. | ||
| 2438 | + eapply interp_sound_open in Happ as (e' & Hsteps & He'). | ||
| 2439 | eexists; split; [|done]. eapply rtc_l; [constructor|]. | ||
| 2440 | rewrite subst_abs_env_insert // in Hsteps. | ||
| 2441 | + destruct (interp_thunk _ _) as [mv'|] eqn:Hthunk; simplify_res. | ||
| 2442 | apply interp_thunk_sound in Hthunk as (et & Htsteps & Het). | ||
| 2443 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 2444 | { eexists; split; [by eapply SAppR_rtc|]. | ||
| 2445 | split; [|inv 1]. destruct Het. | ||
| 2446 | intros [??]; inv_step; eauto using final. } | ||
| 2447 | destruct (maybe VAttr v') as [ts|] eqn:?; simplify_res; last first. | ||
| 2448 | { eexists; repeat split; [by apply SAppR_rtc| |inv 1]. | ||
| 2449 | intros [e'' Hstep]. destruct v'; inv_step; simplify_eq/=. } | ||
| 2450 | destruct v'; simplify_res. | ||
| 2451 | destruct (interp_match _ _ _) as [tαs|] eqn:Hmatch; | ||
| 2452 | simplify_res; last first. | ||
| 2453 | { eexists; repeat split; [by apply SAppR_rtc| |inv 1]. | ||
| 2454 | intros [e'' Hstep]. inv_step. | ||
| 2455 | rewrite map_fmap_compose fmap_attr_expr_Attr in H6. | ||
| 2456 | apply interp_match_Some_2 in H6. rewrite interp_match_subst in H6. | ||
| 2457 | opose proof (interp_match_proper ∅ ∅ | ||
| 2458 | (Thunk ∅ <$> (thunk_to_expr <$> ts)) ts ms ms strict _ _). | ||
| 2459 | { apply map_eq=> x. rewrite !lookup_fmap. | ||
| 2460 | destruct (ts !! x); f_equal/=. by rewrite subst_env_empty. } | ||
| 2461 | { done. } | ||
| 2462 | repeat destruct (interp_match _ _ _); simplify_eq/=. } | ||
| 2463 | pose proof (interp_match_subst E ts ms strict) as Hmatch'. | ||
| 2464 | rewrite Hmatch /= in Hmatch'. | ||
| 2465 | apply interp_match_Some_1 in Hmatch'. | ||
| 2466 | apply interp_sound_open in Happ as (e' & Hsteps & ?). | ||
| 2467 | exists e'; split; [|done]. | ||
| 2468 | etrans; [by apply SAppR_rtc|]. | ||
| 2469 | eapply rtc_l; [constructor; [done|]|]. | ||
| 2470 | { rewrite map_fmap_compose fmap_attr_expr_Attr. done. } | ||
| 2471 | etrans; [|apply Hsteps]. apply reflexive_eq. f_equal. | ||
| 2472 | rewrite subst_env_indirects_env. | ||
| 2473 | rewrite subst_env_indirects_env_attr_to_tattr_empty. | ||
| 2474 | do 2 f_equal. apply map_eq=> y. rewrite !lookup_fmap. | ||
| 2475 | destruct (_ !! y) as [[]|]; f_equal/=. by rewrite subst_env_empty. | ||
| 2476 | + eexists; split; [done|]. split; [|inv 1]. intros [??]; inv_step. | ||
| 2477 | + destruct (ts !! _) eqn:Hfunc; simplify_res; last first. | ||
| 2478 | { eexists; split; [by eapply SAppL_rtc|]. split; [|inv 1]. | ||
| 2479 | intros [??]; inv_step; simplify_map_eq. } | ||
| 2480 | destruct (interp_thunk _ _) as [mv'|] eqn:Hthunk; simplify_res. | ||
| 2481 | apply interp_thunk_sound in Hthunk as (et & Htsteps & Het). | ||
| 2482 | assert (EApp (EAttr (AttrN ∘ thunk_to_expr <$> ts)) (thunk_to_expr t2) | ||
| 2483 | -{SHALLOW}->* | ||
| 2484 | EApp (EApp et (EAttr (AttrN ∘ thunk_to_expr <$> ts))) (thunk_to_expr t2)) | ||
| 2485 | as Hsteps; [|clear Htsteps]. | ||
| 2486 | { eapply rtc_l; [constructor; by simplify_map_eq|]. | ||
| 2487 | eapply SAppL_rtc, SAppL_rtc, Htsteps. } | ||
| 2488 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 2489 | { eexists; split; [exact Hsteps|]. | ||
| 2490 | split; [|inv 1]. intros [??]. destruct Het as [Hnf []]. | ||
| 2491 | inv_step; eauto using final. destruct Hnf; eauto. } | ||
| 2492 | destruct (interp_app _ _ _) as [mv'|] eqn:Happ'; simplify_res. | ||
| 2493 | apply interp_app_sound in Happ' as (e' & Hsteps' & He'). | ||
| 2494 | destruct mv' as [v''|]; simplify_res; last first. | ||
| 2495 | { eexists; split; [etrans; [apply Hsteps|apply SAppL_rtc, Hsteps']|]. | ||
| 2496 | split; [|inv 1]. intros [??]. destruct He' as [Hnf []]. | ||
| 2497 | inv_step; eauto using final. destruct Hnf; eauto. } | ||
| 2498 | apply interp_app_sound in Happ as (e'' & Hsteps'' & He''). | ||
| 2499 | eexists e''; split; [|done]. | ||
| 2500 | etrans; [apply Hsteps|]. etrans; [apply SAppL_rtc, Hsteps'|]. done. | ||
| 2501 | - destruct n as [|n]; [done|]. rewrite force_deep_S. | ||
| 2502 | intros Hforce. destruct v; simplify_res. | ||
| 2503 | + (* VLit *) by eexists. | ||
| 2504 | + (* VAbs *) by eexists. | ||
| 2505 | + (* VAbsMatch *) by eexists. | ||
| 2506 | + (* VList *) | ||
| 2507 | destruct (mapM _ _) as [mvs|] eqn:Hmap; simplify_res. | ||
| 2508 | assert (∃ ts', | ||
| 2509 | EList (thunk_to_expr <$> ts) -{DEEP}->* EList (thunk_to_expr <$> ts') ∧ | ||
| 2510 | if mvs is Some vs then thunk_to_expr <$> ts' = val_to_expr <$> vs | ||
| 2511 | else nf (step DEEP) (EList (thunk_to_expr <$> ts')) ∧ | ||
| 2512 | ¬Forall (final DEEP ∘ thunk_to_expr) ts') | ||
| 2513 | as (ts' & Hsteps & Hts'); last first. | ||
| 2514 | { eexists; split; [done|]. destruct mvs as [vs|]; simplify_eq/=. | ||
| 2515 | * f_equal. rewrite -list_fmap_compose Hts'. | ||
| 2516 | clear. induction vs; f_equal/=; auto. | ||
| 2517 | * destruct Hts' as [Hnf Hfinal]; split; [done|]. | ||
| 2518 | inv 1. by apply Hfinal, Forall_fmap. } | ||
| 2519 | revert mvs Hmap. induction ts as [|t ts IH]; intros mv' Hmap; simplify_res. | ||
| 2520 | { by exists []. } | ||
| 2521 | destruct (interp_thunk _ _) as [mv''|] eqn:Hthunk; simplify_res. | ||
| 2522 | apply interp_thunk_sound in Hthunk as (et & Htsteps & Het). | ||
| 2523 | destruct mv'' as [v''|]; simplify_res; last first. | ||
| 2524 | { exists (Thunk ∅ et :: ts); csimpl. rewrite subst_env_empty. | ||
| 2525 | apply (stuck_shallow_any DEEP) in Het as [??]. split_and!. | ||
| 2526 | * eapply (SList_rtc []); [done|]. | ||
| 2527 | etrans; [by apply steps_shallow_any|done]. | ||
| 2528 | * by apply List_nf_cons. | ||
| 2529 | * rewrite Forall_cons /= subst_env_empty. | ||
| 2530 | naive_solver eauto using final_any_shallow. } | ||
| 2531 | destruct (force_deep _ _) as [mvf|] eqn:Hforce; simplify_res. | ||
| 2532 | pose proof Hforce as Hforce'. | ||
| 2533 | apply force_deep_sound in Hforce' as (e' & Hsteps' & He'). | ||
| 2534 | destruct mvf as [vf|]; simplify_res; last first. | ||
| 2535 | { exists (Thunk ∅ e' :: ts). csimpl. rewrite subst_env_empty. | ||
| 2536 | destruct He'. split_and!. | ||
| 2537 | * eapply (SList_rtc []); [done|]. | ||
| 2538 | etrans; [by apply steps_shallow_any|done]. | ||
| 2539 | * by apply List_nf_cons. | ||
| 2540 | * rewrite Forall_cons /= subst_env_empty. naive_solver. } | ||
| 2541 | destruct (mapM _ _) as [mvs|] eqn:Hmap'; simplify_res. | ||
| 2542 | destruct (IH _ eq_refl) as (ts' & Hsteps'' & Hts'). | ||
| 2543 | exists (Forced vf :: ts'); csimpl. split. | ||
| 2544 | { etrans; [eapply (SList_rtc []); [done..|]; | ||
| 2545 | etrans; [by apply steps_shallow_any|done]|]; simpl. | ||
| 2546 | eapply List_steps_cons; by eauto using final_force_deep. } | ||
| 2547 | destruct mvs as [vs|]; simplify_res. | ||
| 2548 | { by rewrite Hts'. } | ||
| 2549 | split; [|rewrite Forall_cons; naive_solver]. | ||
| 2550 | apply List_nf_cons_final; naive_solver eauto using final_force_deep. | ||
| 2551 | + (* VAttr *) | ||
| 2552 | destruct (map_mapM_sorted _ _) as [mvs|] eqn:Hmap; simplify_res. | ||
| 2553 | assert (∃ ts', | ||
| 2554 | EAttr (AttrN ∘ thunk_to_expr <$> ts) -{DEEP}->* | ||
| 2555 | EAttr (AttrN ∘ thunk_to_expr <$> ts') ∧ | ||
| 2556 | if mvs is Some vs then thunk_to_expr <$> ts' = val_to_expr <$> vs | ||
| 2557 | else nf (step DEEP) (EAttr (AttrN ∘ thunk_to_expr <$> ts')) ∧ | ||
| 2558 | ¬map_Forall (λ _, final DEEP ∘ thunk_to_expr) ts') | ||
| 2559 | as (ts' & Hsteps & Hts'); last first. | ||
| 2560 | { eexists; split; [done|]. destruct mvs as [vs|]; simplify_eq/=. | ||
| 2561 | * f_equal. rewrite map_fmap_compose Hts'. | ||
| 2562 | apply map_eq=> x. rewrite !lookup_fmap. by destruct (vs !! x). | ||
| 2563 | * destruct Hts' as [Hnf Hfinal]; split; [done|]. | ||
| 2564 | inv 1. apply Hfinal=> x t Hx /=. | ||
| 2565 | ospecialize (H2 x _ _); first by rewrite lookup_fmap Hx. done. } | ||
| 2566 | revert mvs Hmap. induction ts as [|x t ts Hx ? IH] | ||
| 2567 | using (map_sorted_ind attr_le); intros mv' Hmap. | ||
| 2568 | { rewrite map_mapM_sorted_empty in Hmap; simplify_res. by exists ∅. } | ||
| 2569 | rewrite map_mapM_sorted_insert //= in Hmap. | ||
| 2570 | assert ((AttrN ∘ thunk_to_expr <$> ts) !! x = None). | ||
| 2571 | { by rewrite lookup_fmap Hx. } | ||
| 2572 | assert (∀ y α, (AttrN ∘ thunk_to_expr <$> ts) !! y = Some α → | ||
| 2573 | final DEEP (attr_expr α) ∨ attr_le x y). | ||
| 2574 | { intros y α. rewrite lookup_fmap. destruct (ts !! y) eqn:?; naive_solver. } | ||
| 2575 | destruct (interp_thunk _ _) as [mv''|] eqn:Hthunk; simplify_res. | ||
| 2576 | apply interp_thunk_sound in Hthunk as (et & Htsteps & Het). | ||
| 2577 | destruct mv'' as [v''|]; simplify_res; last first. | ||
| 2578 | { exists (<[x:=Thunk ∅ et]> ts). | ||
| 2579 | rewrite !fmap_insert /= subst_env_empty. | ||
| 2580 | apply (stuck_shallow_any DEEP) in Het as [??]. split_and!. | ||
| 2581 | * eapply SAttr_lookup_rtc; [done..|]. | ||
| 2582 | etrans; [by apply steps_shallow_any|done]. | ||
| 2583 | * apply Attr_nf_insert; auto. | ||
| 2584 | intros y. rewrite lookup_fmap fmap_is_Some. eauto. | ||
| 2585 | * rewrite map_Forall_insert //= subst_env_empty. | ||
| 2586 | naive_solver eauto using final_any_shallow. } | ||
| 2587 | destruct (force_deep _ _) as [mvf|] eqn:Hforce; simplify_res. | ||
| 2588 | pose proof Hforce as Hforce'. | ||
| 2589 | apply force_deep_sound in Hforce' as (e' & Hsteps' & He'). | ||
| 2590 | destruct mvf as [vf|]; simplify_res; last first. | ||
| 2591 | { exists (<[x:=Thunk ∅ e']> ts). rewrite !fmap_insert /= subst_env_empty. | ||
| 2592 | destruct He'. split_and!. | ||
| 2593 | * eapply SAttr_lookup_rtc; [done..|]. | ||
| 2594 | etrans; [by apply steps_shallow_any|done]. | ||
| 2595 | * apply Attr_nf_insert; auto. | ||
| 2596 | intros y. rewrite lookup_fmap fmap_is_Some. eauto. | ||
| 2597 | * rewrite map_Forall_insert //= subst_env_empty. naive_solver. } | ||
| 2598 | destruct (map_mapM_sorted _ _ _) as [mvs|] eqn:Hmap'; simplify_res. | ||
| 2599 | destruct (IH _ eq_refl) as (ts' & Hsteps'' & Hts'). | ||
| 2600 | exists (<[x:=Forced vf]> ts'). split. | ||
| 2601 | { rewrite !fmap_insert /=. | ||
| 2602 | etrans; [eapply SAttr_lookup_rtc; [done..|]; | ||
| 2603 | etrans; [by apply steps_shallow_any|done]|]. | ||
| 2604 | eapply Attr_steps_insert; by eauto using final_force_deep. } | ||
| 2605 | destruct mvs as [vs|]; simplify_res. | ||
| 2606 | { by rewrite !fmap_insert Hts'. } | ||
| 2607 | assert (∀ y, ts !! y = None ↔ ts' !! y = None) as Hdom. | ||
| 2608 | { intros y. rewrite -!(fmap_None (AttrN ∘ thunk_to_expr)). | ||
| 2609 | rewrite -!lookup_fmap. by eapply Attr_steps_dom. } | ||
| 2610 | split; [|rewrite map_Forall_insert; naive_solver]. | ||
| 2611 | rewrite fmap_insert /=. apply Attr_nf_insert_final; | ||
| 2612 | eauto using final_force_deep. | ||
| 2613 | * rewrite lookup_fmap fmap_None. naive_solver. | ||
| 2614 | * intros y. rewrite lookup_fmap fmap_is_Some. | ||
| 2615 | rewrite -not_eq_None_Some -Hdom not_eq_None_Some. auto. | ||
| 2616 | * naive_solver. | ||
| 2617 | Qed. | ||
| 2618 | |||
| 2619 | Lemma interp_sound_open' n μ E e mv : | ||
| 2620 | interp' n μ E e = Res mv → | ||
| 2621 | ∃ e', subst_env E e -{μ}->* e' ∧ | ||
| 2622 | if mv is Some v' then e' = val_to_expr v' else stuck μ e'. | ||
| 2623 | Proof. | ||
| 2624 | intros Hinterp. destruct μ. | ||
| 2625 | { rewrite interp_shallow' in Hinterp. by eapply interp_sound_open. } | ||
| 2626 | rewrite /interp' /= in Hinterp. | ||
| 2627 | destruct (interp n E e) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 2628 | apply interp_sound_open in Hinterp' as (e' & Hsteps & He'). | ||
| 2629 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 2630 | { eauto using steps_shallow_any, stuck_shallow_any. } | ||
| 2631 | eapply force_deep_sound in Hinterp as (e'' & Hsteps' & He''). | ||
| 2632 | eexists; split; [|done]. etrans; [by eapply steps_shallow_any|done]. | ||
| 2633 | Qed. | ||
| 2634 | |||
| 2635 | Lemma interp_sound n μ e mv : | ||
| 2636 | interp' n μ ∅ e = Res mv → | ||
| 2637 | ∃ e', e -{μ}->* e' ∧ | ||
| 2638 | if mv is Some v then e' = val_to_expr v else stuck μ e'. | ||
| 2639 | Proof. | ||
| 2640 | intros Hsteps%interp_sound_open'. by rewrite subst_env_empty in Hsteps. | ||
| 2641 | Qed. | ||
| 2642 | |||
| 2643 | (** Final theorems *) | ||
| 2644 | Theorem interp_sound_complete_ret e v : | ||
| 2645 | (∃ w n, interp' n SHALLOW ∅ e = mret w ∧ val_to_expr v = val_to_expr w) | ||
| 2646 | ↔ e -{SHALLOW}->* val_to_expr v. | ||
| 2647 | Proof. | ||
| 2648 | split. | ||
| 2649 | - by intros (n & w & (e' & ? & ->)%interp_sound & ->). | ||
| 2650 | - intros Hsteps. apply interp_complete in Hsteps as ([] & ? & ? & ?); | ||
| 2651 | unfold nf, red; | ||
| 2652 | naive_solver eauto using final_val_to_expr, step_not_val_to_expr. | ||
| 2653 | Qed. | ||
| 2654 | |||
| 2655 | Theorem interp_sound_complete_ret_lit μ e bl (Hbl : base_lit_ok bl) : | ||
| 2656 | (∃ n, interp' n μ ∅ e = mret (VLit bl Hbl)) ↔ e -{μ}->* ELit bl. | ||
| 2657 | Proof. | ||
| 2658 | split. | ||
| 2659 | - intros [n (e' & ? & ->)%interp_sound]. done. | ||
| 2660 | - intros Hsteps. apply interp_complete_ret in Hsteps | ||
| 2661 | as ([] & n & ? & Hv); simplify_eq/=; last by constructor. | ||
| 2662 | exists n. by rewrite (proof_irrel Hbl Hbl0). | ||
| 2663 | Qed. | ||
| 2664 | |||
| 2665 | Theorem interp_sound_complete_fail μ e : | ||
| 2666 | (∃ n, interp' n μ ∅ e = mfail) ↔ ∃ e', e -{μ}->* e' ∧ stuck μ e'. | ||
| 2667 | Proof. | ||
| 2668 | split. | ||
| 2669 | - by intros [n ?%interp_sound]. | ||
| 2670 | - intros (e' & Hsteps & Hnf & Hfinal). by eapply interp_complete_fail. | ||
| 2671 | Qed. | ||
| 2672 | |||
| 2673 | Theorem interp_sound_complete_no_fuel μ e : | ||
| 2674 | (∀ n, interp' n μ ∅ e = NoFuel) ↔ all_loop (step μ) e. | ||
| 2675 | Proof. | ||
| 2676 | rewrite all_loop_alt. split. | ||
| 2677 | - intros Hnofuel e' Hsteps. | ||
| 2678 | destruct (red_final_interp μ e') as [|[|He']]; [done|..]. | ||
| 2679 | + apply interp_complete_ret in Hsteps as (w & m & Hinterp & _); last done. | ||
| 2680 | by rewrite Hnofuel in Hinterp. | ||
| 2681 | + apply interp_sound_complete_fail in He' as (e'' & ? & [Hnf _]). | ||
| 2682 | destruct (interp_complete μ e e'') | ||
| 2683 | as (mv & n & Hinterp & _); [by etrans|done|]. | ||
| 2684 | by rewrite Hnofuel in Hinterp. | ||
| 2685 | - intros Hred n. destruct (interp' n μ ∅ e) as [mv|] eqn:Hinterp; [|done]. | ||
| 2686 | destruct (interp_sound _ _ _ _ Hinterp) as (e' & Hsteps & Hstuck). | ||
| 2687 | destruct mv as [v|]; simplify_eq/=. | ||
| 2688 | + apply Hred in Hsteps as []%final_nf. by eapply final_val_to_expr'. | ||
| 2689 | + destruct Hstuck as [[] ?]; eauto. | ||
| 2690 | Qed. | ||
diff --git a/theories/nix/notations.v b/theories/nix/notations.v new file mode 100644 index 0000000..e9995b5 --- /dev/null +++ b/theories/nix/notations.v | |||
| @@ -0,0 +1,43 @@ | |||
| 1 | From mininix Require Export nix.operational. | ||
| 2 | |||
| 3 | (* Influenced by | ||
| 4 | https://gitlab.mpi-sws.org/iris/iris/-/blob/master/iris_heap_lang/notation.v | ||
| 5 | But always uses ":" instead of a scope. *) | ||
| 6 | |||
| 7 | Coercion EId' : string >-> expr. | ||
| 8 | Coercion NInt : Z >-> num. | ||
| 9 | Coercion NFloat : float >-> num. | ||
| 10 | Coercion LitNum : num >-> base_lit. | ||
| 11 | Coercion LitBool : bool >-> base_lit. | ||
| 12 | Coercion ELit : base_lit >-> expr. | ||
| 13 | Coercion EApp : expr >-> Funclass. | ||
| 14 | |||
| 15 | Notation "λattr: a , e" := (EAbsMatch a true e) | ||
| 16 | (at level 200, e, a at level 200, | ||
| 17 | format "'[' 'λattr:' a , '/ ' e ']'"). | ||
| 18 | Notation "λattr: a .., e" := (EAbsMatch a false e) | ||
| 19 | (at level 200, e, a at level 200, | ||
| 20 | format "'[' 'λattr:' a .., '/ ' e ']'"). | ||
| 21 | |||
| 22 | Notation "λ: x .. y , e" := (EAbs x .. (EAbs y e) ..) | ||
| 23 | (at level 200, x, y at level 1, e at level 200, | ||
| 24 | format "'[' 'λ:' x .. y , '/ ' e ']'"). | ||
| 25 | Notation "'let:' x := e1 'in' e2" := (ELet x e1 e2) | ||
| 26 | (at level 200, x at level 1, e1, e2 at level 200, | ||
| 27 | format "'[' 'let:' x := '[' e1 ']' 'in' '/' e2 ']'"). | ||
| 28 | Notation "'with:' a 'in' e" := (EWith a e) | ||
| 29 | (at level 200, a, e at level 200, | ||
| 30 | format "'[' 'with:' a 'in' '/' e ']'"). | ||
| 31 | |||
| 32 | Notation "'if:' e1 'then' e2 'else' e3" := (EIf e1 e2 e3) | ||
| 33 | (at level 200, e1, e2, e3 at level 200). | ||
| 34 | |||
| 35 | Notation "e1 .: e2" := (ESelect e1 e2) (at level 70, no associativity). | ||
| 36 | |||
| 37 | Notation "e1 +: e2" := (EBinOp AddOp e1 e2) (at level 50, left associativity). | ||
| 38 | Notation "e1 *: e2" := (EBinOp MulOp e1 e2). | ||
| 39 | Notation "e1 -: e2" := (EBinOp SubOp e1 e2) (at level 50, left associativity). | ||
| 40 | Notation "e1 /: e2" := (EBinOp DivOp e1 e2) (at level 40). | ||
| 41 | Notation "e1 =: e2" := (EBinOp EqOp e1 e2) (at level 70, no associativity). | ||
| 42 | Notation "e1 <: e2" := (EBinOp LtOp e1 e2) (at level 70, no associativity). | ||
| 43 | Notation "'ceil:' e" := (EBinOp (RoundOp Ceil) e LitNull) (at level 10). | ||
diff --git a/theories/nix/operational.v b/theories/nix/operational.v new file mode 100644 index 0000000..d3f0777 --- /dev/null +++ b/theories/nix/operational.v | |||
| @@ -0,0 +1,527 @@ | |||
| 1 | From mininix Require Export utils nix.floats. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | (** Our development does not rely on a particular order on attribute set names. | ||
| 5 | It can be any decidable total order. We pick something concrete (lexicographic | ||
| 6 | order on strings) to avoid having to parametrize the whole development. *) | ||
| 7 | Definition attr_le := String.le. | ||
| 8 | Global Instance attr_le_dec : RelDecision attr_le := _. | ||
| 9 | Global Instance attr_le_po : PartialOrder attr_le := _. | ||
| 10 | Global Instance attr_le_total : Total attr_le := _. | ||
| 11 | Global Typeclasses Opaque attr_le. | ||
| 12 | |||
| 13 | Inductive mode := SHALLOW | DEEP. | ||
| 14 | Inductive kind := ABS | WITH. | ||
| 15 | Inductive rec := REC | NONREC. | ||
| 16 | |||
| 17 | Global Instance rec_eq_dec : EqDecision rec. | ||
| 18 | Proof. solve_decision. Defined. | ||
| 19 | |||
| 20 | Inductive num := | ||
| 21 | | NInt (n : Z) | ||
| 22 | | NFloat (f : float). | ||
| 23 | |||
| 24 | Inductive base_lit := | ||
| 25 | | LitNum (n : num) | ||
| 26 | | LitBool (b : bool) | ||
| 27 | | LitString (s : string) | ||
| 28 | | LitNull. | ||
| 29 | |||
| 30 | Global Instance num_inhabited : Inhabited num := populate (NInt 0). | ||
| 31 | Global Instance base_lit_inhabited : Inhabited base_lit := populate LitNull. | ||
| 32 | |||
| 33 | Global Instance num_eq_dec : EqDecision num. | ||
| 34 | Proof. solve_decision. Defined. | ||
| 35 | Global Instance base_lit_eq_dec : EqDecision base_lit. | ||
| 36 | Proof. solve_decision. Defined. | ||
| 37 | |||
| 38 | Global Instance maybe_NInt : Maybe NInt := λ n, | ||
| 39 | if n is NInt i then Some i else None. | ||
| 40 | Global Instance maybe_NFloat : Maybe NFloat := λ n, | ||
| 41 | if n is NFloat f then Some f else None. | ||
| 42 | Global Instance maybe_LitNum : Maybe LitNum := λ bl, | ||
| 43 | if bl is LitNum n then Some n else None. | ||
| 44 | Global Instance maybe_LitBool : Maybe LitBool := λ bl, | ||
| 45 | if bl is LitBool b then Some b else None. | ||
| 46 | Global Instance maybe_LitString : Maybe LitString := λ bl, | ||
| 47 | if bl is LitString s then Some s else None. | ||
| 48 | |||
| 49 | Inductive bin_op : Set := | ||
| 50 | | AddOp | SubOp | MulOp | DivOp | AndOp | OrOp | XOrOp (* Arithmetic *) | ||
| 51 | | LtOp | EqOp (* Relations *) | ||
| 52 | | RoundOp (m : round_mode) (* Conversions *) | ||
| 53 | | MatchStringOp (* Strings *) | ||
| 54 | | MatchListOp | AppendListOp (* Lists *) | ||
| 55 | | SelectAttrOp | UpdateAttrOp | HasAttrOp | ||
| 56 | | DeleteAttrOp | SingletonAttrOp | MatchAttrOp (* Attribute sets *) | ||
| 57 | | FunctionArgsOp | TypeOfOp. | ||
| 58 | |||
| 59 | Global Instance bin_op_eq_dec : EqDecision bin_op. | ||
| 60 | Proof. solve_decision. Defined. | ||
| 61 | |||
| 62 | Global Instance maybe_RoundOp : Maybe RoundOp := λ op, | ||
| 63 | if op is RoundOp m then Some m else None. | ||
| 64 | |||
| 65 | Section expr. | ||
| 66 | Local Unset Elimination Schemes. | ||
| 67 | Inductive expr := | ||
| 68 | | ELit (bl : base_lit) | ||
| 69 | | EId (x : string) (mke : option (kind * expr)) | ||
| 70 | | EAbs (x : string) (e : expr) | ||
| 71 | | EAbsMatch (ms : gmap string (option expr)) (strict : bool) (e : expr) | ||
| 72 | | EApp (e1 e2 : expr) | ||
| 73 | | ESeq (μ : mode) (e1 e2 : expr) | ||
| 74 | | EList (es : list expr) | ||
| 75 | | EAttr (αs : gmap string attr) | ||
| 76 | | ELetAttr (k : kind) (e1 e2 : expr) | ||
| 77 | | EBinOp (op : bin_op) (e1 e2 : expr) | ||
| 78 | | EIf (e1 e2 e3 : expr) | ||
| 79 | with attr := | ||
| 80 | | Attr (τ : rec) (e : expr). | ||
| 81 | End expr. | ||
| 82 | |||
| 83 | Definition EId' x := EId x None. | ||
| 84 | Notation AttrR := (Attr REC). | ||
| 85 | Notation AttrN := (Attr NONREC). | ||
| 86 | Notation ESelect e x := (EBinOp SelectAttrOp e (ELit (LitString x))). | ||
| 87 | Notation ELet x e := (ELetAttr ABS (EAttr {[ x := AttrN e ]})). | ||
| 88 | Notation EWith := (ELetAttr WITH). | ||
| 89 | |||
| 90 | Definition attr_expr (α : attr) : expr := match α with Attr _ e => e end. | ||
| 91 | Definition attr_rec (α : attr) : rec := match α with Attr μ _ => μ end. | ||
| 92 | Definition attr_map (f : expr → expr) (α : attr) : attr := | ||
| 93 | match α with Attr μ e => Attr μ (f e) end. | ||
| 94 | |||
| 95 | Definition from_attr {A} (f g : expr → A) (α : attr) : A := | ||
| 96 | match α with AttrR e => f e | AttrN e => g e end. | ||
| 97 | |||
| 98 | Definition merge_kinded {A} (new old : kind * A) : option (kind * A) := | ||
| 99 | match new.1, old.1 with | ||
| 100 | | WITH, ABS => Some old | ||
| 101 | | _, _ => Some new | ||
| 102 | end. | ||
| 103 | Arguments merge_kinded {_} !_ !_ / : simpl nomatch. | ||
| 104 | Notation union_kinded := (union_with merge_kinded). | ||
| 105 | |||
| 106 | Definition no_recs : gmap string attr → Prop := | ||
| 107 | map_Forall (λ _ α, attr_rec α = NONREC). | ||
| 108 | |||
| 109 | Definition indirects (αs : gmap string attr) : gmap string (kind * expr) := | ||
| 110 | map_imap (λ x _, Some (ABS, ESelect (EAttr αs) x)) αs. | ||
| 111 | |||
| 112 | Fixpoint subst (ds : gmap string (kind * expr)) (e : expr) : expr := | ||
| 113 | match e with | ||
| 114 | | ELit b => ELit b | ||
| 115 | | EId x mkd => EId x $ union_kinded (ds !! x) mkd | ||
| 116 | | EAbs x e => EAbs x (subst ds e) | ||
| 117 | | EAbsMatch ms strict e => | ||
| 118 | EAbsMatch (fmap (M:=option) (subst ds) <$> ms) strict (subst ds e) | ||
| 119 | | EApp e1 e2 => EApp (subst ds e1) (subst ds e2) | ||
| 120 | | ESeq μ e1 e2 => ESeq μ (subst ds e1) (subst ds e2) | ||
| 121 | | EList es => EList (subst ds <$> es) | ||
| 122 | | EAttr αs => EAttr (attr_map (subst ds) <$> αs) | ||
| 123 | | ELetAttr k e1 e2 => ELetAttr k (subst ds e1) (subst ds e2) | ||
| 124 | | EBinOp op e1 e2 => EBinOp op (subst ds e1) (subst ds e2) | ||
| 125 | | EIf e1 e2 e3 => EIf (subst ds e1) (subst ds e2) (subst ds e3) | ||
| 126 | end. | ||
| 127 | |||
| 128 | Notation attr_subst ds := (attr_map (subst ds)). | ||
| 129 | |||
| 130 | Definition int_min : Z := -(1 ≪ 63). | ||
| 131 | Definition int_max : Z := 1 ≪ 63 - 1. | ||
| 132 | |||
| 133 | Definition int_ok (i : Z) : bool := bool_decide (int_min ≤ i ≤ int_max)%Z. | ||
| 134 | Definition num_ok (n : num) : bool := | ||
| 135 | match n with NInt i => int_ok i | _ => true end. | ||
| 136 | Definition base_lit_ok (bl : base_lit) : bool := | ||
| 137 | match bl with LitNum n => num_ok n | _ => true end. | ||
| 138 | |||
| 139 | Inductive final : mode → expr → Prop := | ||
| 140 | | ELitFinal μ bl : base_lit_ok bl → final μ (ELit bl) | ||
| 141 | | EAbsFinal μ x e : final μ (EAbs x e) | ||
| 142 | | EAbsMatchFinal μ ms strict e : final μ (EAbsMatch ms strict e) | ||
| 143 | | EListShallowFinal es : final SHALLOW (EList es) | ||
| 144 | | EListDeepFinal es : Forall (final DEEP) es → final DEEP (EList es) | ||
| 145 | | EAttrShallowFinal αs : no_recs αs → final SHALLOW (EAttr αs) | ||
| 146 | | EAttrDeepFinal αs : | ||
| 147 | no_recs αs → | ||
| 148 | map_Forall (λ _, final DEEP ∘ attr_expr) αs → | ||
| 149 | final DEEP (EAttr αs). | ||
| 150 | |||
| 151 | Fixpoint sem_eq_list (es1 es2 : list expr) : expr := | ||
| 152 | match es1, es2 with | ||
| 153 | | [], [] => ELit (LitBool true) | ||
| 154 | | e1 :: es1, e2 :: es2 => | ||
| 155 | EIf (EBinOp EqOp e1 e2) (sem_eq_list es1 es2) (ELit (LitBool false)) | ||
| 156 | | _, _ => ELit (LitBool false) | ||
| 157 | end. | ||
| 158 | |||
| 159 | Fixpoint sem_lt_list (es1 es2 : list expr) : expr := | ||
| 160 | match es1, es2 with | ||
| 161 | | [], _ => ELit (LitBool true) | ||
| 162 | | e1 :: es1, e2 :: es2 => | ||
| 163 | EIf (EBinOp LtOp e1 e2) (ELit (LitBool true)) $ | ||
| 164 | EIf (EBinOp EqOp e1 e2) (sem_lt_list es1 es2) (ELit (LitBool false)) | ||
| 165 | | _ :: _, [] => ELit (LitBool false) | ||
| 166 | end. | ||
| 167 | |||
| 168 | Definition sem_and_attr (es : gmap string expr) : expr := | ||
| 169 | map_fold_sorted attr_le | ||
| 170 | (λ _ e1 e2, EIf e1 e2 (ELit (LitBool false))) | ||
| 171 | (ELit (LitBool true)) es. | ||
| 172 | |||
| 173 | Definition sem_eq_attr (αs1 αs2 : gmap string attr) : expr := | ||
| 174 | sem_and_attr $ merge (λ mα1 mα2, | ||
| 175 | α1 ← mα1; α2 ← mα2; Some (EBinOp EqOp (attr_expr α1) (attr_expr α2))) αs1 αs2. | ||
| 176 | |||
| 177 | Definition num_to_float (n : num) : float := | ||
| 178 | match n with | ||
| 179 | | NInt i => Float.of_Z i | ||
| 180 | | NFloat f => f | ||
| 181 | end. | ||
| 182 | |||
| 183 | Definition sem_bin_op_lift | ||
| 184 | (fint : Z → Z → Z) (ffloat : float → float → float) | ||
| 185 | (n1 n2 : num) : option num := | ||
| 186 | match n1, n2 with | ||
| 187 | | NInt i1, NInt i2 => | ||
| 188 | let i := fint i1 i2 in | ||
| 189 | guard (int_ok i);; | ||
| 190 | Some (NInt i) | ||
| 191 | | _, _ => Some $ NFloat $ ffloat (num_to_float n1) (num_to_float n2) | ||
| 192 | end. | ||
| 193 | |||
| 194 | Definition sem_bin_rel_lift | ||
| 195 | (fint : Z → Z → bool) (ffloat : float → float → bool) | ||
| 196 | (n1 n2 : num) : bool := | ||
| 197 | match n1, n2 with | ||
| 198 | | NInt i1, NInt i2 => fint i1 i2 | ||
| 199 | | _, _ => ffloat (num_to_float n1) (num_to_float n2) | ||
| 200 | end. | ||
| 201 | |||
| 202 | Definition sem_eq_base_lit (bl1 bl2 : base_lit) : bool := | ||
| 203 | match bl1, bl2 with | ||
| 204 | | LitNum n1, LitNum n2 => sem_bin_rel_lift Z.eqb Float.eqb n1 n2 | ||
| 205 | | LitBool b1, LitBool b2 => bool_decide (b1 = b2) | ||
| 206 | | LitString s1, LitString s2 => bool_decide (s1 = s2) | ||
| 207 | | LitNull, LitNull => true | ||
| 208 | | _, _ => false | ||
| 209 | end. | ||
| 210 | |||
| 211 | (** Precondition e1 and e2 are final *) | ||
| 212 | Definition sem_eq (e1 e2 : expr) : option expr := | ||
| 213 | match e1, e2 with | ||
| 214 | | ELit bl1, ELit bl2 => Some $ ELit (LitBool (sem_eq_base_lit bl1 bl2)) | ||
| 215 | | EAbs _ _, EAbs _ _ => None | ||
| 216 | | EList es1, EList es2 => Some $ | ||
| 217 | if decide (length es1 = length es2) then sem_eq_list es1 es2 | ||
| 218 | else ELit $ LitBool false | ||
| 219 | | EAttr αs1, EAttr αs2 => Some $ | ||
| 220 | if decide (dom αs1 = dom αs2) then sem_eq_attr αs1 αs2 | ||
| 221 | else ELit $ LitBool false | ||
| 222 | | _, _ => Some $ ELit (LitBool false) | ||
| 223 | end. | ||
| 224 | |||
| 225 | Definition div_allowed (n : num) : bool := | ||
| 226 | match n with | ||
| 227 | | NInt n => bool_decide (n ≠ 0%Z) | ||
| 228 | | NFloat f => negb (Float.eqb f (Float.of_Z 0)) (* TODO: Check NaNs *) | ||
| 229 | end. | ||
| 230 | |||
| 231 | Definition sem_bin_op_num (op : bin_op) (n1 : num) : option (num → option base_lit) := | ||
| 232 | match op with | ||
| 233 | | AddOp => Some $ λ n2, | ||
| 234 | LitNum <$> sem_bin_op_lift Z.add Float.add n1 n2 | ||
| 235 | | SubOp => Some $ λ n2, | ||
| 236 | LitNum <$> sem_bin_op_lift Z.sub Float.sub n1 n2 | ||
| 237 | | MulOp => Some $ λ n2, | ||
| 238 | LitNum <$> sem_bin_op_lift Z.mul Float.mul n1 n2 | ||
| 239 | | DivOp => Some $ λ n2, | ||
| 240 | (* Quot can overflow: [MIN_INT `quot` -1] equals [MAX_INT + 1] *) | ||
| 241 | guard (div_allowed n2);; | ||
| 242 | LitNum <$> sem_bin_op_lift Z.quot Float.div n1 n2 | ||
| 243 | | AndOp => | ||
| 244 | i1 ← maybe NInt n1; | ||
| 245 | Some $ λ n2, i2 ← maybe NInt n2; | ||
| 246 | Some $ LitNum $ NInt $ Z.land i1 i2 | ||
| 247 | | OrOp => | ||
| 248 | i1 ← maybe NInt n1; | ||
| 249 | Some $ λ n2, i2 ← maybe NInt n2; | ||
| 250 | Some $ LitNum $ NInt $ Z.lor i1 i2 | ||
| 251 | | XOrOp => | ||
| 252 | i1 ← maybe NInt n1; | ||
| 253 | Some $ λ n2, i2 ← maybe NInt n2; | ||
| 254 | Some $ LitNum $ NInt $ Z.lxor i1 i2 | ||
| 255 | | LtOp => Some $ λ n2, | ||
| 256 | Some $ LitBool (sem_bin_rel_lift Z.ltb Float.ltb n1 n2) | ||
| 257 | | _ => None | ||
| 258 | end%Z. | ||
| 259 | |||
| 260 | Definition sem_bin_op_string (op : bin_op) : option (string → string → base_lit) := | ||
| 261 | match op with | ||
| 262 | | AddOp => Some $ λ s1 s2, LitString (s1 +:+ s2) | ||
| 263 | | LtOp => Some $ λ s1 s2, LitBool (bool_decide (strict attr_le s1 s2)) | ||
| 264 | | _ => None | ||
| 265 | end. | ||
| 266 | |||
| 267 | Definition type_of_num (n : num) : string := | ||
| 268 | match n with | ||
| 269 | | NInt _ => "int" | ||
| 270 | | NFloat _ => "float" | ||
| 271 | end. | ||
| 272 | |||
| 273 | Definition type_of_base_lit (bl : base_lit) : string := | ||
| 274 | match bl with | ||
| 275 | | LitNum n => type_of_num n | ||
| 276 | | LitBool _ => "bool" | ||
| 277 | | LitString _ => "string" | ||
| 278 | | LitNull => "null" | ||
| 279 | end. | ||
| 280 | |||
| 281 | Definition type_of_expr (e : expr) := | ||
| 282 | match e with | ||
| 283 | | ELit bl => Some (type_of_base_lit bl) | ||
| 284 | | EAbs _ _ | EAbsMatch _ _ _ => Some "lambda" | ||
| 285 | | EList _ => Some "list" | ||
| 286 | | EAttr _ => Some "set" | ||
| 287 | | _ => None | ||
| 288 | end. | ||
| 289 | |||
| 290 | (* Used for [RoundOp] *) | ||
| 291 | Definition float_to_bounded_Z (f : float) : Z := | ||
| 292 | match Float.to_Z f with | ||
| 293 | | Some x => if decide (int_ok x) then x else int_min | ||
| 294 | | None => int_min | ||
| 295 | end. | ||
| 296 | |||
| 297 | Inductive sem_bin_op : bin_op → expr → (expr → expr → Prop) → Prop := | ||
| 298 | | EqSem e1 : | ||
| 299 | sem_bin_op EqOp e1 (λ e2 e, sem_eq e1 e2 = Some e) | ||
| 300 | | LitNumSem op n1 f : | ||
| 301 | sem_bin_op_num op n1 = Some f → | ||
| 302 | sem_bin_op op (ELit (LitNum n1)) (λ e2 e, ∃ n2 bl, | ||
| 303 | e2 = ELit (LitNum n2) ∧ f n2 = Some bl ∧ e = ELit bl) | ||
| 304 | | RoundSem m n1 : | ||
| 305 | sem_bin_op (RoundOp m) (ELit (LitNum n1)) (λ e2 e, | ||
| 306 | e2 = ELit LitNull ∧ | ||
| 307 | e = ELit $ LitNum $ NInt $ float_to_bounded_Z $ Float.round m $ num_to_float n1) | ||
| 308 | | LitStringSem op s1 f : | ||
| 309 | sem_bin_op_string op = Some f → | ||
| 310 | sem_bin_op op (ELit (LitString s1)) (λ e2 e, ∃ s2, | ||
| 311 | e2 = ELit (LitString s2) ∧ e = ELit (f s1 s2)) | ||
| 312 | | MatchStringSem s : | ||
| 313 | sem_bin_op MatchStringOp (ELit (LitString s)) (λ e2 e, | ||
| 314 | e2 = ELit LitNull ∧ | ||
| 315 | match s with | ||
| 316 | | EmptyString => e = EAttr {[ | ||
| 317 | "empty" := AttrN (ELit (LitBool true)); | ||
| 318 | "head" := AttrN (ELit LitNull); | ||
| 319 | "tail" := AttrN (ELit LitNull) ]} | ||
| 320 | | String a s => e = EAttr {[ | ||
| 321 | "empty" := AttrN (ELit (LitBool false)); | ||
| 322 | "head" := AttrN (ELit (LitString (String a EmptyString))); | ||
| 323 | "tail" := AttrN (ELit (LitString s)) ]} | ||
| 324 | end) | ||
| 325 | | LtListSem es : | ||
| 326 | sem_bin_op LtOp (EList es) (λ e2 e, ∃ es', | ||
| 327 | e2 = EList es' ∧ | ||
| 328 | e = sem_lt_list es es') | ||
| 329 | | MatchListSem es : | ||
| 330 | sem_bin_op MatchListOp (EList es) (λ e2 e, | ||
| 331 | e2 = ELit LitNull ∧ | ||
| 332 | match es with | ||
| 333 | | [] => e = EAttr {[ | ||
| 334 | "empty" := AttrN (ELit (LitBool true)); | ||
| 335 | "head" := AttrN (ELit LitNull); | ||
| 336 | "tail" := AttrN (ELit LitNull) ]} | ||
| 337 | | e' :: es => e = EAttr {[ | ||
| 338 | "empty" := AttrN (ELit (LitBool false)); | ||
| 339 | "head" := AttrN e'; | ||
| 340 | "tail" := AttrN (EList es) ]} | ||
| 341 | end) | ||
| 342 | | AppendListSem es : | ||
| 343 | sem_bin_op AppendListOp (EList es) (λ e2 e, ∃ es', | ||
| 344 | e2 = EList es' ∧ | ||
| 345 | e = EList (es ++ es')) | ||
| 346 | | SelectAttrSem αs : | ||
| 347 | no_recs αs → | ||
| 348 | sem_bin_op SelectAttrOp (EAttr αs) (λ e2 e, ∃ x, | ||
| 349 | e2 = ELit (LitString x) ∧ αs !! x = Some (AttrN e)) | ||
| 350 | | UpdateAttrSem αs1 : | ||
| 351 | no_recs αs1 → | ||
| 352 | sem_bin_op UpdateAttrOp (EAttr αs1) (λ e2 e, ∃ αs2, | ||
| 353 | e2 = EAttr αs2 ∧ no_recs αs2 ∧ e = EAttr (αs2 ∪ αs1)) | ||
| 354 | | HasAttrSem αs : | ||
| 355 | no_recs αs → | ||
| 356 | sem_bin_op HasAttrOp (EAttr αs) (λ e2 e, ∃ x, | ||
| 357 | e2 = ELit (LitString x) ∧ e = ELit (LitBool (bool_decide (is_Some (αs !! x))))) | ||
| 358 | | DeleteAttrSem αs : | ||
| 359 | no_recs αs → | ||
| 360 | sem_bin_op DeleteAttrOp (EAttr αs) (λ e2 e, ∃ x, | ||
| 361 | e2 = ELit (LitString x) ∧ e = EAttr (delete x αs)) | ||
| 362 | | SingletonAttrSem x : | ||
| 363 | sem_bin_op SingletonAttrOp (ELit (LitString x)) (λ e2 e, | ||
| 364 | e2 = ELit LitNull ∧ | ||
| 365 | e = EAbs "t" (EAttr {[ x := AttrN (EId' "t") ]})) | ||
| 366 | | MatchAttrSem αs : | ||
| 367 | no_recs αs → | ||
| 368 | sem_bin_op MatchAttrOp (EAttr αs) (λ e2 e, | ||
| 369 | e2 = ELit LitNull ∧ | ||
| 370 | ((αs = ∅ ∧ | ||
| 371 | e = EAttr {[ | ||
| 372 | "empty" := AttrN (ELit (LitBool true)); | ||
| 373 | "key" := AttrN (ELit LitNull); | ||
| 374 | "head" := AttrN (ELit LitNull); | ||
| 375 | "tail" := AttrN (ELit LitNull) ]}) ∨ | ||
| 376 | (∃ x e', | ||
| 377 | αs !! x = Some (AttrN e') ∧ | ||
| 378 | (∀ y, is_Some (αs !! y) → attr_le x y) ∧ | ||
| 379 | e = EAttr {[ | ||
| 380 | "empty" := AttrN (ELit (LitBool false)); | ||
| 381 | "key" := AttrN (ELit (LitString x)); | ||
| 382 | "head" := AttrN e'; | ||
| 383 | "tail" := AttrN (EAttr (delete x αs)) ]}))) | ||
| 384 | | FunctionArgsAbsSem x e' : | ||
| 385 | sem_bin_op FunctionArgsOp (EAbs x e') (λ e2 e, | ||
| 386 | e2 = ELit LitNull ∧ | ||
| 387 | e = EAttr ∅) | ||
| 388 | | FunctionArgsAbsMatchSem ms strict e' : | ||
| 389 | sem_bin_op FunctionArgsOp (EAbsMatch ms strict e') (λ e2 e, | ||
| 390 | e2 = ELit LitNull ∧ | ||
| 391 | e = EAttr (AttrN ∘ ELit ∘ LitBool ∘ from_option (λ _, true) false <$> ms)) | ||
| 392 | | TypeOfSem e1 : | ||
| 393 | sem_bin_op TypeOfOp e1 (λ e2 e, ∃ x, | ||
| 394 | e2 = ELit LitNull ∧ | ||
| 395 | type_of_expr e1 = Some x ∧ | ||
| 396 | e = ELit (LitString x)). | ||
| 397 | |||
| 398 | Inductive matches : | ||
| 399 | gmap string expr → gmap string (option expr) → bool → gmap string attr → Prop := | ||
| 400 | | MatchEmpty strict : | ||
| 401 | matches ∅ ∅ strict ∅ | ||
| 402 | | MatchAny es : | ||
| 403 | matches es ∅ false ∅ | ||
| 404 | | MatchAvail x e es ms md strict βs : | ||
| 405 | es !! x = None → | ||
| 406 | ms !! x = None → | ||
| 407 | matches es ms strict βs → | ||
| 408 | matches (<[x:=e]> es) (<[x:=md]> ms) strict (<[x:=AttrN e]> βs) | ||
| 409 | | MatchOptDefault x es ms d strict βs : | ||
| 410 | es !! x = None → | ||
| 411 | ms !! x = None → | ||
| 412 | matches es ms strict βs → | ||
| 413 | matches es (<[x:=Some d]> ms) strict (<[x:=AttrR d]> βs). | ||
| 414 | |||
| 415 | Reserved Notation "e1 -{ μ }-> e2" | ||
| 416 | (right associativity, at level 55, μ at level 1, format "e1 -{ μ }-> e2"). | ||
| 417 | |||
| 418 | Inductive ctx1 : mode → mode → (expr → expr) → Prop := | ||
| 419 | | CList es1 es2 : | ||
| 420 | Forall (final DEEP) es1 → | ||
| 421 | ctx1 DEEP DEEP (λ e, EList (es1 ++ e :: es2)) | ||
| 422 | | CAttr αs x : | ||
| 423 | no_recs αs → | ||
| 424 | αs !! x = None → | ||
| 425 | (∀ y α, αs !! y = Some α → final DEEP (attr_expr α) ∨ attr_le x y) → | ||
| 426 | ctx1 DEEP DEEP (λ e, EAttr (<[x:=AttrN e]> αs)) | ||
| 427 | | CAppL μ e2 : | ||
| 428 | ctx1 SHALLOW μ (λ e1, EApp e1 e2) | ||
| 429 | | CAppR μ ms strict e1 : | ||
| 430 | ctx1 SHALLOW μ (EApp (EAbsMatch ms strict e1)) | ||
| 431 | | CSeq μ μ' e2 : | ||
| 432 | ctx1 μ' μ (λ e1, ESeq μ' e1 e2) | ||
| 433 | | CLetAttr μ k e2 : | ||
| 434 | ctx1 SHALLOW μ (λ e1, ELetAttr k e1 e2) | ||
| 435 | | CBinOpL μ op e2 : | ||
| 436 | ctx1 SHALLOW μ (λ e1, EBinOp op e1 e2) | ||
| 437 | | CBinOpR μ op e1 Φ : | ||
| 438 | final SHALLOW e1 → | ||
| 439 | sem_bin_op op e1 Φ → | ||
| 440 | ctx1 SHALLOW μ (EBinOp op e1) | ||
| 441 | | CIf μ e2 e3 : | ||
| 442 | ctx1 SHALLOW μ (λ e1, EIf e1 e2 e3). | ||
| 443 | |||
| 444 | Inductive step : mode → relation expr := | ||
| 445 | | Sβ μ x e1 e2 : | ||
| 446 | EApp (EAbs x e1) e2 -{μ}-> subst {[x:=(ABS, e2)]} e1 | ||
| 447 | | SβMatch μ ms strict e1 αs βs : | ||
| 448 | no_recs αs → | ||
| 449 | matches (attr_expr <$> αs) ms strict βs → | ||
| 450 | EApp (EAbsMatch ms strict e1) (EAttr αs) -{μ}-> | ||
| 451 | subst (indirects βs) e1 | ||
| 452 | | SFunctor μ αs e1 e2 : | ||
| 453 | no_recs αs → | ||
| 454 | αs !! "__functor" = Some (AttrN e1) → | ||
| 455 | EApp (EAttr αs) e2 -{μ}-> EApp (EApp e1 (EAttr αs)) e2 | ||
| 456 | | SSeqFinal μ μ' e1 e2 : | ||
| 457 | final μ' e1 → ESeq μ' e1 e2 -{μ}-> e2 | ||
| 458 | | SLetAttrAttr μ k αs e : | ||
| 459 | no_recs αs → | ||
| 460 | ELetAttr k (EAttr αs) e -{μ}-> subst ((k,.) ∘ attr_expr <$> αs) e | ||
| 461 | | SAttr_rec μ αs : | ||
| 462 | ¬no_recs αs → | ||
| 463 | EAttr αs -{μ}-> | ||
| 464 | EAttr (AttrN ∘ from_attr (subst (indirects αs)) id <$> αs) | ||
| 465 | | SBinOp μ op e1 Φ e2 e : | ||
| 466 | final SHALLOW e1 → | ||
| 467 | final SHALLOW e2 → | ||
| 468 | sem_bin_op op e1 Φ → Φ e2 e → | ||
| 469 | EBinOp op e1 e2 -{μ}-> e | ||
| 470 | | SIfBool μ b e2 e3 : | ||
| 471 | EIf (ELit (LitBool b)) e2 e3 -{μ}-> if b then e2 else e3 | ||
| 472 | | SId μ x k e : | ||
| 473 | EId x (Some (k,e)) -{μ}-> e | ||
| 474 | | SCtx K μ μ' e e' : | ||
| 475 | ctx1 μ μ' K → e -{μ}-> e' → K e -{μ'}-> K e' | ||
| 476 | where "e1 -{ μ }-> e2" := (step μ e1 e2). | ||
| 477 | |||
| 478 | Notation "e1 -{ μ }->* e2" := (rtc (step μ) e1 e2) | ||
| 479 | (right associativity, at level 55, μ at level 1, format "e1 -{ μ }->* e2"). | ||
| 480 | Notation "e1 -{ μ }->+ e2" := (tc (step μ) e1 e2) | ||
| 481 | (right associativity, at level 55, μ at level 1, format "e1 -{ μ }->+ e2"). | ||
| 482 | |||
| 483 | Definition stuck (μ : mode) (e : expr) : Prop := | ||
| 484 | nf (step μ) e ∧ ¬final μ e. | ||
| 485 | |||
| 486 | (** Induction *) | ||
| 487 | Fixpoint expr_size (e : expr) : nat := | ||
| 488 | match e with | ||
| 489 | | ELit _ => 1 | ||
| 490 | | EId _ mkd => S (from_option (expr_size ∘ snd) 1 mkd) | ||
| 491 | | EAbs _ d => S (expr_size d) | ||
| 492 | | EAbsMatch ms _ e => | ||
| 493 | S (map_sum_with (from_option expr_size 1) ms + expr_size e) | ||
| 494 | | EApp e1 e2 | ESeq _ e1 e2 => S (expr_size e1 + expr_size e2) | ||
| 495 | | EList es => S (sum_list_with expr_size es) | ||
| 496 | | EAttr eτs => S (map_sum_with (expr_size ∘ attr_expr) eτs) | ||
| 497 | | ELetAttr _ e1 e2 => S (expr_size e1 + expr_size e2) | ||
| 498 | | EBinOp _ e1 e2 => S (expr_size e1 + expr_size e2) | ||
| 499 | | EIf e1 e2 e3 => S (expr_size e1 + expr_size e2 + expr_size e3) | ||
| 500 | end. | ||
| 501 | |||
| 502 | Lemma expr_ind (P : expr → Prop) : | ||
| 503 | (∀ b, P (ELit b)) → | ||
| 504 | (∀ x mkd, from_option (P ∘ snd) True mkd → P (EId x mkd)) → | ||
| 505 | (∀ x e, P e → P (EAbs x e)) → | ||
| 506 | (∀ ms strict e, | ||
| 507 | map_Forall (λ _, from_option P True) ms → P e → P (EAbsMatch ms strict e)) → | ||
| 508 | (∀ e1 e2, P e1 → P e2 → P (EApp e1 e2)) → | ||
| 509 | (∀ μ e1 e2, P e1 → P e2 → P (ESeq μ e1 e2)) → | ||
| 510 | (∀ es, Forall P es → P (EList es)) → | ||
| 511 | (∀ αs, map_Forall (λ _, P ∘ attr_expr) αs → P (EAttr αs)) → | ||
| 512 | (∀ k e1 e2, P e1 → P e2 → P (ELetAttr k e1 e2)) → | ||
| 513 | (∀ op e1 e2, P e1 → P e2 → P (EBinOp op e1 e2)) → | ||
| 514 | (∀ e1 e2 e3, P e1 → P e2 → P e3 → P (EIf e1 e2 e3)) → | ||
| 515 | ∀ e, P e. | ||
| 516 | Proof. | ||
| 517 | intros Hlit Hid Habs Hmatch Happ Hseq Hlist Hattr Hlet Hop Hif e. | ||
| 518 | induction (Nat.lt_wf_0_projected expr_size e) as [e _ IH]. | ||
| 519 | destruct e; repeat destruct select (option _); simpl in *; eauto with lia. | ||
| 520 | - apply Hmatch; [|by eauto with lia]=> y [e'|] Hx //=. apply IH, Nat.lt_succ_r. | ||
| 521 | etrans; [|apply Nat.le_add_r]. | ||
| 522 | eapply (map_sum_with_lookup_le (from_option expr_size 1) _ _ _ Hx). | ||
| 523 | - apply Hlist, Forall_forall=> e ?. apply IH, Nat.lt_succ_r. | ||
| 524 | by apply sum_list_with_in. | ||
| 525 | - apply Hattr, map_Forall_lookup=> y e ?. apply IH, Nat.lt_succ_r. | ||
| 526 | by eapply (map_sum_with_lookup_le (expr_size ∘ attr_expr)). | ||
| 527 | Qed. | ||
diff --git a/theories/nix/operational_props.v b/theories/nix/operational_props.v new file mode 100644 index 0000000..4041bfe --- /dev/null +++ b/theories/nix/operational_props.v | |||
| @@ -0,0 +1,680 @@ | |||
| 1 | From mininix Require Export utils nix.operational. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | (** Properties of operational semantics *) | ||
| 5 | Lemma float_to_bounded_Z_ok f : int_ok (float_to_bounded_Z f). | ||
| 6 | Proof. | ||
| 7 | rewrite /float_to_bounded_Z. | ||
| 8 | destruct (Float.to_Z f); simplify_option_eq; done. | ||
| 9 | Qed. | ||
| 10 | |||
| 11 | Lemma int_ok_alt i : | ||
| 12 | int_ok i ↔ ∀ n, (63 ≤ n)%Z → Z.testbit i n = bool_decide (i < 0)%Z. | ||
| 13 | Proof. | ||
| 14 | rewrite -Z.bounded_iff_bits //. | ||
| 15 | rewrite /int_ok bool_decide_spec /int_min /int_max Z.shiftl_1_l. lia. | ||
| 16 | Qed. | ||
| 17 | |||
| 18 | Lemma int_ok_land i1 i2 : int_ok i1 → int_ok i2 → int_ok (Z.land i1 i2). | ||
| 19 | Proof. | ||
| 20 | rewrite !int_ok_alt=> Hi1 Hi2 n ?. rewrite Z.land_spec Hi1 // Hi2 //. | ||
| 21 | apply eq_bool_prop_intro. rewrite andb_True !bool_decide_spec Z.land_neg //. | ||
| 22 | Qed. | ||
| 23 | |||
| 24 | Lemma int_ok_lor i1 i2 : int_ok i1 → int_ok i2 → int_ok (Z.lor i1 i2). | ||
| 25 | Proof. | ||
| 26 | rewrite !int_ok_alt=> Hi1 Hi2 n ?. rewrite Z.lor_spec Hi1 // Hi2 //. | ||
| 27 | apply eq_bool_prop_intro. rewrite orb_True !bool_decide_spec Z.lor_neg //. | ||
| 28 | Qed. | ||
| 29 | |||
| 30 | Lemma int_ok_lxor i1 i2 : int_ok i1 → int_ok i2 → int_ok (Z.lxor i1 i2). | ||
| 31 | Proof. | ||
| 32 | rewrite !int_ok_alt=> Hi1 Hi2 n ?. rewrite Z.lxor_spec Hi1 // Hi2 //. | ||
| 33 | apply eq_bool_prop_intro. rewrite xorb_True !bool_decide_spec. | ||
| 34 | rewrite !Z.lt_nge Z.lxor_nonneg. lia. | ||
| 35 | Qed. | ||
| 36 | |||
| 37 | Lemma sem_bin_op_num_ok {op f n1 n2 bl} : | ||
| 38 | num_ok n1 → num_ok n2 → | ||
| 39 | sem_bin_op_num op n1 = Some f → f n2 = Some bl → base_lit_ok bl. | ||
| 40 | Proof. | ||
| 41 | intros; destruct op, n1, n2; simplify_option_eq; | ||
| 42 | try (by apply (bool_decide_pack _)); | ||
| 43 | auto using int_ok_land, int_ok_lor, int_ok_lxor. | ||
| 44 | Qed. | ||
| 45 | |||
| 46 | Lemma sem_bin_op_string_ok {op f s1 s2} : | ||
| 47 | sem_bin_op_string op = Some f → base_lit_ok (f s1 s2). | ||
| 48 | Proof. intros; destruct op; naive_solver. Qed. | ||
| 49 | |||
| 50 | Global Hint Extern 0 (no_recs (_ <$> _)) => by apply map_Forall_fmap : core. | ||
| 51 | |||
| 52 | Ltac inv_step := repeat | ||
| 53 | match goal with | ||
| 54 | | H : ¬no_recs (_ <$> _) |- _ => destruct H; by apply map_Forall_fmap | ||
| 55 | | H : ?e -{_}-> _ |- _ => assert_succeeds (is_app_constructor e); inv H | ||
| 56 | | H : ctx1 _ _ ?K |- _ => is_var K; inv H | ||
| 57 | end. | ||
| 58 | |||
| 59 | Global Instance Attr_inj τ : Inj (=) (=) (Attr τ). | ||
| 60 | Proof. by injection 1. Qed. | ||
| 61 | |||
| 62 | Lemma fmap_attr_expr_Attr τ (es : gmap string expr) : | ||
| 63 | attr_expr <$> (Attr τ <$> es) = es. | ||
| 64 | Proof. apply map_eq=> x. rewrite !lookup_fmap. by destruct (_ !! _). Qed. | ||
| 65 | |||
| 66 | Lemma no_recs_insert αs x e : no_recs αs → no_recs (<[x:=AttrN e]> αs). | ||
| 67 | Proof. by apply map_Forall_insert_2. Qed. | ||
| 68 | Lemma no_recs_insert_inv αs x τ e : | ||
| 69 | αs !! x = None → no_recs (<[x:=Attr τ e]> αs) → no_recs αs. | ||
| 70 | Proof. intros ??%map_Forall_insert; naive_solver. Qed. | ||
| 71 | Lemma no_recs_lookup αs x τ e : no_recs αs → αs !! x = Some (Attr τ e) → τ = NONREC. | ||
| 72 | Proof. intros Hall. apply Hall. Qed. | ||
| 73 | |||
| 74 | Lemma no_recs_attr_subst αs ds : no_recs αs → no_recs (attr_subst ds <$> αs). | ||
| 75 | Proof. | ||
| 76 | intros. eapply map_Forall_fmap, map_Forall_impl; [done|]. by intros ? [[]] [=]. | ||
| 77 | Qed. | ||
| 78 | |||
| 79 | Lemma from_attr_no_recs {A} (f g : expr → A) (αs : gmap string attr) : | ||
| 80 | no_recs αs → from_attr f g <$> αs = g ∘ attr_expr <$> αs. | ||
| 81 | Proof. | ||
| 82 | intros Hrecs. apply map_eq=> x. rewrite !lookup_fmap. specialize (Hrecs x). | ||
| 83 | destruct (αs !! x) as [[]|] eqn:?; naive_solver. | ||
| 84 | Qed. | ||
| 85 | |||
| 86 | Lemma sem_and_attr_empty : sem_and_attr ∅ = ELit (LitBool true). | ||
| 87 | Proof. done. Qed. | ||
| 88 | Lemma sem_and_attr_insert es x e : | ||
| 89 | es !! x = None → (∀ y, is_Some (es !! y) → attr_le x y) → | ||
| 90 | sem_and_attr (<[x:=e]> es) = EIf e (sem_and_attr es) (ELit (LitBool false)). | ||
| 91 | Proof. intros. by rewrite /sem_and_attr map_fold_sorted_insert. Qed. | ||
| 92 | |||
| 93 | Lemma matches_strict es ms ds x e : | ||
| 94 | es !! x = None → | ||
| 95 | ms !! x = None → | ||
| 96 | matches es ms false ds → | ||
| 97 | matches (<[x:=e]> es) ms false ds. | ||
| 98 | Proof. | ||
| 99 | remember false as strict. | ||
| 100 | induction 3; simplify_eq/=; | ||
| 101 | repeat match goal with | ||
| 102 | | H : <[ _ := _ ]> _ !! _ = None |- _ => apply lookup_insert_None in H as [??] | ||
| 103 | | _ => rewrite (insert_commute _ x) // | ||
| 104 | | _ => constructor | ||
| 105 | | _ => apply lookup_insert_None | ||
| 106 | end; eauto. | ||
| 107 | Qed. | ||
| 108 | |||
| 109 | Lemma subst_empty e : subst ∅ e = e. | ||
| 110 | Proof. | ||
| 111 | induction e; repeat destruct select (option _); do 2 f_equal/=; auto. | ||
| 112 | - apply map_eq=> x. rewrite lookup_fmap. | ||
| 113 | destruct (_ !! x) as [[e'|]|] eqn:Hx; do 2 f_equal/=. apply (H _ _ Hx). | ||
| 114 | - induction H; f_equal/=; auto. | ||
| 115 | - apply map_eq; intros i. rewrite lookup_fmap. | ||
| 116 | destruct (_ !! i) as [[τ e]|] eqn:?; do 2 f_equal/=. | ||
| 117 | by eapply (H _ (Attr _ _)). | ||
| 118 | Qed. | ||
| 119 | |||
| 120 | Lemma subst_union ds1 ds2 e : | ||
| 121 | subst (union_kinded ds1 ds2) e = subst ds1 (subst ds2 e). | ||
| 122 | Proof. | ||
| 123 | revert ds1 ds2. induction e; intros ds1 ds2; f_equal/=; auto. | ||
| 124 | - rewrite lookup_union_with. | ||
| 125 | destruct mkd as [[[]]|], | ||
| 126 | (ds1 !! x) as [[[] t1]|], (ds2 !! x) as [[[] t2]|]; naive_solver. | ||
| 127 | - apply map_eq=> y. rewrite !lookup_fmap. | ||
| 128 | destruct (_ !! y) as [[e'|]|] eqn:Hy; do 2 f_equal/=. | ||
| 129 | rewrite -(H _ _ Hy) //. | ||
| 130 | - induction H; f_equal/=; auto. | ||
| 131 | - apply map_eq=> y. rewrite !lookup_fmap. | ||
| 132 | destruct (_ !! y) as [[τ e]|] eqn:Hy; do 2 f_equal/=. | ||
| 133 | rewrite -(H _ _ Hy) //. | ||
| 134 | Qed. | ||
| 135 | |||
| 136 | Lemma SAppL μ e1 e1' e2 : | ||
| 137 | e1 -{SHALLOW}-> e1' → EApp e1 e2 -{μ}-> EApp e1' e2. | ||
| 138 | Proof. apply (SCtx (λ e, EApp e _)). constructor. Qed. | ||
| 139 | Lemma SAppR μ ms strict e1 e2 e2' : | ||
| 140 | e2 -{SHALLOW}-> e2' → | ||
| 141 | EApp (EAbsMatch ms strict e1) e2 -{μ}-> EApp (EAbsMatch ms strict e1) e2'. | ||
| 142 | Proof. apply SCtx. constructor. Qed. | ||
| 143 | Lemma SSeq μ μ' e1 e1' e2 : | ||
| 144 | e1 -{μ'}-> e1' → ESeq μ' e1 e2 -{μ}-> ESeq μ' e1' e2. | ||
| 145 | Proof. apply (SCtx (λ e, ESeq _ e _)). constructor. Qed. | ||
| 146 | Lemma SList es1 e e' es2 : | ||
| 147 | Forall (final DEEP) es1 → | ||
| 148 | e -{DEEP}-> e' → | ||
| 149 | EList (es1 ++ e :: es2) -{DEEP}-> EList (es1 ++ e' :: es2). | ||
| 150 | Proof. intros ?. apply (SCtx (λ e, EList (_ ++ e :: _))). by constructor. Qed. | ||
| 151 | Lemma SAttr αs x e e' : | ||
| 152 | no_recs αs → | ||
| 153 | αs !! x = None → | ||
| 154 | (∀ y α, αs !! y = Some α → final DEEP (attr_expr α) ∨ attr_le x y) → | ||
| 155 | e -{DEEP}-> e' → | ||
| 156 | EAttr (<[x:=AttrN e]> αs) -{DEEP}-> EAttr (<[x:=AttrN e']> αs). | ||
| 157 | Proof. intros ???. apply (SCtx (λ e, EAttr (<[x:=AttrN e]> _))). by constructor. Qed. | ||
| 158 | Lemma SLetAttr μ k e1 e1' e2 : | ||
| 159 | e1 -{SHALLOW}-> e1' → ELetAttr k e1 e2 -{μ}-> ELetAttr k e1' e2. | ||
| 160 | Proof. apply (SCtx (λ e, ELetAttr _ e _)). constructor. Qed. | ||
| 161 | Lemma SBinOpL μ op e1 e1' e2 : | ||
| 162 | e1 -{SHALLOW}-> e1' → EBinOp op e1 e2 -{μ}-> EBinOp op e1' e2. | ||
| 163 | Proof. apply (SCtx (λ e, EBinOp _ e _)). constructor. Qed. | ||
| 164 | Lemma SBinOpR μ op e1 Φ e2 e2' : | ||
| 165 | final SHALLOW e1 → sem_bin_op op e1 Φ → | ||
| 166 | e2 -{SHALLOW}-> e2' → EBinOp op e1 e2 -{μ}-> EBinOp op e1 e2'. | ||
| 167 | Proof. intros ??. apply SCtx. by econstructor. Qed. | ||
| 168 | Lemma SIf μ e1 e1' e2 e3 : | ||
| 169 | e1 -{SHALLOW}-> e1' → EIf e1 e2 e3 -{μ}-> EIf e1' e2 e3. | ||
| 170 | Proof. apply (SCtx (λ e, EIf e _ _)). constructor. Qed. | ||
| 171 | |||
| 172 | Global Hint Constructors step : step. | ||
| 173 | Global Hint Resolve SAppL SAppR SSeq SList SAttr SLetAttr SBinOpL SBinOpR SIf : step. | ||
| 174 | |||
| 175 | Lemma step_not_final μ e1 e2 : e1 -{μ}-> e2 → ¬final μ e1. | ||
| 176 | Proof. | ||
| 177 | assert (∀ (αs : gmap string attr) x μ e, | ||
| 178 | map_Forall (λ _, final DEEP ∘ attr_expr) (<[x:=Attr μ e]> αs) → final DEEP e). | ||
| 179 | { intros αs x μ' e Hall. eapply (Hall _ (Attr _ _)), lookup_insert. } | ||
| 180 | induction 1; inv 1; inv_step; decompose_Forall; naive_solver. | ||
| 181 | Qed. | ||
| 182 | Lemma final_nf μ e : final μ e → nf (step μ) e. | ||
| 183 | Proof. by intros ? [??%step_not_final]. Qed. | ||
| 184 | |||
| 185 | Lemma step_any_shallow μ e1 e2 : | ||
| 186 | e1 -{μ}-> e2 → e1 -{SHALLOW}-> e2 ∨ final SHALLOW e1. | ||
| 187 | Proof. | ||
| 188 | induction 1; inv_step; | ||
| 189 | naive_solver eauto using final, no_recs_insert with step. | ||
| 190 | Qed. | ||
| 191 | |||
| 192 | Lemma step_shallow_any μ e1 e2 : e1 -{SHALLOW}-> e2 → e1 -{μ}-> e2. | ||
| 193 | Proof. | ||
| 194 | remember SHALLOW as μ'. induction 1; inv_step; simplify_eq/=; eauto with step. | ||
| 195 | Qed. | ||
| 196 | Lemma steps_shallow_any μ e1 e2 : e1 -{SHALLOW}->* e2 → e1 -{μ}->* e2. | ||
| 197 | Proof. induction 1; eauto using rtc, step_shallow_any. Qed. | ||
| 198 | Lemma final_any_shallow μ e : final μ e → final SHALLOW e. | ||
| 199 | Proof. destruct μ; [done|]. induction 1; simplify_eq/=; eauto using final. Qed. | ||
| 200 | Lemma stuck_shallow_any μ e : stuck SHALLOW e → stuck μ e. | ||
| 201 | Proof. | ||
| 202 | intros [Hnf Hfinal]. split; [|naive_solver eauto using final_any_shallow]. | ||
| 203 | intros [e' Hstep%step_any_shallow]; naive_solver. | ||
| 204 | Qed. | ||
| 205 | |||
| 206 | Lemma step_final_shallow μ e1 e2 : | ||
| 207 | final SHALLOW e1 → e1 -{μ}-> e2 → final SHALLOW e2. | ||
| 208 | Proof. | ||
| 209 | induction 1; intros; inv_step; decompose_Forall; | ||
| 210 | eauto using step, final, no_recs_insert; try done. | ||
| 211 | - by odestruct step_not_final. | ||
| 212 | - apply map_Forall_insert in H0 as [??]; simpl in *; last done. | ||
| 213 | by odestruct step_not_final. | ||
| 214 | Qed. | ||
| 215 | |||
| 216 | Lemma SAppL_rtc μ e1 e1' e2 : | ||
| 217 | e1 -{SHALLOW}->* e1' → EApp e1 e2 -{μ}->* EApp e1' e2. | ||
| 218 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 219 | Lemma SAppR_rtc μ ms strict e1 e2 e2' : | ||
| 220 | e2 -{SHALLOW}->* e2' → | ||
| 221 | EApp (EAbsMatch ms strict e1) e2 -{μ}->* EApp (EAbsMatch ms strict e1) e2'. | ||
| 222 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 223 | Lemma SSeq_rtc μ μ' e1 e1' e2 : | ||
| 224 | e1 -{μ'}->* e1' → ESeq μ' e1 e2 -{μ}->* ESeq μ' e1' e2. | ||
| 225 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 226 | Lemma SList_rtc es1 e e' es2 : | ||
| 227 | Forall (final DEEP) es1 → | ||
| 228 | e -{DEEP}->* e' → | ||
| 229 | EList (es1 ++ e :: es2) -{DEEP}->* EList (es1 ++ e' :: es2). | ||
| 230 | Proof. induction 2; econstructor; eauto with step. Qed. | ||
| 231 | Lemma SLetAttr_rtc μ k e1 e1' e2 : | ||
| 232 | e1 -{SHALLOW}->* e1' → ELetAttr k e1 e2 -{μ}->* ELetAttr k e1' e2. | ||
| 233 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 234 | Lemma SBinOpL_rtc μ op e1 e1' e2 : | ||
| 235 | e1 -{SHALLOW}->* e1' → EBinOp op e1 e2 -{μ}->* EBinOp op e1' e2. | ||
| 236 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 237 | Lemma SBinOpR_rtc μ op e1 Φ e2 e2' : | ||
| 238 | final SHALLOW e1 → sem_bin_op op e1 Φ → | ||
| 239 | e2 -{SHALLOW}->* e2' → EBinOp op e1 e2 -{μ}->* EBinOp op e1 e2'. | ||
| 240 | Proof. induction 3; econstructor; eauto with step. Qed. | ||
| 241 | Lemma SIf_rtc μ e1 e1' e2 e3 : | ||
| 242 | e1 -{SHALLOW}->* e1' → EIf e1 e2 e3 -{μ}->* EIf e1' e2 e3. | ||
| 243 | Proof. induction 1; econstructor; eauto with step. Qed. | ||
| 244 | |||
| 245 | Lemma SApp_tc μ e1 e1' e2 : | ||
| 246 | e1 -{SHALLOW}->+ e1' → EApp e1 e2 -{μ}->+ EApp e1' e2. | ||
| 247 | Proof. induction 1; by econstructor; eauto with step. Qed. | ||
| 248 | Lemma SSeq_tc μ μ' e1 e1' e2 : | ||
| 249 | e1 -{μ'}->+ e1' → ESeq μ' e1 e2 -{μ}->+ ESeq μ' e1' e2. | ||
| 250 | Proof. induction 1; by econstructor; eauto with step. Qed. | ||
| 251 | Lemma SList_tc es1 e e' es2 : | ||
| 252 | Forall (final DEEP) es1 → | ||
| 253 | e -{DEEP}->+ e' → | ||
| 254 | EList (es1 ++ e :: es2) -{DEEP}->+ EList (es1 ++ e' :: es2). | ||
| 255 | Proof. induction 2; by econstructor; eauto with step. Qed. | ||
| 256 | Lemma SLetAttr_tc μ k e1 e1' e2 : | ||
| 257 | e1 -{SHALLOW}->+ e1' → ELetAttr k e1 e2 -{μ}->+ ELetAttr k e1' e2. | ||
| 258 | Proof. induction 1; by econstructor; eauto with step. Qed. | ||
| 259 | Lemma SBinOpL_tc μ op e1 e1' e2 : | ||
| 260 | e1 -{SHALLOW}->+ e1' → EBinOp op e1 e2 -{μ}->+ EBinOp op e1' e2. | ||
| 261 | Proof. induction 1; by econstructor; eauto with step. Qed. | ||
| 262 | Lemma SBinOpR_tc μ op e1 Φ e2 e2' : | ||
| 263 | final SHALLOW e1 → sem_bin_op op e1 Φ → | ||
| 264 | e2 -{SHALLOW}->+ e2' → EBinOp op e1 e2 -{μ}->+ EBinOp op e1 e2'. | ||
| 265 | Proof. induction 3; by econstructor; eauto with step. Qed. | ||
| 266 | Lemma SIf_tc μ e1 e1' e2 e3 : | ||
| 267 | e1 -{SHALLOW}->+ e1' → EIf e1 e2 e3 -{μ}->+ EIf e1' e2 e3. | ||
| 268 | Proof. induction 1; by econstructor; eauto with step. Qed. | ||
| 269 | |||
| 270 | Lemma SList_inv es1 e2 : | ||
| 271 | EList es1 -{DEEP}-> e2 ↔ ∃ ds1 ds2 e e', | ||
| 272 | es1 = ds1 ++ e :: ds2 ∧ e2 = EList (ds1 ++ e' :: ds2) ∧ | ||
| 273 | Forall (final DEEP) ds1 ∧ | ||
| 274 | e -{DEEP}-> e'. | ||
| 275 | Proof. split; intros; inv_step; naive_solver eauto using SList. Qed. | ||
| 276 | |||
| 277 | Lemma List_nf_cons_final es e : | ||
| 278 | final DEEP e → | ||
| 279 | nf (step DEEP) (EList es) → | ||
| 280 | nf (step DEEP) (EList (e :: es)). | ||
| 281 | Proof. | ||
| 282 | intros Hfinal Hnf [e' (ds1 & ds2 & e1 & e2 & ? & -> & Hds1 & Hstep)%SList_inv]. | ||
| 283 | destruct Hds1; simplify_eq/=. | ||
| 284 | - by apply step_not_final in Hstep. | ||
| 285 | - naive_solver eauto with step. | ||
| 286 | Qed. | ||
| 287 | Lemma List_nf_cons es e : | ||
| 288 | ¬final DEEP e → | ||
| 289 | nf (step DEEP) e → | ||
| 290 | nf (step DEEP) (EList (e :: es)). | ||
| 291 | Proof. | ||
| 292 | intros Hfinal Hnf [e' (ds1 & ds2 & e1 & e2 & ? & -> & Hds1 & Hstep)%SList_inv]. | ||
| 293 | destruct Hds1; naive_solver. | ||
| 294 | Qed. | ||
| 295 | |||
| 296 | Lemma List_steps_cons es1 es2 e : | ||
| 297 | final DEEP e → | ||
| 298 | EList es1 -{DEEP}->* EList es2 → | ||
| 299 | EList (e :: es1) -{DEEP}->* EList (e :: es2). | ||
| 300 | Proof. | ||
| 301 | intros ? Hstep. | ||
| 302 | remember (EList es1) as e1 eqn:He1; remember (EList es2) as e2 eqn:He2. | ||
| 303 | revert es1 es2 He1 He2. | ||
| 304 | induction Hstep as [|e1 e2 e3 Hstep Hstep' IH]; | ||
| 305 | intros es1 es3 ??; simplify_eq/=; [done|]. | ||
| 306 | inv_step. eapply rtc_l; [apply (SList (_ :: _))|]; naive_solver. | ||
| 307 | Qed. | ||
| 308 | |||
| 309 | Lemma SAttr_rec_rtc μ αs : | ||
| 310 | EAttr αs -{μ}->* | ||
| 311 | EAttr (AttrN ∘ from_attr (subst (indirects αs)) id <$> αs). | ||
| 312 | Proof. | ||
| 313 | destruct (decide (no_recs αs)) as [Hαs|]; [|by eauto using rtc_once, step]. | ||
| 314 | eapply reflexive_eq. f_equal. apply map_eq=> x. rewrite lookup_fmap. | ||
| 315 | destruct (αs !! x) as [[τ e]|] eqn:?; [|done]. | ||
| 316 | assert (τ = NONREC) as -> by eauto using no_recs_lookup. done. | ||
| 317 | Qed. | ||
| 318 | |||
| 319 | Lemma SAttr_lookup_rtc αs x e e' : | ||
| 320 | no_recs αs → | ||
| 321 | αs !! x = None → | ||
| 322 | (∀ y α, αs !! y = Some α → final DEEP (attr_expr α) ∨ attr_le x y) → | ||
| 323 | e -{DEEP}->* e' → | ||
| 324 | EAttr (<[x:=AttrN e]> αs) -{DEEP}->* EAttr (<[x:=AttrN e']> αs). | ||
| 325 | Proof. | ||
| 326 | intros Hrecs Hx Hfirst He. revert αs Hrecs Hx Hfirst. | ||
| 327 | induction He as [e|e1 e2 e3 He12 He23 IH]; intros eτs Hrec Hx Hfirst; [done|]. | ||
| 328 | eapply rtc_l; first by eapply SAttr. apply IH; [done..|]. | ||
| 329 | apply step_not_final in He12. naive_solver. | ||
| 330 | Qed. | ||
| 331 | |||
| 332 | Lemma SAttr_inv αs1 e2 : | ||
| 333 | no_recs αs1 → | ||
| 334 | EAttr αs1 -{DEEP}-> e2 ↔ ∃ αs x e e', | ||
| 335 | αs1 = <[x:=AttrN e]> αs ∧ e2 = EAttr (<[x:=AttrN e']> αs) ∧ | ||
| 336 | αs !! x = None ∧ | ||
| 337 | (∀ y α, αs !! y = Some α → final DEEP (attr_expr α) ∨ attr_le x y) ∧ | ||
| 338 | e -{DEEP}-> e'. | ||
| 339 | Proof. | ||
| 340 | split; [intros; inv_step|]; | ||
| 341 | naive_solver eauto using SAttr, no_recs_insert_inv. | ||
| 342 | Qed. | ||
| 343 | |||
| 344 | Lemma Attr_nf_insert_final αs x e : | ||
| 345 | no_recs αs → | ||
| 346 | αs !! x = None → | ||
| 347 | final DEEP e → | ||
| 348 | (∀ y, is_Some (αs !! y) → attr_le x y) → | ||
| 349 | nf (step DEEP) (EAttr αs) → | ||
| 350 | nf (step DEEP) (EAttr (<[x:=AttrN e]> αs)). | ||
| 351 | Proof. | ||
| 352 | intros Hrecs Hx Hfinal Hleast Hnf | ||
| 353 | [? (αs'&x'&e'&e''&Hαs&->&Hx'&?&Hstep)%SAttr_inv]; | ||
| 354 | last by eauto using no_recs_insert. | ||
| 355 | assert (x ≠ x'). | ||
| 356 | { intros ->. apply (f_equal (.!! x')) in Hαs. rewrite !lookup_insert in Hαs. | ||
| 357 | apply step_not_final in Hstep. naive_solver. } | ||
| 358 | destruct Hnf. exists (EAttr (<[x':=AttrN e'']> (delete x αs'))). | ||
| 359 | rewrite -(delete_insert αs x (AttrN e)) // Hαs delete_insert_ne //. | ||
| 360 | refine (SCtx _ _ _ _ _ (CAttr _ _ _ _ _) _); | ||
| 361 | [|by rewrite lookup_delete_ne| |done]. | ||
| 362 | - apply (no_recs_insert _ x e) in Hrecs. rewrite Hαs in Hrecs. | ||
| 363 | apply no_recs_insert_inv in Hrecs; last done. by apply map_Forall_delete. | ||
| 364 | - intros ?? ?%lookup_delete_Some; naive_solver. | ||
| 365 | Qed. | ||
| 366 | Lemma Attr_nf_insert αs x e : | ||
| 367 | no_recs αs → | ||
| 368 | αs !! x = None → | ||
| 369 | ¬final DEEP e → | ||
| 370 | (∀ y, is_Some (αs !! y) → attr_le x y) → | ||
| 371 | nf (step DEEP) e → | ||
| 372 | nf (step DEEP) (EAttr (<[x:=AttrN e]> αs)). | ||
| 373 | Proof. | ||
| 374 | intros Hrecs Hx ?? Hnf [? (αs'&x'&e'&e''&Hαs&->&Hx'&Hleast'&Hstep)%SAttr_inv]; | ||
| 375 | last eauto using no_recs_insert. | ||
| 376 | assert (x ≠ x') as Hxx'. | ||
| 377 | { intros ->. apply (f_equal (.!! x')) in Hαs. rewrite !lookup_insert in Hαs. | ||
| 378 | naive_solver. } | ||
| 379 | odestruct (Hleast' x (AttrN e)); [|done|]. | ||
| 380 | - apply (f_equal (.!! x)) in Hαs. | ||
| 381 | by rewrite lookup_insert lookup_insert_ne in Hαs. | ||
| 382 | - apply (f_equal (.!! x')) in Hαs. | ||
| 383 | rewrite lookup_insert lookup_insert_ne // in Hαs. | ||
| 384 | destruct Hxx'. apply (anti_symm attr_le); naive_solver. | ||
| 385 | Qed. | ||
| 386 | |||
| 387 | Lemma Attr_step_dom μ αs1 e2 : | ||
| 388 | EAttr αs1 -{μ}-> e2 → | ||
| 389 | ∃ αs2, e2 = EAttr αs2 ∧ ∀ i, αs1 !! i = None ↔ αs2 !! i = None. | ||
| 390 | Proof. | ||
| 391 | intros; inv_step; (eexists; split; [done|]). | ||
| 392 | - intros i. by rewrite lookup_fmap fmap_None. | ||
| 393 | - intros i. rewrite !lookup_insert_None; naive_solver. | ||
| 394 | Qed. | ||
| 395 | Lemma Attr_steps_dom μ αs1 αs2 : | ||
| 396 | EAttr αs1 -{μ}->* EAttr αs2 → ∀ i, αs1 !! i = None ↔ αs2 !! i = None. | ||
| 397 | Proof. | ||
| 398 | intros Hstep. | ||
| 399 | remember (EAttr αs1) as e1 eqn:He1; remember (EAttr αs2) as e2 eqn:He2. | ||
| 400 | revert αs1 αs2 He1 He2. induction Hstep as [|e1 e2 e3 Hstep Hstep' IH]; | ||
| 401 | intros αs1 αs3 ??; simplify_eq/=; [done|]. | ||
| 402 | apply Attr_step_dom in Hstep; naive_solver. | ||
| 403 | Qed. | ||
| 404 | |||
| 405 | Lemma Attr_step_recs αs1 αs2 : | ||
| 406 | EAttr αs1 -{DEEP}-> EAttr αs2 → no_recs αs1 → no_recs αs2. | ||
| 407 | Proof. intros. inv_step; by eauto using no_recs_insert. Qed. | ||
| 408 | Lemma Attr_steps_recs αs1 αs2 : | ||
| 409 | EAttr αs1 -{DEEP}->* EAttr αs2 → no_recs αs1 → no_recs αs2. | ||
| 410 | Proof. | ||
| 411 | intros Hstep. | ||
| 412 | remember (EAttr αs1) as e1 eqn:He1; remember (EAttr αs2) as e2 eqn:He2. | ||
| 413 | revert αs1 αs2 He1 He2. induction Hstep as [|e1 e2 e3 Hstep Hstep' IH]; | ||
| 414 | intros αs1 αs3 ???; simplify_eq/=; [done|]. | ||
| 415 | pose proof (Attr_step_dom _ _ _ Hstep) as (es2 & -> & ?). | ||
| 416 | apply Attr_step_recs in Hstep; naive_solver. | ||
| 417 | Qed. | ||
| 418 | |||
| 419 | Lemma Attr_step_insert αs1 αs2 x e : | ||
| 420 | no_recs αs1 → | ||
| 421 | αs1 !! x = None → | ||
| 422 | final DEEP e → | ||
| 423 | EAttr αs1 -{DEEP}-> EAttr αs2 → | ||
| 424 | EAttr (<[x:=AttrN e]> αs1) -{DEEP}-> EAttr (<[x:=AttrN e]> αs2). | ||
| 425 | Proof. | ||
| 426 | intros Hrecs Hx ? | ||
| 427 | (αs' & x' & e1 & e1' & ? & ? & ? & ? & ?)%SAttr_inv; [|done]; simplify_eq. | ||
| 428 | apply lookup_insert_None in Hx as [??]. rewrite !(insert_commute _ x) //. | ||
| 429 | eapply SAttr; [|by rewrite lookup_insert_ne| |done]. | ||
| 430 | - by eapply no_recs_insert, no_recs_insert_inv. | ||
| 431 | - intros y e' ?%lookup_insert_Some; naive_solver. | ||
| 432 | Qed. | ||
| 433 | Lemma Attr_steps_insert αs1 αs2 x e : | ||
| 434 | no_recs αs1 → | ||
| 435 | αs1 !! x = None → | ||
| 436 | final DEEP e → | ||
| 437 | EAttr αs1 -{DEEP}->* EAttr αs2 → | ||
| 438 | EAttr (<[x:=AttrN e]> αs1) -{DEEP}->* EAttr (<[x:=AttrN e]> αs2). | ||
| 439 | Proof. | ||
| 440 | intros Hrecs Hx ? Hstep. | ||
| 441 | remember (EAttr αs1) as e1 eqn:He1; remember (EAttr αs2) as e2 eqn:He2. | ||
| 442 | revert αs1 αs2 Hx Hrecs He1 He2. | ||
| 443 | induction Hstep as [|e1 e2 e3 Hstep Hstep' IH]; | ||
| 444 | intros αs1 αs3 ????; simplify_eq/=; [done|]. | ||
| 445 | pose proof (Attr_step_dom _ _ _ Hstep) as (αs2 & -> & Hdom). | ||
| 446 | eapply rtc_l; first by eapply Attr_step_insert. | ||
| 447 | eapply IH; naive_solver eauto using Attr_step_recs. | ||
| 448 | Qed. | ||
| 449 | |||
| 450 | Reserved Infix "=D=>" (right associativity, at level 55). | ||
| 451 | |||
| 452 | Inductive step_delayed : relation expr := | ||
| 453 | | RDrefl e : | ||
| 454 | e =D=> e | ||
| 455 | | RDId x e1 e2 : | ||
| 456 | e1 =D=> e2 → | ||
| 457 | EId x (Some (ABS, e1)) =D=> e2 | ||
| 458 | | RDBinOp op e1 e1' e2 e2' : | ||
| 459 | e1 =D=> e1' → e2 =D=> e2' → EBinOp op e1 e2 =D=> EBinOp op e1' e2' | ||
| 460 | | RDIf e1 e1' e2 e2' e3 e3' : | ||
| 461 | e1 =D=> e1' → e2 =D=> e2' → e3 =D=> e3' → EIf e1 e2 e3 =D=> EIf e1' e2' e3' | ||
| 462 | where "e1 =D=> e2" := (step_delayed e1 e2). | ||
| 463 | |||
| 464 | Global Instance step_delayed_po : PreOrder step_delayed. | ||
| 465 | Proof. | ||
| 466 | split; [constructor|]. | ||
| 467 | intros e1 e2 e3 Hstep. revert e3. | ||
| 468 | induction Hstep; inv 1; eauto using step_delayed. | ||
| 469 | Qed. | ||
| 470 | Hint Extern 0 (_ =D=> _) => reflexivity : core. | ||
| 471 | |||
| 472 | Lemma delayed_final_l e1 e2 : | ||
| 473 | final SHALLOW e1 → | ||
| 474 | e1 =D=> e2 → | ||
| 475 | e1 = e2. | ||
| 476 | Proof. intros Hfinal. induction 1; try by inv Hfinal. Qed. | ||
| 477 | |||
| 478 | Lemma delayed_final_r μ e1 e2 : | ||
| 479 | final μ e2 → | ||
| 480 | e1 =D=> e2 → | ||
| 481 | e1 -{μ}->* e2. | ||
| 482 | Proof. | ||
| 483 | intros Hfinal. induction 1; try by inv Hfinal. | ||
| 484 | eapply rtc_l; [constructor|]. eauto. | ||
| 485 | Qed. | ||
| 486 | |||
| 487 | Lemma delayed_step_l μ e1 e1' e2 : | ||
| 488 | e1 =D=> e1' → | ||
| 489 | e1 -{μ}-> e2 → | ||
| 490 | ∃ e2', e1' -{μ}->* e2' ∧ e2 =D=> e2'. | ||
| 491 | Proof. | ||
| 492 | intros Hrem. revert μ e2. | ||
| 493 | induction Hrem; intros μ ? Hstep. | ||
| 494 | - eauto using rtc_once. | ||
| 495 | - inv_step. by exists e2. | ||
| 496 | - inv_step. | ||
| 497 | + eapply delayed_final_l in Hrem1 as ->, Hrem2 as ->; [|by eauto..]. | ||
| 498 | eexists; split; [|done]. eapply rtc_once. by econstructor. | ||
| 499 | + apply IHHrem1 in H2 as (e1'' & ? & ?). | ||
| 500 | eexists; split; [by eapply SBinOpL_rtc|]. by constructor. | ||
| 501 | + eapply delayed_final_l in Hrem1 as ->; [|by eauto..]. | ||
| 502 | apply IHHrem2 in H2 as (e2'' & ? & ?). | ||
| 503 | eexists (EBinOp _ e1' e2''); split; [|by constructor]. | ||
| 504 | by eapply SBinOpR_rtc. | ||
| 505 | - inv_step. | ||
| 506 | + eapply delayed_final_l in Hrem1 as <-; [|by repeat constructor]. | ||
| 507 | eexists; split; [eapply rtc_once; constructor|]. by destruct b. | ||
| 508 | + apply IHHrem1 in H2 as (e1'' & ? & ?). | ||
| 509 | eexists; split; [by eapply SIf_rtc|]. by constructor. | ||
| 510 | Qed. | ||
| 511 | |||
| 512 | Lemma delayed_steps_l μ e1 e1' e2 : | ||
| 513 | e1 =D=> e1' → | ||
| 514 | e1 -{μ}->* e2 → | ||
| 515 | ∃ e2', e1' -{μ}->* e2' ∧ e2 =D=> e2'. | ||
| 516 | Proof. | ||
| 517 | intros Hdel Hsteps. revert e1' Hdel. | ||
| 518 | induction Hsteps as [e|e1 e2 e3 Hstep Hsteps IH]; intros e1' Hdel. | ||
| 519 | { eexists; by split. } | ||
| 520 | eapply delayed_step_l in Hstep as (e2' & Hstep2 & Hdel2); [|done]. | ||
| 521 | apply IH in Hdel2 as (e3' & ? & ?). eexists; by split; [etrans|]. | ||
| 522 | Qed. | ||
| 523 | |||
| 524 | Lemma delayed_step_r μ e1 e1' e2 : | ||
| 525 | e1' =D=> e1 → | ||
| 526 | e1 -{μ}-> e2 → | ||
| 527 | ∃ e2', e1' -{μ}->+ e2' ∧ e2' =D=> e2. | ||
| 528 | Proof. | ||
| 529 | intros Hrem. revert μ e2. | ||
| 530 | induction Hrem; intros μ ? Hstep. | ||
| 531 | - eauto using tc_once. | ||
| 532 | - apply IHHrem in Hstep as (e1' & ? & ?). | ||
| 533 | eexists. split; [|done]. eapply tc_l; [econstructor|done]. | ||
| 534 | - inv_step. | ||
| 535 | + exists e0; split; [|done]. | ||
| 536 | eapply tc_rtc_l; [by eapply SBinOpL_rtc, delayed_final_r, Hrem1|]. | ||
| 537 | eapply tc_rtc_l; [by eapply SBinOpR_rtc, delayed_final_r, Hrem2|]. | ||
| 538 | eapply tc_once. by econstructor. | ||
| 539 | + apply IHHrem1 in H2 as (e1'' & ? & ?). | ||
| 540 | eexists; split; [by eapply SBinOpL_tc|]. by constructor. | ||
| 541 | + apply IHHrem2 in H2 as (e2'' & ? & ?). | ||
| 542 | eexists (EBinOp _ e1' e2''); split; [|by apply RDBinOp]. | ||
| 543 | eapply tc_rtc_l; [by eapply SBinOpL_rtc, delayed_final_r, Hrem1|]. | ||
| 544 | by eapply SBinOpR_tc. | ||
| 545 | - inv_step. | ||
| 546 | + exists (if b then e2 else e3). split; [|by destruct b]. | ||
| 547 | eapply tc_rtc_l; | ||
| 548 | [eapply SIf_rtc, delayed_final_r, Hrem1; by repeat constructor|]. | ||
| 549 | eapply tc_once; constructor. | ||
| 550 | + apply IHHrem1 in H2 as (e1'' & ? & ?). | ||
| 551 | eexists; split; [by eapply SIf_tc|]. by constructor. | ||
| 552 | Qed. | ||
| 553 | |||
| 554 | Lemma delayed_steps_r μ e1 e1' e2 : | ||
| 555 | e1' =D=> e1 → | ||
| 556 | e1 -{μ}->* e2 → | ||
| 557 | ∃ e2', e1' -{μ}->* e2' ∧ e2' =D=> e2. | ||
| 558 | Proof. | ||
| 559 | intros Hdel Hsteps. revert e1' Hdel. | ||
| 560 | induction Hsteps as [e|e1 e2 e3 Hstep Hsteps IH]; intros e1' Hdel. | ||
| 561 | { eexists; by split. } | ||
| 562 | eapply delayed_step_r in Hstep as (e2' & Hstep2%tc_rtc & Hdel2); [|done]. | ||
| 563 | apply IH in Hdel2 as (e3' & ? & ?). eexists; by split; [etrans|]. | ||
| 564 | Qed. | ||
| 565 | |||
| 566 | (** Determinism *) | ||
| 567 | |||
| 568 | Lemma bin_op_det op e Φ Ψ : | ||
| 569 | sem_bin_op op e Φ → | ||
| 570 | sem_bin_op op e Ψ → | ||
| 571 | Φ = Ψ. | ||
| 572 | Proof. by destruct 1; inv 1. Qed. | ||
| 573 | |||
| 574 | Lemma bin_op_rel_det op e1 Φ e2 d1 d2 : | ||
| 575 | sem_bin_op op e1 Φ → | ||
| 576 | Φ e2 d1 → | ||
| 577 | Φ e2 d2 → | ||
| 578 | d1 = d2. | ||
| 579 | Proof. | ||
| 580 | assert (AntiSymm eq attr_le) by apply _. | ||
| 581 | unfold AntiSymm in *. inv 1; repeat case_match; naive_solver. | ||
| 582 | Qed. | ||
| 583 | |||
| 584 | Lemma matches_present x e md es ms strict βs : | ||
| 585 | es !! x = Some e → ms !! x = Some md → | ||
| 586 | matches es ms strict βs → | ||
| 587 | βs !! x = Some (AttrN e). | ||
| 588 | Proof. | ||
| 589 | intros Hes Hms. induction 1; try done. | ||
| 590 | - by apply lookup_insert_Some in Hes as [[]|[]]; simplify_map_eq. | ||
| 591 | - by simplify_map_eq. | ||
| 592 | Qed. | ||
| 593 | |||
| 594 | Lemma matches_default x es ms d strict βs : | ||
| 595 | es !! x = None → | ||
| 596 | ms !! x = Some (Some d) → | ||
| 597 | matches es ms strict βs → | ||
| 598 | βs !! x = Some (AttrR d). | ||
| 599 | Proof. | ||
| 600 | intros Hes Hms. induction 1; try done. | ||
| 601 | - by apply lookup_insert_None in Hes as []; simplify_map_eq. | ||
| 602 | - by apply lookup_insert_Some in Hms as [[]|[]]; simplify_map_eq. | ||
| 603 | Qed. | ||
| 604 | |||
| 605 | Lemma matches_weaken x es ms strict βs : | ||
| 606 | matches es ms strict βs → | ||
| 607 | matches (delete x es) (delete x ms) strict (delete x βs). | ||
| 608 | Proof. | ||
| 609 | induction 1; [constructor|constructor|..]; rename x0 into y; | ||
| 610 | (destruct (decide (x = y)) as [->|Hxy]; | ||
| 611 | [ rewrite !delete_insert_delete // | ||
| 612 | | rewrite !delete_insert_ne //; constructor; | ||
| 613 | by simplify_map_eq ]). | ||
| 614 | Qed. | ||
| 615 | |||
| 616 | Lemma matches_det es ms strict βs1 βs2 : | ||
| 617 | matches es ms strict βs1 → | ||
| 618 | matches es ms strict βs2 → | ||
| 619 | βs1 = βs2. | ||
| 620 | Proof. | ||
| 621 | intros Hβs1. revert βs2. induction Hβs1; intros βs2 Hβs2; | ||
| 622 | try (inv Hβs2; done || (by exfalso; eapply (insert_non_empty (M:=stringmap)))). | ||
| 623 | - eapply (matches_weaken x) in Hβs2 as Hβs2'. | ||
| 624 | rewrite !delete_insert // in Hβs2'. | ||
| 625 | rewrite (IHHβs1 _ Hβs2') insert_delete //. | ||
| 626 | eapply matches_present; eauto; apply lookup_insert. | ||
| 627 | - eapply (matches_weaken x) in Hβs2 as Hβs2'. | ||
| 628 | rewrite delete_notin // delete_insert // in Hβs2'. | ||
| 629 | rewrite (IHHβs1 _ Hβs2') insert_delete //. | ||
| 630 | eapply matches_default; eauto. apply lookup_insert. | ||
| 631 | Qed. | ||
| 632 | |||
| 633 | Lemma ctx_det K1 K2 e1 e2 μ μ1' μ2' : | ||
| 634 | K1 e1 = K2 e2 → | ||
| 635 | ctx1 μ1' μ K1 → | ||
| 636 | ctx1 μ2' μ K2 → | ||
| 637 | red (step μ1') e1 → | ||
| 638 | red (step μ2') e2 → | ||
| 639 | K1 = K2 ∧ e1 = e2 ∧ μ1' = μ2'. | ||
| 640 | Proof. | ||
| 641 | intros Hes HK1 HK2 Hred1 Hred2. | ||
| 642 | induction HK1; inv HK2; try done. | ||
| 643 | - apply not_elem_of_app_cons_inv_l in Hes as [<- [<- <-]]; first done. | ||
| 644 | + intros He1. apply (proj1 (Forall_forall _ _) H0) in He1. | ||
| 645 | inv Hred1. by apply step_not_final in H1. | ||
| 646 | + intros He2. apply (proj1 (Forall_forall _ _) H) in He2. | ||
| 647 | inv Hred2. by apply step_not_final in H1. | ||
| 648 | - destruct (decide (x = x0)) as [<-|]. | ||
| 649 | { by apply map_insert_inv_eq in Hes as [[= ->] [= ->]]. } | ||
| 650 | apply map_insert_inv_ne in Hes as (Hx0 & Hx & Hαs); try done. | ||
| 651 | apply H1 in Hx0 as [contra | Hxlex0]. | ||
| 652 | + inv Hred2. by apply step_not_final in H5. | ||
| 653 | + apply H4 in Hx as [contra | Hx0lex]. | ||
| 654 | * inv Hred1. by apply step_not_final in H5. | ||
| 655 | * assert (Hasym : AntiSymm eq attr_le) by apply _. | ||
| 656 | by pose proof (Hasym _ _ Hxlex0 Hx0lex). | ||
| 657 | - inv Hred1. inv_step. | ||
| 658 | - inv Hred2. inv_step. | ||
| 659 | - inv Hred1. by apply step_not_final in H1. | ||
| 660 | - inv Hred2. by apply step_not_final in H1. | ||
| 661 | Qed. | ||
| 662 | |||
| 663 | Lemma step_det μ e d1 d2 : | ||
| 664 | e -{μ}-> d1 → | ||
| 665 | e -{μ}-> d2 → | ||
| 666 | d1 = d2. | ||
| 667 | Proof. | ||
| 668 | intros Hred1. revert d2. | ||
| 669 | induction Hred1; intros d2 Hred2; inv Hred2; try by inv_step. | ||
| 670 | - by apply (matches_det _ _ _ _ _ H0) in H8 as <-. | ||
| 671 | - inv_step. by apply step_not_final in H3. | ||
| 672 | - inv_step. destruct H. by apply no_recs_insert. | ||
| 673 | - assert (Φ = Φ0) as <- by (by eapply bin_op_det). | ||
| 674 | by eapply bin_op_rel_det. | ||
| 675 | - inv_step; by apply step_not_final in H6. | ||
| 676 | - inv_step. by apply step_not_final in Hred1. | ||
| 677 | - inv_step. destruct H2. by apply no_recs_insert. | ||
| 678 | - inv_step; by apply step_not_final in Hred1. | ||
| 679 | - eapply ctx_det in H0 as (?&?&?); [|by eauto..]; naive_solver. | ||
| 680 | Qed. | ||
diff --git a/theories/nix/tests.v b/theories/nix/tests.v new file mode 100644 index 0000000..cbce874 --- /dev/null +++ b/theories/nix/tests.v | |||
| @@ -0,0 +1,185 @@ | |||
| 1 | From mininix Require Export nix.interp nix.notations. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | Open Scope Z_scope. | ||
| 4 | |||
| 5 | (** Compare base vals without comparing the proofs. Since we do not have | ||
| 6 | definitional proof irrelevance, comparing the proofs would fail (and in practice | ||
| 7 | make Coq loop). *) | ||
| 8 | Definition res_eq (rv : res val) (bl2 : base_lit) := | ||
| 9 | match rv with | ||
| 10 | | Res (Some (VLit bl1 _)) => bl1 = bl2 | ||
| 11 | | _ => False | ||
| 12 | end. | ||
| 13 | Infix "=?" := res_eq. | ||
| 14 | |||
| 15 | Definition float_1 := | ||
| 16 | ceil: (Float.of_Z 20 /: 3). | ||
| 17 | Goal interp 100 ∅ float_1 =? 7. | ||
| 18 | Proof. by vm_compute. Qed. | ||
| 19 | |||
| 20 | Definition float_2 := | ||
| 21 | Float.of_Z 100000000000000000000000000000000000000000000000000000000000000000000000000000 *: | ||
| 22 | Float.of_Z 100000000000000000000000000000000000000000000000000000000000000000000000000000 *: | ||
| 23 | Float.of_Z 100000000000000000000000000000000000000000000000000000000000000000000000000000 *: | ||
| 24 | Float.of_Z 100000000000000000000000000000000000000000000000000000000000000000000000000000 *: | ||
| 25 | Float.of_Z 100000000000000000000000000000000000000000000000000000000000000000000000000000. | ||
| 26 | Goal interp 100 ∅ float_2 =? NFloat (Binary.B754_infinity false). | ||
| 27 | Proof. by vm_compute. Qed. | ||
| 28 | |||
| 29 | Definition float_3 := float_2 /: float_2. | ||
| 30 | Goal interp 100 ∅ float_3 =? NFloat (`Float.indef_nan). | ||
| 31 | Proof. by vm_compute. Qed. | ||
| 32 | |||
| 33 | Definition let_let := | ||
| 34 | let: "x" := 1 in let: "x" := 2 in "x". | ||
| 35 | Goal interp 100 ∅ let_let =? 2. | ||
| 36 | Proof. by vm_compute. Qed. | ||
| 37 | |||
| 38 | Definition with_let := | ||
| 39 | with: EAttr {[ "x" := AttrN 1 ]} in let: "x" := 2 in "x". | ||
| 40 | Goal interp 100 ∅ with_let =? 2. | ||
| 41 | Proof. by vm_compute. Qed. | ||
| 42 | |||
| 43 | Definition let_with := | ||
| 44 | let: "x" := 1 in with: EAttr {[ "x" := AttrN 2 ]} in "x". | ||
| 45 | Goal interp 100 ∅ let_with =? 1. | ||
| 46 | Proof. by vm_compute. Qed. | ||
| 47 | |||
| 48 | Definition with_with := | ||
| 49 | with: EAttr {[ "x" := AttrN 1 ]} in with: EAttr {[ "x" := AttrN 2 ]} in "x". | ||
| 50 | Goal interp 100 ∅ with_with =? 2. | ||
| 51 | Proof. by vm_compute. Qed. | ||
| 52 | |||
| 53 | Definition with_with_inherit := | ||
| 54 | with: EAttr {[ "x" := AttrN 1 ]} in with: EAttr {[ "x" := AttrN "x" ]} in "x". | ||
| 55 | Goal interp 100 ∅ with_with_inherit =? 1. | ||
| 56 | Proof. by vm_compute. Qed. | ||
| 57 | |||
| 58 | Definition with_loop := | ||
| 59 | with: EAttr {[ "x" := AttrR "x" ]} in "x". | ||
| 60 | Goal interp 100 ∅ with_loop = NoFuel. | ||
| 61 | Proof. by vm_compute. Qed. | ||
| 62 | |||
| 63 | Definition rec_attr_shadow_1 := | ||
| 64 | let: "foo" := EAttr {[ "bar" := AttrN 10 ]} in | ||
| 65 | EAttr {[ | ||
| 66 | "bar" := AttrR ("foo" .: "bar"); | ||
| 67 | "foo" := AttrR (EAttr {[ "bar" := AttrN 20 ]}) | ||
| 68 | ]} .: "bar". | ||
| 69 | Goal interp 100 ∅ rec_attr_shadow_1 =? 20. | ||
| 70 | Proof. by vm_compute. Qed. | ||
| 71 | |||
| 72 | Definition rec_attr_shadow_2 := | ||
| 73 | EAttr {[ | ||
| 74 | "y" := AttrR (EAttr {[ "y" := AttrN "z" ]} .: "y"); | ||
| 75 | "z" := AttrR 20 | ||
| 76 | ]} .: "y". | ||
| 77 | Goal interp 100 ∅ rec_attr_shadow_2 =? 20. | ||
| 78 | Proof. by vm_compute. Qed. | ||
| 79 | |||
| 80 | Definition nested_functor_1 := | ||
| 81 | EAttr {[ "__functor" := AttrN $ λ: "self", | ||
| 82 | EAttr {[ "__functor" := AttrN $ λ: "self" "x", 10 ]} ]} 10. | ||
| 83 | Goal interp 100 ∅ nested_functor_1 =? 10. | ||
| 84 | Proof. by vm_compute. Qed. | ||
| 85 | |||
| 86 | Definition nested_functor_2 := | ||
| 87 | EAttr {[ "__functor" := AttrN $ | ||
| 88 | EAttr {[ "__functor" := AttrN $ λ: "self" "self" "x", 10 ]} ]} 10. | ||
| 89 | Goal interp 100 ∅ nested_functor_2 =? 10. | ||
| 90 | Proof. by vm_compute. Qed. | ||
| 91 | |||
| 92 | Definition functor_loop_1 := | ||
| 93 | EAttr {[ "__functor" := AttrN $ | ||
| 94 | λ: "self", "self" "self" | ||
| 95 | ]} 10. | ||
| 96 | Goal interp 1000 ∅ functor_loop_1 = NoFuel. | ||
| 97 | Proof. by vm_compute. Qed. | ||
| 98 | |||
| 99 | Definition functor_loop_2 := | ||
| 100 | EAttr {[ "__functor" := AttrN $ | ||
| 101 | λ: "self" "f", "f" ("self" "f") | ||
| 102 | ]} (λ: "go" "x", "go" "x") 10. | ||
| 103 | Goal interp 1000 ∅ functor_loop_2 = NoFuel. | ||
| 104 | Proof. by vm_compute. Qed. | ||
| 105 | |||
| 106 | Fixpoint many_lets (i : nat) (e : expr) : expr := | ||
| 107 | match i with | ||
| 108 | | O => e | ||
| 109 | | S i => let: "x" +:+ pretty i := 0 in many_lets i e | ||
| 110 | end. | ||
| 111 | |||
| 112 | Fixpoint many_adds (i : nat) : expr := | ||
| 113 | match i with | ||
| 114 | | O => 0 | ||
| 115 | | S i => ("x" +:+ pretty i) +: many_adds i | ||
| 116 | end. | ||
| 117 | |||
| 118 | Definition big_prog (i : nat) : expr := many_lets i $ many_adds i. | ||
| 119 | |||
| 120 | Definition big := big_prog 1000. | ||
| 121 | |||
| 122 | Goal interp 5000 ∅ big =? 0. | ||
| 123 | Proof. by vm_compute. Qed. | ||
| 124 | |||
| 125 | Definition matching_1 := | ||
| 126 | (λattr: {[ "x" := None; "y" := None ]}, "x" +: "y") | ||
| 127 | (EAttr {[ "x" := AttrN 10; "y" := AttrN 11 ]}). | ||
| 128 | Goal interp 1000 ∅ matching_1 =? 21. | ||
| 129 | Proof. by vm_compute. Qed. | ||
| 130 | |||
| 131 | Definition matching_2 := | ||
| 132 | (λattr: {[ "x" := None; "y" := Some (EId' "x") ]}, "x" +: "y") | ||
| 133 | (EAttr {[ "x" := AttrN 10 ]}). | ||
| 134 | Goal interp 1000 ∅ matching_2 =? 20. | ||
| 135 | Proof. by vm_compute. Qed. | ||
| 136 | |||
| 137 | Definition matching_3 := | ||
| 138 | (λattr: {[ "x" := None; "y" := None ]}, "x" +: "y") | ||
| 139 | (EAttr {[ "x" := AttrN 10 ]}). | ||
| 140 | Goal interp 1000 ∅ matching_3 = mfail. | ||
| 141 | Proof. by vm_compute. Qed. | ||
| 142 | |||
| 143 | Definition matching_4 := | ||
| 144 | (λattr: {[ "x" := None; "y" := None ]}, "x" +: "y") | ||
| 145 | (EAttr {[ "x" := AttrN 10; "y" := AttrN 11; "z" := AttrN 12 ]}). | ||
| 146 | Goal interp 1000 ∅ matching_4 = mfail. | ||
| 147 | Proof. by vm_compute. Qed. | ||
| 148 | |||
| 149 | Definition matching_5 := | ||
| 150 | (λattr: {[ "x" := None; "y" := None ]} .., "x" +: "y") | ||
| 151 | (EAttr {[ "x" := AttrN 10; "y" := AttrN 11; "z" := AttrN 12 ]}). | ||
| 152 | Goal interp 1000 ∅ matching_5 =? 21. | ||
| 153 | Proof. by vm_compute. Qed. | ||
| 154 | |||
| 155 | Definition matching_6 := | ||
| 156 | (λattr: {[ "y" := Some (EId' "y") ]}, "y") | ||
| 157 | (EAttr {[ "y" := AttrN 10 ]}). | ||
| 158 | Goal interp 1000 ∅ matching_6 =? 10. | ||
| 159 | Proof. by vm_compute. Qed. | ||
| 160 | |||
| 161 | Definition matching_7 := | ||
| 162 | (λattr: {[ "y" := Some (EId' "y") ]}, "y") (EAttr ∅). | ||
| 163 | Goal interp 1000 ∅ matching_7 = NoFuel. | ||
| 164 | Proof. by vm_compute. Qed. | ||
| 165 | |||
| 166 | Definition matching_8 := | ||
| 167 | (λattr: {[ "y" := Some (EId' "y") ]}.., "y") | ||
| 168 | (EAttr {[ "x" := AttrN 10 ]}). | ||
| 169 | Goal interp 1000 ∅ matching_8 = NoFuel. | ||
| 170 | Proof. by vm_compute. Qed. | ||
| 171 | |||
| 172 | Definition list_lt_1 := | ||
| 173 | EList [ELit 2; ELit 3] <: EList [ELit 3]. | ||
| 174 | Goal interp 1000 ∅ list_lt_1 =? true. | ||
| 175 | Proof. by vm_compute. Qed. | ||
| 176 | |||
| 177 | Definition list_lt_2 := | ||
| 178 | EList [ELit 2; ELit 3] <: EList [ELit 2]. | ||
| 179 | Goal interp 1000 ∅ list_lt_2 =? false. | ||
| 180 | Proof. by vm_compute. Qed. | ||
| 181 | |||
| 182 | Definition list_lt_3 := | ||
| 183 | EList [ELit 2] <: EList [ELit 2; ELit 3]. | ||
| 184 | Goal interp 1000 ∅ list_lt_3 =? true. | ||
| 185 | Proof. by vm_compute. Qed. | ||
diff --git a/theories/nix/wp.v b/theories/nix/wp.v new file mode 100644 index 0000000..0eca6e1 --- /dev/null +++ b/theories/nix/wp.v | |||
| @@ -0,0 +1,143 @@ | |||
| 1 | From mininix Require Export nix.operational_props. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | Definition wp (μ : mode) (e : expr) (Φ : expr → Prop) : Prop := | ||
| 5 | ∃ e', e -{μ}->* e' ∧ final μ e' ∧ Φ e'. | ||
| 6 | |||
| 7 | Lemma Lit_wp μ Φ bl : | ||
| 8 | base_lit_ok bl → | ||
| 9 | Φ (ELit bl) → | ||
| 10 | wp μ (ELit bl) Φ. | ||
| 11 | Proof. exists (ELit bl). by repeat constructor. Qed. | ||
| 12 | |||
| 13 | Lemma Abs_wp μ Φ x e : | ||
| 14 | Φ (EAbs x e) → | ||
| 15 | wp μ (EAbs x e) Φ. | ||
| 16 | Proof. exists (EAbs x e). by repeat constructor. Qed. | ||
| 17 | |||
| 18 | Lemma AbsMatch_wp μ Φ ms strict e : | ||
| 19 | Φ (EAbsMatch ms strict e) → | ||
| 20 | wp μ (EAbsMatch ms strict e) Φ. | ||
| 21 | Proof. exists (EAbsMatch ms strict e). by repeat constructor. Qed. | ||
| 22 | |||
| 23 | Lemma LetAttr_no_recs_wp μ Φ k αs e : | ||
| 24 | no_recs αs → | ||
| 25 | wp μ (subst ((k,.) ∘ attr_expr <$> αs) e) Φ → | ||
| 26 | wp μ (ELetAttr k (EAttr αs) e) Φ. | ||
| 27 | Proof. | ||
| 28 | intros Hαs (e' & Hsteps & ? & HΦ). exists e'. split; [|done]. | ||
| 29 | etrans; [|apply Hsteps]. apply rtc_once. by constructor. | ||
| 30 | Qed. | ||
| 31 | |||
| 32 | Lemma BinOp_wp μ Φ op e1 e2 : | ||
| 33 | wp SHALLOW e1 (λ e1', ∃ Φop, | ||
| 34 | sem_bin_op op e1' Φop ∧ | ||
| 35 | wp SHALLOW e2 (λ e2', ∃ e', Φop e2' e' ∧ wp μ e' Φ)) → | ||
| 36 | wp μ (EBinOp op e1 e2) Φ. | ||
| 37 | Proof. | ||
| 38 | intros (e1' & Hsteps1 & ? & Φop & Hop1 & e2' & Hsteps2 & ? | ||
| 39 | & e' & Hop2 & e'' & Hsteps & ? & HΦ). | ||
| 40 | exists e''. split; [|done]. | ||
| 41 | etrans; [by apply SBinOpL_rtc|]. | ||
| 42 | etrans; [by eapply SBinOpR_rtc|]. | ||
| 43 | eapply rtc_l; [by econstructor|]. done. | ||
| 44 | Qed. | ||
| 45 | |||
| 46 | Lemma Id_wp μ Φ x k e : | ||
| 47 | wp μ e Φ → | ||
| 48 | wp μ (EId x (Some (k,e))) Φ. | ||
| 49 | Proof. | ||
| 50 | intros (e' & Hsteps & ? & HΦ). exists e'. split; [|done]. | ||
| 51 | etrans; [|apply Hsteps]. apply rtc_once. constructor. | ||
| 52 | Qed. | ||
| 53 | |||
| 54 | Lemma App_wp μ Φ e1 e2 : | ||
| 55 | wp SHALLOW e1 (λ e1', wp μ (EApp e1' e2) Φ) ↔ | ||
| 56 | wp μ (EApp e1 e2) Φ. | ||
| 57 | Proof. | ||
| 58 | split. | ||
| 59 | - intros (e1' & Hsteps1 & ? & e' & Hsteps2 & ? & HΦ). | ||
| 60 | exists e'; split; [|done]. etrans; [|apply Hsteps2]. | ||
| 61 | by apply SAppL_rtc. | ||
| 62 | - intros (e' & Hsteps & Hfinal & HΦ). | ||
| 63 | cut (∃ e1', e1 -{SHALLOW}->* e1' ∧ final SHALLOW e1' ∧ EApp e1' e2 -{μ}->* e'). | ||
| 64 | { intros (e1'&?&?&?). exists e1'. split_and!; [done..|]. by exists e'. } | ||
| 65 | clear Φ HΦ. apply rtc_nsteps in Hsteps as [n Hsteps]. | ||
| 66 | revert e1 Hsteps. induction n as [|n IH]; intros e1 Hsteps. | ||
| 67 | { inv Hsteps. inv Hfinal. } | ||
| 68 | inv Hsteps. inv H0. | ||
| 69 | + eexists; split_and!; [done|by constructor|]. | ||
| 70 | eapply rtc_l; [by constructor|by eapply rtc_nsteps_2]. | ||
| 71 | + eexists; split_and!; [done|by constructor|]. | ||
| 72 | eapply rtc_l; [by constructor|by eapply rtc_nsteps_2]. | ||
| 73 | + eexists; split_and!; [done|by constructor|]. | ||
| 74 | eapply rtc_l; [by constructor|by eapply rtc_nsteps_2]. | ||
| 75 | + inv H2. | ||
| 76 | * apply IH in H1 as (e'' & Hsteps & ? & ?). | ||
| 77 | exists e''; split; [|done]. by eapply rtc_l. | ||
| 78 | * eexists; split_and!; [done|by constructor|]. | ||
| 79 | eapply rtc_l; [by eapply SAppR|]. by eapply rtc_nsteps_2. | ||
| 80 | Qed. | ||
| 81 | |||
| 82 | Lemma Attr_wp_shallow Φ αs : | ||
| 83 | Φ (EAttr (AttrN ∘ from_attr (subst (indirects αs)) id <$> αs)) → | ||
| 84 | wp SHALLOW (EAttr αs) Φ. | ||
| 85 | Proof. | ||
| 86 | eexists (EAttr (AttrN ∘ _ <$> αs)); split_and!; [ |by constructor|done]. | ||
| 87 | destruct (decide (no_recs αs)); [|apply rtc_once; by constructor]. | ||
| 88 | apply reflexive_eq; f_equal. apply map_eq=> x. rewrite lookup_fmap. | ||
| 89 | destruct (αs !! x) as [[? e]|] eqn:?; f_equal/=. | ||
| 90 | by assert (τ = NONREC) as -> by eauto using no_recs_lookup. | ||
| 91 | Qed. | ||
| 92 | |||
| 93 | Lemma β_wp μ Φ x e1 e2 : | ||
| 94 | wp μ (subst {[x:=(ABS, e2)]} e1) Φ → | ||
| 95 | wp μ (EApp (EAbs x e1) e2) Φ. | ||
| 96 | Proof. | ||
| 97 | intros (e' & Hsteps & ? & ?). exists e'. split; [|done]. | ||
| 98 | eapply rtc_l; [|done]. by constructor. | ||
| 99 | Qed. | ||
| 100 | |||
| 101 | Lemma βMatch_wp μ Φ ms strict e1 αs βs : | ||
| 102 | no_recs αs → | ||
| 103 | matches (attr_expr <$> αs) ms strict βs → | ||
| 104 | wp μ (subst (indirects βs) e1) Φ → | ||
| 105 | wp μ (EApp (EAbsMatch ms strict e1) (EAttr αs)) Φ. | ||
| 106 | Proof. | ||
| 107 | intros ?? (e' & Hsteps & ? & ?). exists e'. split; [|done]. | ||
| 108 | eapply rtc_l; [|done]. by constructor. | ||
| 109 | Qed. | ||
| 110 | |||
| 111 | Lemma Functor_wp μ Φ αs e1 e2 : | ||
| 112 | no_recs αs → | ||
| 113 | αs !! "__functor" = Some (AttrN e1) → | ||
| 114 | wp μ (EApp (EApp e1 (EAttr αs)) e2) Φ → | ||
| 115 | wp μ (EApp (EAttr αs) e2) Φ. | ||
| 116 | Proof. | ||
| 117 | intros ?? (e' & Hsteps & ? & ?). exists e'. split; [|done]. | ||
| 118 | eapply rtc_l; [|done]. by constructor. | ||
| 119 | Qed. | ||
| 120 | |||
| 121 | Lemma If_wp μ Φ e1 e2 e3 : | ||
| 122 | wp SHALLOW e1 (λ e1', ∃ b : bool, | ||
| 123 | e1' = ELit (LitBool b) ∧ wp μ (if b then e2 else e3) Φ) → | ||
| 124 | wp μ (EIf e1 e2 e3) Φ. | ||
| 125 | Proof. | ||
| 126 | intros (e1' & Hsteps & ? & b & -> & e' & Hsteps' & ? & HΦ). | ||
| 127 | exists e'; split; [|done]. etrans; [by apply SIf_rtc|]. | ||
| 128 | eapply rtc_l; [|done]. destruct b; constructor. | ||
| 129 | Qed. | ||
| 130 | |||
| 131 | Lemma wp_mono μ e Φ Ψ : | ||
| 132 | wp μ e Φ → | ||
| 133 | (∀ e', Φ e' → Ψ e') → | ||
| 134 | wp μ e Ψ. | ||
| 135 | Proof. intros (e' & ? & ? & ?) ?. exists e'. naive_solver. Qed. | ||
| 136 | |||
| 137 | Lemma union_kinded_abs {A} mkv (v2 : A) : | ||
| 138 | union_kinded (pair WITH <$> mkv) (Some (ABS, v2)) = Some (ABS, v2). | ||
| 139 | Proof. by destruct mkv. Qed. | ||
| 140 | |||
| 141 | Lemma union_kinded_with {A} (v : A) mkv2 : | ||
| 142 | union_kinded (Some (WITH, v)) (pair WITH <$> mkv2) = Some (WITH, v). | ||
| 143 | Proof. by destruct mkv2. Qed. | ||
diff --git a/theories/nix/wp_examples.v b/theories/nix/wp_examples.v new file mode 100644 index 0000000..7bc2109 --- /dev/null +++ b/theories/nix/wp_examples.v | |||
| @@ -0,0 +1,164 @@ | |||
| 1 | From mininix Require Import nix.wp nix.notations. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | Local Open Scope Z_scope. | ||
| 4 | |||
| 5 | Definition test αs := | ||
| 6 | let: "x" := 1 in | ||
| 7 | with: EAttr αs in | ||
| 8 | with: EAttr {[ "y" := AttrN 2 ]} in | ||
| 9 | "x" =: "y". | ||
| 10 | |||
| 11 | Example test_wp μ αs : | ||
| 12 | no_recs αs → | ||
| 13 | wp μ (test αs) (.= false). | ||
| 14 | Proof. | ||
| 15 | intros Hαs. rewrite /test. apply LetAttr_no_recs_wp. | ||
| 16 | { by apply no_recs_insert. } | ||
| 17 | rewrite /= !map_fmap_singleton /= right_id_L lookup_singleton lookup_singleton_ne //=. | ||
| 18 | apply LetAttr_no_recs_wp. | ||
| 19 | { by apply no_recs_attr_subst. } | ||
| 20 | rewrite /= !map_fmap_singleton /= right_id_L. | ||
| 21 | rewrite (map_fmap_compose attr_expr) lookup_fmap union_kinded_abs. | ||
| 22 | rewrite !lookup_fmap. | ||
| 23 | apply LetAttr_no_recs_wp. | ||
| 24 | { by apply no_recs_insert. } | ||
| 25 | rewrite /= map_fmap_singleton lookup_singleton lookup_singleton_ne //=. | ||
| 26 | rewrite union_kinded_with. | ||
| 27 | apply BinOp_wp. | ||
| 28 | apply Id_wp, Lit_wp; first done. eexists; split; [constructor|]. | ||
| 29 | apply Id_wp, Lit_wp; first done. | ||
| 30 | eexists; split; [done|]. by apply Lit_wp. | ||
| 31 | Qed. | ||
| 32 | |||
| 33 | Definition neg := λ: "b", if: "b" then false else true. | ||
| 34 | |||
| 35 | Lemma neg_wp μ (Φ : expr → Prop) e : | ||
| 36 | wp SHALLOW e (λ e', ∃ b : bool, e' = b ∧ Φ (negb b)) → | ||
| 37 | wp μ (neg e) Φ. | ||
| 38 | Proof. | ||
| 39 | intros Hwp. apply β_wp. rewrite /= lookup_singleton /=. | ||
| 40 | apply If_wp, Id_wp. eapply wp_mono; [done|]. | ||
| 41 | intros ? (b & -> & ?). exists b; split; [done|]. | ||
| 42 | destruct b; by apply Lit_wp. | ||
| 43 | Qed. | ||
| 44 | |||
| 45 | (* rec { f = x: if x = 0 then true else !(f (x - 1)); }.f n *) | ||
| 46 | Definition even_rec_attr := | ||
| 47 | EAttr {[ "f" := AttrR (λ: "x", if: "x" =: 0 then true else neg ("f" ("x" -: 1))) ]} .: "f". | ||
| 48 | |||
| 49 | Lemma even_rec_attr_wp e n : | ||
| 50 | 0 ≤ n ≤ int_max → | ||
| 51 | wp SHALLOW e (.= n) → | ||
| 52 | wp SHALLOW (even_rec_attr e) (.= Z.even n). | ||
| 53 | Proof. | ||
| 54 | intros Hn Hwp. apply App_wp. | ||
| 55 | revert e Hwp. induction (Z.lt_wf 0 n) as [n _ IH]; intros e Hwp. | ||
| 56 | apply BinOp_wp. apply Attr_wp_shallow. | ||
| 57 | eexists; split; [by constructor|]. | ||
| 58 | apply Lit_wp; [done|]. eexists; split; [by eexists|]. | ||
| 59 | rewrite /=. apply Abs_wp, β_wp. | ||
| 60 | rewrite /= !lookup_singleton /= !lookup_singleton_ne //=. | ||
| 61 | rewrite !union_with_None_l !union_with_None_r. | ||
| 62 | rewrite /indirects map_imap_insert map_imap_empty lookup_insert. | ||
| 63 | rewrite -/even_rec_attr -/neg. | ||
| 64 | apply If_wp, BinOp_wp, Id_wp. | ||
| 65 | eapply wp_mono; [apply Hwp|]; intros ? ->. | ||
| 66 | eexists; split; [by constructor|]. | ||
| 67 | apply Lit_wp; [done|]. eexists; split; [by eexists|]. simpl. | ||
| 68 | destruct (n =? 0) eqn:Hn0; (apply Lit_wp; [done|]; eexists; split; [done|]; simpl). | ||
| 69 | { apply Lit_wp; [done|]. by apply Z.eqb_eq in Hn0 as ->. } | ||
| 70 | apply neg_wp, App_wp, Id_wp. | ||
| 71 | eapply wp_mono; [apply (IH (n-1))|]; [lia..| |]. | ||
| 72 | 2:{ intros e' He'. eapply wp_mono; [apply He'|]. | ||
| 73 | intros ? ->. eexists; split; [done|]. | ||
| 74 | by rewrite Z.negb_even Z.sub_1_r Z.odd_pred. } | ||
| 75 | eapply BinOp_wp, Id_wp. eapply wp_mono; [apply Hwp|]. intros ? ->. | ||
| 76 | eexists; split; [by constructor|]. apply Lit_wp; [done|]. | ||
| 77 | eexists; split; [eexists _, _; split_and!; [done| |done]|]. | ||
| 78 | - rewrite /= option_guard_True //. apply bool_decide_pack. | ||
| 79 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 80 | - apply Lit_wp; [|done]. apply bool_decide_pack. | ||
| 81 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 82 | Qed. | ||
| 83 | |||
| 84 | Lemma even_rec_attr_wp' n : | ||
| 85 | 0 ≤ n ≤ int_max → | ||
| 86 | wp SHALLOW (even_rec_attr n) (.= Z.even n). | ||
| 87 | Proof. | ||
| 88 | intros ?. apply even_rec_attr_wp; [done|]. apply Lit_wp; [|done]. | ||
| 89 | rewrite /= /int_ok. apply bool_decide_pack. | ||
| 90 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 91 | Qed. | ||
| 92 | |||
| 93 | (* { "__functor " = r: x: if x == 0 then true else !(r (x - 1)); } n *) | ||
| 94 | Definition even_rec_functor := | ||
| 95 | EAttr {[ "__functor" := | ||
| 96 | AttrN (λ: "r" "x", if: "x" =: 0 then true else neg ("r" ("x" -: 1))) ]}. | ||
| 97 | |||
| 98 | Lemma even_rec_functor_wp e n : | ||
| 99 | 0 ≤ n ≤ int_max → | ||
| 100 | wp SHALLOW e (.= n) → | ||
| 101 | wp SHALLOW (even_rec_functor e) (.= Z.even n). | ||
| 102 | Proof. | ||
| 103 | intros Hn Hwp. apply App_wp. | ||
| 104 | revert e Hwp. induction (Z.lt_wf 0 n) as [n _ IH]; intros e Hwp. | ||
| 105 | apply Attr_wp_shallow. rewrite map_fmap_singleton /=. eapply Functor_wp. | ||
| 106 | { by apply no_recs_insert. } | ||
| 107 | { done. } | ||
| 108 | apply App_wp. apply β_wp. | ||
| 109 | rewrite /= !lookup_singleton !lookup_singleton_ne //=. apply Abs_wp, β_wp. | ||
| 110 | rewrite /= !lookup_singleton /= !lookup_singleton_ne //=. | ||
| 111 | rewrite -/even_rec_functor -/neg. | ||
| 112 | apply If_wp, BinOp_wp, Id_wp. | ||
| 113 | eapply wp_mono; [apply Hwp|]; intros ? ->. | ||
| 114 | eexists; split; [by constructor|]. | ||
| 115 | apply Lit_wp; [done|]. eexists; split; [by eexists|]. simpl. | ||
| 116 | destruct (n =? 0) eqn:Hn0; (apply Lit_wp; [done|]; eexists; split; [done|]; simpl). | ||
| 117 | { apply Lit_wp; [done|]. by apply Z.eqb_eq in Hn0 as ->. } | ||
| 118 | apply neg_wp, App_wp, Id_wp. | ||
| 119 | eapply wp_mono; [apply (IH (n-1))|]; [lia..| |]. | ||
| 120 | 2:{ intros e' He'. eapply wp_mono; [apply He'|]. | ||
| 121 | intros ? ->. eexists; split; [done|]. | ||
| 122 | by rewrite Z.negb_even Z.sub_1_r Z.odd_pred. } | ||
| 123 | eapply BinOp_wp, Id_wp. eapply wp_mono; [apply Hwp|]. intros ? ->. | ||
| 124 | eexists; split; [by constructor|]. apply Lit_wp; [done|]. | ||
| 125 | eexists; split; [eexists _, _; split_and!; [done| |done]|]. | ||
| 126 | - rewrite /= option_guard_True //. apply bool_decide_pack. | ||
| 127 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 128 | - apply Lit_wp; [|done]. apply bool_decide_pack. | ||
| 129 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 130 | Qed. | ||
| 131 | |||
| 132 | Lemma even_rec_functor_wp' n : | ||
| 133 | 0 ≤ n ≤ int_max → | ||
| 134 | wp SHALLOW (even_rec_functor n) (.= Z.even n). | ||
| 135 | Proof. | ||
| 136 | intros ?. apply even_rec_functor_wp; [done|]. apply Lit_wp; [|done]. | ||
| 137 | rewrite /= /int_ok. apply bool_decide_pack. | ||
| 138 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 139 | Qed. | ||
| 140 | |||
| 141 | (* ({ f ? (x: if x == 0 then true else !(f (x - 1))) }: f) {} n *) | ||
| 142 | Definition even_rec_default := | ||
| 143 | (λattr: | ||
| 144 | {[ "f" := Some (λ: "x", if: "x" =: 0 then true else neg ("f" ("x" -: 1))) ]}, "f") | ||
| 145 | (EAttr ∅). | ||
| 146 | |||
| 147 | Lemma even_rec_default_wp e n : | ||
| 148 | 0 ≤ n ≤ int_max → | ||
| 149 | wp SHALLOW e (.= n) → | ||
| 150 | wp SHALLOW (even_rec_default e) (.= Z.even n). | ||
| 151 | Proof. | ||
| 152 | intros Hn Hwp. apply App_wp. | ||
| 153 | eapply βMatch_wp; [done|repeat econstructor|]. simplify_map_eq. | ||
| 154 | rewrite -/even_rec_attr. by apply Id_wp, App_wp, even_rec_attr_wp. | ||
| 155 | Qed. | ||
| 156 | |||
| 157 | Lemma even_rec_default_wp' n : | ||
| 158 | 0 ≤ n ≤ int_max → | ||
| 159 | wp SHALLOW (even_rec_default n) (.= Z.even n). | ||
| 160 | Proof. | ||
| 161 | intros ?. apply even_rec_default_wp; [done|]. apply Lit_wp; [|done]. | ||
| 162 | rewrite /= /int_ok. apply bool_decide_pack. | ||
| 163 | rewrite /int_min Z.shiftl_mul_pow2 //. lia. | ||
| 164 | Qed. | ||