diff options
Diffstat (limited to 'theories/evallang')
| -rw-r--r-- | theories/evallang/interp.v | 52 | ||||
| -rw-r--r-- | theories/evallang/interp_proofs.v | 478 | ||||
| -rw-r--r-- | theories/evallang/operational.v | 140 | ||||
| -rw-r--r-- | theories/evallang/operational_props.v | 33 | ||||
| -rw-r--r-- | theories/evallang/tests.v | 33 |
5 files changed, 736 insertions, 0 deletions
diff --git a/theories/evallang/interp.v b/theories/evallang/interp.v new file mode 100644 index 0000000..d98b87f --- /dev/null +++ b/theories/evallang/interp.v | |||
| @@ -0,0 +1,52 @@ | |||
| 1 | From mininix Require Export res evallang.operational_props. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | Module Import evallang. | ||
| 5 | Export evallang. | ||
| 6 | |||
| 7 | Inductive thunk := Thunk { thunk_env : gmap string thunk; thunk_expr : expr }. | ||
| 8 | Add Printing Constructor thunk. | ||
| 9 | Notation env := (gmap string thunk). | ||
| 10 | |||
| 11 | Inductive val := | ||
| 12 | | VString (s : string) | ||
| 13 | | VClo (x : string) (E : env) (e : expr). | ||
| 14 | |||
| 15 | Global Instance maybe_VString : Maybe VString := λ v, | ||
| 16 | if v is VString s then Some s else None. | ||
| 17 | Global Instance maybe_VClo : Maybe3 VClo := λ v, | ||
| 18 | if v is VClo x E e then Some (x, E, e) else None. | ||
| 19 | |||
| 20 | Definition interp1 (interp : env → expr → res val) (E : env) (e : expr) : res val := | ||
| 21 | match e with | ||
| 22 | | EString s => | ||
| 23 | mret (VString s) | ||
| 24 | | EId ds x => | ||
| 25 | t ← Res $ (E !! x) ∪ (Thunk ∅ <$> ds); | ||
| 26 | interp (thunk_env t) (thunk_expr t) | ||
| 27 | | EEval ds e => | ||
| 28 | v ← interp E e; | ||
| 29 | s ← Res $ maybe VString v; | ||
| 30 | e ← Res $ parse s; | ||
| 31 | interp (E ∪ (Thunk ∅ <$> ds)) e | ||
| 32 | | EAbs ex e => | ||
| 33 | v ← interp E ex; | ||
| 34 | x ← Res $ maybe VString v; | ||
| 35 | mret (VClo x E e) | ||
| 36 | | EApp e1 e2 => | ||
| 37 | v1 ← interp E e1; | ||
| 38 | '(x, E', e') ← Res (maybe3 VClo v1); | ||
| 39 | interp (<[x:=Thunk E e2]> E') e' | ||
| 40 | end. | ||
| 41 | |||
| 42 | Fixpoint interp (n : nat) (E : env) (e : expr) : res val := | ||
| 43 | match n with | ||
| 44 | | O => NoFuel | ||
| 45 | | S n => interp1 (interp n) E e | ||
| 46 | end. | ||
| 47 | |||
| 48 | Global Opaque interp. | ||
| 49 | |||
| 50 | End evallang. | ||
| 51 | |||
| 52 | Add Printing Constructor evallang.thunk. | ||
diff --git a/theories/evallang/interp_proofs.v b/theories/evallang/interp_proofs.v new file mode 100644 index 0000000..0a26dd1 --- /dev/null +++ b/theories/evallang/interp_proofs.v | |||
| @@ -0,0 +1,478 @@ | |||
| 1 | From mininix Require Export evallang.interp. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | Module Import evallang. | ||
| 5 | Export evallang. | ||
| 6 | |||
| 7 | Lemma interp_S n : interp (S n) = interp1 (interp n). | ||
| 8 | Proof. done. Qed. | ||
| 9 | |||
| 10 | Fixpoint thunk_size (t : thunk) : nat := | ||
| 11 | S (map_sum_with thunk_size (thunk_env t)). | ||
| 12 | Definition env_size (E : env) : nat := | ||
| 13 | map_sum_with thunk_size E. | ||
| 14 | |||
| 15 | Lemma env_ind (P : env → Prop) : | ||
| 16 | (∀ E, map_Forall (λ i, P ∘ thunk_env) E → P E) → | ||
| 17 | ∀ E : env, P E. | ||
| 18 | Proof. | ||
| 19 | intros Pbs E. | ||
| 20 | induction (Nat.lt_wf_0_projected env_size E) as [E _ IH]. | ||
| 21 | apply Pbs, map_Forall_lookup=> y [E' e'] Hy. | ||
| 22 | apply (map_sum_with_lookup_le thunk_size) in Hy. | ||
| 23 | apply IH. by rewrite -Nat.le_succ_l. | ||
| 24 | Qed. | ||
| 25 | |||
| 26 | (** Correspondence to operational semantics *) | ||
| 27 | Definition subst_env' (thunk_to_expr : thunk → expr) (E : env) : expr → expr := | ||
| 28 | subst (thunk_to_expr <$> E). | ||
| 29 | Fixpoint thunk_to_expr (t : thunk) : expr := | ||
| 30 | subst_env' thunk_to_expr (thunk_env t) (thunk_expr t). | ||
| 31 | Notation subst_env := (subst_env' thunk_to_expr). | ||
| 32 | |||
| 33 | Lemma subst_env_eq e E : | ||
| 34 | subst_env E e = | ||
| 35 | match e with | ||
| 36 | | EString s => EString s | ||
| 37 | | EId ds x => EId ((thunk_to_expr <$> E !! x) ∪ ds) x | ||
| 38 | | EEval ds e => EEval ((thunk_to_expr <$> E) ∪ ds) (subst_env E e) | ||
| 39 | | EAbs ex e => EAbs (subst_env E ex) (subst_env E e) | ||
| 40 | | EApp e1 e2 => EApp (subst_env E e1) (subst_env E e2) | ||
| 41 | end. | ||
| 42 | Proof. destruct e; rewrite /subst_env' /= ?lookup_fmap //. Qed. | ||
| 43 | |||
| 44 | Lemma subst_env_alt E e : subst_env E e = subst (thunk_to_expr <$> E) e. | ||
| 45 | Proof. done. Qed. | ||
| 46 | |||
| 47 | (* Use the unfolding lemmas, don't rely on conversion *) | ||
| 48 | Opaque subst_env'. | ||
| 49 | |||
| 50 | Definition val_to_expr (v : val) : expr := | ||
| 51 | match v with | ||
| 52 | | VString s => EString s | ||
| 53 | | VClo x E e => EAbs (EString x) (subst_env E e) | ||
| 54 | end. | ||
| 55 | |||
| 56 | Lemma final_val_to_expr v : final (val_to_expr v). | ||
| 57 | Proof. by destruct v. Qed. | ||
| 58 | Lemma step_not_val_to_expr v e : val_to_expr v --> e → False. | ||
| 59 | Proof. intros []%step_not_final. apply final_val_to_expr. Qed. | ||
| 60 | |||
| 61 | Lemma subst_empty e : subst ∅ e = e. | ||
| 62 | Proof. induction e; f_equal/=; rewrite ?lookup_empty ?left_id_L //. Qed. | ||
| 63 | |||
| 64 | Lemma subst_env_empty e : subst_env ∅ e = e. | ||
| 65 | Proof. rewrite subst_env_alt. apply subst_empty. Qed. | ||
| 66 | |||
| 67 | Lemma interp_le {n1 n2 E e mv} : | ||
| 68 | interp n1 E e = Res mv → n1 ≤ n2 → interp n2 E e = Res mv. | ||
| 69 | Proof. | ||
| 70 | revert n2 E e mv. | ||
| 71 | induction n1 as [|n1 IH]; intros [|n2] E e mv He ?; [by (done || lia)..|]. | ||
| 72 | rewrite interp_S in He; rewrite interp_S; destruct e; | ||
| 73 | repeat match goal with | ||
| 74 | | _ => case_match | ||
| 75 | | H : context [(_ <$> ?mx)] |- _ => destruct mx eqn:?; simplify_res | ||
| 76 | | H : ?r ≫= _ = _ |- _ => destruct r as [[]|] eqn:?; simplify_res | ||
| 77 | | H : _ <$> ?r = _ |- _ => destruct r as [[]|] eqn:?; simplify_res | ||
| 78 | | |- interp ?n ?E ?e ≫= _ = _ => erewrite (IH n E e) by (done || lia) | ||
| 79 | | _ => progress simplify_res | ||
| 80 | | _ => progress simplify_option_eq | ||
| 81 | end; eauto with lia. | ||
| 82 | Qed. | ||
| 83 | |||
| 84 | Lemma interp_agree {n1 n2 E e mv1 mv2} : | ||
| 85 | interp n1 E e = Res mv1 → interp n2 E e = Res mv2 → mv1 = mv2. | ||
| 86 | Proof. | ||
| 87 | intros He1 He2. apply (inj Res). destruct (total (≤) n1 n2). | ||
| 88 | - rewrite -He2. symmetry. eauto using interp_le. | ||
| 89 | - rewrite -He1. eauto using interp_le. | ||
| 90 | Qed. | ||
| 91 | |||
| 92 | Lemma subst_env_union E1 E2 e : | ||
| 93 | subst_env (E1 ∪ E2) e = subst_env E1 (subst_env E2 e). | ||
| 94 | Proof. | ||
| 95 | revert E1 E2. induction e; intros E1 E2; rewrite subst_env_eq /=. | ||
| 96 | - done. | ||
| 97 | - rewrite !subst_env_eq lookup_union. by destruct (E1 !! _), (E2 !! _), ds. | ||
| 98 | - rewrite !(subst_env_eq (EEval _ _)) IHe. f_equal. | ||
| 99 | by rewrite assoc_L map_fmap_union. | ||
| 100 | - rewrite !(subst_env_eq (EAbs _ _)) /=. f_equal; auto. | ||
| 101 | - rewrite !(subst_env_eq (EApp _ _)) /=. f_equal; auto. | ||
| 102 | Qed. | ||
| 103 | |||
| 104 | Lemma subst_env_insert E x e t : | ||
| 105 | subst_env (<[x:=t]> E) e = subst {[x:=thunk_to_expr t]} (subst_env E e). | ||
| 106 | Proof. | ||
| 107 | rewrite insert_union_singleton_l subst_env_union subst_env_alt. | ||
| 108 | by rewrite map_fmap_singleton. | ||
| 109 | Qed. | ||
| 110 | |||
| 111 | Lemma subst_env_insert_eq e1 e2 E1 E2 x E1' E2' e1' e2' : | ||
| 112 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 113 | subst_env E1' e1' = subst_env E2' e2' → | ||
| 114 | subst_env (<[x:=Thunk E1' e1']> E1) e1 = subst_env (<[x:=Thunk E2' e2']> E2) e2. | ||
| 115 | Proof. intros He He'. by rewrite !subst_env_insert //= He' He. Qed. | ||
| 116 | |||
| 117 | Lemma option_fmap_thunk_to_expr_Thunk (me : option expr) : | ||
| 118 | thunk_to_expr <$> (Thunk ∅ <$> me) = me. | ||
| 119 | Proof. destruct me; f_equal/=. by rewrite subst_env_empty. Qed. | ||
| 120 | |||
| 121 | Lemma map_fmap_thunk_to_expr_Thunk (es : gmap string expr) : | ||
| 122 | thunk_to_expr <$> (Thunk ∅ <$> es) = es. | ||
| 123 | Proof. | ||
| 124 | apply map_eq=> x. by rewrite !lookup_fmap option_fmap_thunk_to_expr_Thunk. | ||
| 125 | Qed. | ||
| 126 | |||
| 127 | Lemma subst_env_eval_eq E1 E2 ds1 ds2 e : | ||
| 128 | (thunk_to_expr <$> E1) ∪ ds1 = (thunk_to_expr <$> E2) ∪ ds2 → | ||
| 129 | subst_env (E1 ∪ (Thunk ∅ <$> ds1)) e = subst_env (E2 ∪ (Thunk ∅ <$> ds2)) e. | ||
| 130 | Proof. | ||
| 131 | intros HE. | ||
| 132 | by rewrite !subst_env_alt !map_fmap_union !map_fmap_thunk_to_expr_Thunk HE. | ||
| 133 | Qed. | ||
| 134 | |||
| 135 | Lemma interp_proper n E1 E2 e1 e2 mv : | ||
| 136 | subst_env E1 e1 = subst_env E2 e2 → | ||
| 137 | interp n E1 e1 = Res mv → | ||
| 138 | ∃ mw m, interp m E2 e2 = Res mw ∧ | ||
| 139 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 140 | Proof. | ||
| 141 | revert n E1 E2 e1 e2 mv. induction n as [|n IHn]; [done|]. | ||
| 142 | intros E1 E2 e1 e2 mv Hsubst Hinterp. | ||
| 143 | rewrite 2!subst_env_eq in Hsubst. | ||
| 144 | rewrite interp_S in Hinterp. destruct e1, e2; simplify_res; try done. | ||
| 145 | - eexists (Some (VString _)), 1. by rewrite interp_S. | ||
| 146 | - assert (thunk_to_expr <$> E1 !! x0 ∪ (Thunk ∅ <$> ds) = | ||
| 147 | thunk_to_expr <$> E2 !! x0 ∪ (Thunk ∅ <$> ds0)). | ||
| 148 | { destruct (E1 !! _), (E2 !! _), ds, ds0; simplify_eq/=; | ||
| 149 | f_equal/=; by rewrite ?subst_env_empty. } | ||
| 150 | destruct (E1 !! x0 ∪ (Thunk ∅ <$> ds)) as [[E1' e1']|], | ||
| 151 | (E2 !! x0 ∪ (Thunk ∅ <$> ds0)) as [[E2' e2']|] eqn:HE2; | ||
| 152 | simplify_res; last first. | ||
| 153 | { exists None, 1. by rewrite interp_S /= HE2. } | ||
| 154 | eapply IHn in Hinterp as (mw & m & Hinterp2 & ?); [|by eauto..]. | ||
| 155 | exists mw, (S m). split; [|done]. rewrite interp_S /= HE2 /=. done. | ||
| 156 | - destruct (interp n _ e1) as [mv1|] eqn:Hinterp'; simplify_eq/=. | ||
| 157 | eapply IHn in Hinterp' as (mw1 & m1 & Hinterp1 & ?); last done. | ||
| 158 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 159 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 160 | destruct (maybe VString v1) as [x|] eqn:Hv1; | ||
| 161 | simplify_res; last first. | ||
| 162 | { exists None, (S m1). split; [|done]. rewrite interp_S /= Hinterp1 /=. | ||
| 163 | destruct v1, w1; repeat destruct select base_lit; by simplify_eq/=. } | ||
| 164 | destruct v1, w1; repeat destruct select base_lit; simplify_eq/=. | ||
| 165 | destruct (parse _) as [e|] eqn:Hparse; simplify_res; last first. | ||
| 166 | { exists None, (S m1). split; [|done]. rewrite interp_S /= Hinterp1 /=. | ||
| 167 | by rewrite Hparse. } | ||
| 168 | eapply IHn in Hinterp | ||
| 169 | as (mw & m2 & Hinterp2 & ?); last by apply subst_env_eval_eq. | ||
| 170 | exists mw, (S (m1 `max` m2)). split; [|done]. rewrite interp_S /=. | ||
| 171 | rewrite (interp_le Hinterp1) /=; last lia. rewrite Hparse /=. | ||
| 172 | eauto using interp_le with lia. | ||
| 173 | - destruct (interp n _ _) as [mv1|] eqn:Hinterp'; simplify_eq/=. | ||
| 174 | eapply IHn in Hinterp' as (mw1 & m1 & Hinterp1 & ?); last done. | ||
| 175 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 176 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 177 | destruct (maybe VString _) eqn:Hstring; simplify_res; last first. | ||
| 178 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 179 | by assert (maybe VString w1 = None) as -> by (by destruct v1, w1). } | ||
| 180 | destruct v1, w1; simplify_eq/=. | ||
| 181 | eexists (Some (VClo _ _ _)), (S m1). | ||
| 182 | rewrite interp_S /= Hinterp1 /=. split; [done|]. by do 2 f_equal/=. | ||
| 183 | - destruct (interp n _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 184 | eapply IHn in Hinterp' as (mw' & m1 & Hinterp1 & ?); last done. | ||
| 185 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 186 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 187 | destruct (maybe3 VClo _) eqn:Hclo; simplify_res; last first. | ||
| 188 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 189 | by assert (maybe3 VClo w' = None) as -> by (by destruct v', w'). } | ||
| 190 | destruct v', w'; simplify_eq/=. | ||
| 191 | eapply IHn with (E2 := <[x0:=Thunk E2 e2_2]> E0) in Hinterp | ||
| 192 | as (w & m2 & Hinterp2 & ?); last by apply subst_env_insert_eq. | ||
| 193 | exists w, (S (m1 `max` m2)). rewrite interp_S /=. | ||
| 194 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 195 | rewrite (interp_le Hinterp2) /=; last lia. done. | ||
| 196 | Qed. | ||
| 197 | |||
| 198 | Lemma subst_as_subst_env x e1 e2 : | ||
| 199 | subst {[x:=e2]} e1 = subst_env (<[x:=Thunk ∅ e2]> ∅) e1. | ||
| 200 | Proof. rewrite subst_env_insert //= !subst_env_empty //. Qed. | ||
| 201 | |||
| 202 | Lemma interp_subst_abs n x e1 e2 mv : | ||
| 203 | interp n ∅ (subst {[x:=e2]} e1) = Res mv → | ||
| 204 | ∃ mw m, interp m (<[x:=Thunk ∅ e2]> ∅) e1 = Res mw ∧ | ||
| 205 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 206 | Proof. | ||
| 207 | apply interp_proper. by rewrite subst_env_empty subst_as_subst_env. | ||
| 208 | Qed. | ||
| 209 | |||
| 210 | Lemma interp_subst_eval n e ds mv : | ||
| 211 | interp n ∅ (subst ds e) = Res mv → | ||
| 212 | ∃ mw m, interp m (Thunk ∅ <$> ds) e = Res mw ∧ | ||
| 213 | val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 214 | Proof. | ||
| 215 | apply interp_proper. | ||
| 216 | by rewrite subst_env_empty subst_env_alt map_fmap_thunk_to_expr_Thunk. | ||
| 217 | Qed. | ||
| 218 | |||
| 219 | Lemma interp_step e1 e2 n mv : | ||
| 220 | e1 --> e2 → | ||
| 221 | interp n ∅ e2 = Res mv → | ||
| 222 | ∃ mw m, interp m ∅ e1 = Res mw ∧ val_to_expr <$> mv = val_to_expr <$> mw. | ||
| 223 | Proof. | ||
| 224 | intros Hstep. revert mv n. | ||
| 225 | induction Hstep; intros mv n Hinterp. | ||
| 226 | - apply interp_subst_abs in Hinterp as (mw & [|m] & Hinterp & Hv); | ||
| 227 | simplify_eq/=; [|done..]. | ||
| 228 | exists mw, (S (S (S m))). rewrite !interp_S /= -!interp_S. | ||
| 229 | eauto using interp_le with lia. | ||
| 230 | - exists mv, (S n). rewrite !interp_S /=. | ||
| 231 | rewrite lookup_empty left_id_L /=. done. | ||
| 232 | - apply interp_subst_eval in Hinterp as (mw & [|m] & Hinterp & Hv); | ||
| 233 | simplify_eq/=; [|done..]. | ||
| 234 | exists mw, (S (S m)). rewrite !interp_S /= -interp_S. | ||
| 235 | rewrite left_id_L H /=. done. | ||
| 236 | - destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 237 | destruct (interp n _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 238 | apply IHHstep in Hinterp' as (mw' & m1 & Hinterp1 & ?); simplify_res. | ||
| 239 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 240 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 241 | destruct (maybe VString _) eqn:Hstring; simplify_res; last first. | ||
| 242 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 243 | by assert (maybe VString w' = None) as -> by (by destruct v', w'). } | ||
| 244 | destruct v', w'; simplify_eq/=. | ||
| 245 | eexists (Some (VClo _ _ _)), (S m1). rewrite !interp_S /=. | ||
| 246 | rewrite (interp_le Hinterp1) /=; last lia. done. | ||
| 247 | - destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 248 | destruct (interp n _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 249 | apply IHHstep in Hinterp' as (mw' & m1 & Hinterp1 & ?); simplify_res. | ||
| 250 | destruct mv' as [v'|], mw' as [w'|]; simplify_res; last first. | ||
| 251 | { exists None, (S m1). by rewrite interp_S /= Hinterp1. } | ||
| 252 | destruct (maybe3 VClo _) eqn:Hclo; simplify_res; last first. | ||
| 253 | { exists None, (S m1). rewrite interp_S /= Hinterp1 /=. | ||
| 254 | by assert (maybe3 VClo w' = None) as -> by (by destruct v', w'). } | ||
| 255 | destruct v', w'; simplify_eq/=. | ||
| 256 | eapply interp_proper in Hinterp as (mw & m2 & Hinterp2 & Hv); | ||
| 257 | last apply subst_env_insert_eq; try done. | ||
| 258 | exists mw, (S (m1 `max` m2)). rewrite !interp_S /=. | ||
| 259 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 260 | by rewrite (interp_le Hinterp2) /=; last lia. | ||
| 261 | - destruct n as [|n]; [done|rewrite interp_S /= in Hinterp]. | ||
| 262 | destruct (interp n _ e1') as [mv1|] eqn:Hinterp1; simplify_eq/=. | ||
| 263 | apply IHHstep in Hinterp1 as (mw1 & m & Hinterp1 & Hw1). | ||
| 264 | destruct mv1 as [v1|], mw1 as [w1|]; simplify_res; last first. | ||
| 265 | { exists None, (S m). by rewrite interp_S /= Hinterp1. } | ||
| 266 | destruct (maybe VString _) eqn:Hstring; simplify_res; last first. | ||
| 267 | { exists None, (S m). rewrite interp_S /= Hinterp1 /=. | ||
| 268 | by assert (maybe VString w1 = None) as -> by (by destruct v1, w1). } | ||
| 269 | destruct v1, w1; simplify_eq/=. | ||
| 270 | exists mv, (S (n `max` m)). split; [|done]. rewrite interp_S /=. | ||
| 271 | rewrite (interp_le Hinterp1) /=; last lia. | ||
| 272 | destruct (parse _); simplify_res; eauto using interp_le with lia. | ||
| 273 | Qed. | ||
| 274 | |||
| 275 | Lemma final_interp e : | ||
| 276 | final e → | ||
| 277 | ∃ w m, interp m ∅ e = mret w ∧ e = val_to_expr w. | ||
| 278 | Proof. | ||
| 279 | induction e as [| | |[]|]; inv 1. | ||
| 280 | - eexists (VString _), 1. by rewrite interp_S /=. | ||
| 281 | - eexists (VClo _ _ _), 2. rewrite interp_S /=. split; [done|]. | ||
| 282 | by rewrite subst_env_empty. | ||
| 283 | Qed. | ||
| 284 | |||
| 285 | Lemma red_final_interp e : | ||
| 286 | red step e ∨ final e ∨ ∃ m, interp m ∅ e = mfail. | ||
| 287 | Proof. | ||
| 288 | induction e. | ||
| 289 | - (* ENat *) right; left. constructor. | ||
| 290 | - (* EId *) destruct ds as [e|]. | ||
| 291 | + left. by repeat econstructor. | ||
| 292 | + do 2 right. by exists 1. | ||
| 293 | - (* EEval *) destruct IHe as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 294 | + left. by repeat econstructor. | ||
| 295 | + apply final_interp in Hfinal as (w & m & Hinterp & ->). | ||
| 296 | destruct (maybe VString w) as [x|] eqn:Hw; last first. | ||
| 297 | { do 2 right. exists (S m). rewrite interp_S /= Hinterp /=. | ||
| 298 | by rewrite Hw. } | ||
| 299 | destruct w; simplify_eq/=. | ||
| 300 | destruct (parse x) as [e|] eqn:Hparse; last first. | ||
| 301 | { do 2 right. exists (S m). rewrite interp_S /= Hinterp /=. | ||
| 302 | by rewrite Hparse. } | ||
| 303 | left. by repeat econstructor. | ||
| 304 | + do 2 right. exists (S m). rewrite interp_S /= Hinterp. done. | ||
| 305 | - (* EAbs *) destruct IHe1 as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 306 | + left. by repeat econstructor. | ||
| 307 | + apply final_interp in Hfinal as (w & m & Hinterp & ->). | ||
| 308 | destruct (maybe VString w) as [x|] eqn:Hw; last first. | ||
| 309 | { do 2 right. exists (S m). rewrite interp_S /= Hinterp /=. | ||
| 310 | by rewrite Hw. } | ||
| 311 | destruct w; naive_solver. | ||
| 312 | + do 2 right. exists (S m). rewrite interp_S /= Hinterp. done. | ||
| 313 | - (* EApp *) destruct IHe1 as [[??]|[Hfinal|[m Hinterp]]]. | ||
| 314 | + left. by repeat econstructor. | ||
| 315 | + apply final_interp in Hfinal as (w & m & Hinterp & ->). | ||
| 316 | destruct (maybe3 VClo w) eqn:Hw. | ||
| 317 | { destruct w; simplify_eq/=. left. by repeat econstructor. } | ||
| 318 | do 2 right. exists (S m). by rewrite interp_S /= Hinterp /= Hw. | ||
| 319 | + do 2 right. exists (S m). by rewrite interp_S /= Hinterp. | ||
| 320 | Qed. | ||
| 321 | |||
| 322 | Lemma interp_complete e1 e2 : | ||
| 323 | e1 -->* e2 → | ||
| 324 | nf step e2 → | ||
| 325 | ∃ mw m, interp m ∅ e1 = Res mw ∧ | ||
| 326 | if mw is Some w then e2 = val_to_expr w else ¬final e2. | ||
| 327 | Proof. | ||
| 328 | intros Hsteps Hnf. induction Hsteps as [e|e1 e2 e3 Hstep _ IH]. | ||
| 329 | { destruct (red_final_interp e) as [?|[Hfinal|[m Hinterp]]]; [done|..]. | ||
| 330 | - apply final_interp in Hfinal as (w & m & ? & ?). | ||
| 331 | by exists (Some w), m. | ||
| 332 | - exists None, m. split; [done|]. intros Hfinal. | ||
| 333 | apply final_interp in Hfinal as (w & m' & ? & _). | ||
| 334 | by assert (mfail = mret w) by eauto using interp_agree. } | ||
| 335 | destruct IH as (mw & m & Hinterp & ?); try done. | ||
| 336 | eapply interp_step in Hinterp as (mw' & m' & ? & ?); last done. | ||
| 337 | destruct mw, mw'; naive_solver. | ||
| 338 | Qed. | ||
| 339 | |||
| 340 | Lemma interp_complete_ret e1 e2 : | ||
| 341 | e1 -->* e2 → final e2 → | ||
| 342 | ∃ w m, interp m ∅ e1 = mret w ∧ e2 = val_to_expr w. | ||
| 343 | Proof. | ||
| 344 | intros Hsteps Hfinal. apply interp_complete in Hsteps | ||
| 345 | as ([w|] & m & ? & ?); naive_solver eauto using final_nf. | ||
| 346 | Qed. | ||
| 347 | Lemma interp_complete_fail e1 e2 : | ||
| 348 | e1 -->* e2 → nf step e2 → ¬final e2 → | ||
| 349 | ∃ m, interp m ∅ e1 = mfail. | ||
| 350 | Proof. | ||
| 351 | intros Hsteps Hnf Hforce. | ||
| 352 | apply interp_complete in Hsteps as ([w|] & m & ? & ?); simplify_eq/=; try by eauto. | ||
| 353 | destruct Hforce. apply final_val_to_expr. | ||
| 354 | Qed. | ||
| 355 | |||
| 356 | Lemma interp_sound_open E e n mv : | ||
| 357 | interp n E e = Res mv → | ||
| 358 | ∃ e', subst_env E e -->* e' ∧ | ||
| 359 | if mv is Some v then e' = val_to_expr v else stuck e'. | ||
| 360 | Proof. | ||
| 361 | revert E e mv. | ||
| 362 | induction n as [|n IH]; intros E e mv Hinterp; first done. | ||
| 363 | rewrite subst_env_eq. rewrite interp_S in Hinterp. | ||
| 364 | destruct e; simplify_res. | ||
| 365 | - (* EString *) by eexists. | ||
| 366 | - (* EId *) | ||
| 367 | assert (thunk_to_expr <$> (E !! x) ∪ (Thunk ∅ <$> ds) | ||
| 368 | = (thunk_to_expr <$> E !! x) ∪ ds). | ||
| 369 | { destruct (_ !! _), ds; f_equal/=. by rewrite subst_env_empty. } | ||
| 370 | destruct (_ ∪ (_ <$> _)) as [[E1 e1]|], (_ ∪ _) as [e2|]; simplify_res. | ||
| 371 | * apply IH in Hinterp as (e'' & Hsteps & He''). | ||
| 372 | exists e''; split; [|done]. | ||
| 373 | eapply rtc_l; [|done]. by econstructor. | ||
| 374 | * eexists; split; [done|]. split; [|inv 1]. | ||
| 375 | intros [? Hstep]. inv_step; simplify_eq/=; congruence. | ||
| 376 | - (* EEval *) | ||
| 377 | destruct (interp _ _ _) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 378 | apply IH in Hinterp1 as (e1' & Hsteps1 & He1'). | ||
| 379 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 380 | { eexists; split; [by eapply SEval_rtc|]. split; [|inv 1]. | ||
| 381 | intros [??]. destruct He1' as [Hnf []]. | ||
| 382 | inv_step; simpl; eauto. destruct Hnf; eauto. } | ||
| 383 | destruct (maybe VString _) as [x|] eqn:Hv1; simplify_res; last first. | ||
| 384 | { eexists; split; [by eapply SEval_rtc|]. split; [|inv 1]. | ||
| 385 | intros [??]. destruct v1; inv_step. } | ||
| 386 | destruct v1; simplify_eq/=. | ||
| 387 | destruct (parse x) as [ex|] eqn:Hparse; simplify_res; last first. | ||
| 388 | { eexists; split; [by eapply SEval_rtc|]. | ||
| 389 | split; [|inv 1]. intros [??]. inv_step. } | ||
| 390 | apply IH in Hinterp as (e'' & Hsteps & He''). | ||
| 391 | exists e''; split; [|done]. etrans; [by eapply SEval_rtc|]. | ||
| 392 | eapply rtc_l; [by econstructor|]. | ||
| 393 | by rewrite subst_env_alt map_fmap_union | ||
| 394 | map_fmap_thunk_to_expr_Thunk in Hsteps. | ||
| 395 | - (* EAbs *) | ||
| 396 | destruct (interp _ _ _) as [mv1|] eqn:Hinterp1; simplify_res. | ||
| 397 | apply IH in Hinterp1 as (e1' & Hsteps1 & He1'). | ||
| 398 | destruct mv1 as [v1|]; simplify_res; last first. | ||
| 399 | { eexists; split; [by eapply SAbsL_rtc|]. split. | ||
| 400 | + intros [??]. destruct He1' as [Hnf []]. | ||
| 401 | inv_step; simpl; eauto. destruct Hnf; eauto. | ||
| 402 | + intros ?. destruct He1' as [_ []]. by destruct e1'. } | ||
| 403 | eexists; split; [by eapply SAbsL_rtc|]. | ||
| 404 | destruct (maybe VString _) as [x|] eqn:Hv1; simplify_res; last first. | ||
| 405 | { split; [|destruct v1; inv 1]. intros [??]. destruct v1; inv_step. } | ||
| 406 | by destruct v1; simplify_eq/=. | ||
| 407 | - (* EApp *) destruct (interp _ _ _) as [mv'|] eqn:Hinterp'; simplify_res. | ||
| 408 | apply IH in Hinterp' as (e' & Hsteps & He'); try done. | ||
| 409 | destruct mv' as [v'|]; simplify_res; last first. | ||
| 410 | { eexists; repeat split; [by apply SAppL_rtc| |inv 1]. | ||
| 411 | intros [e'' Hstep]. destruct He' as [Hnf Hfinal]. | ||
| 412 | inv Hstep; [by destruct Hfinal; constructor|]. destruct Hnf. eauto. } | ||
| 413 | destruct (maybe3 VClo v') eqn:?; simplify_res; last first. | ||
| 414 | { eexists; repeat split; [by apply SAppL_rtc| |inv 1]. | ||
| 415 | intros [e'' Hstep]. inv Hstep; destruct v'; by repeat inv_step. } | ||
| 416 | destruct v'; simplify_res. | ||
| 417 | apply IH in Hinterp as (e'' & Hsteps' & He''). | ||
| 418 | eexists; split; [|done]. etrans; [by apply SAppL_rtc|]. | ||
| 419 | eapply rtc_l; first by constructor. | ||
| 420 | rewrite subst_env_insert // in Hsteps'. | ||
| 421 | Qed. | ||
| 422 | |||
| 423 | Lemma interp_sound n e mv : | ||
| 424 | interp n ∅ e = Res mv → | ||
| 425 | ∃ e', e -->* e' ∧ if mv is Some v then e' = val_to_expr v else stuck e'. | ||
| 426 | Proof. | ||
| 427 | intros Hsteps%interp_sound_open; try done. | ||
| 428 | by rewrite subst_env_empty in Hsteps. | ||
| 429 | Qed. | ||
| 430 | |||
| 431 | (** Final theorems *) | ||
| 432 | Theorem interp_sound_complete_ret e v : | ||
| 433 | (∃ w n, interp n ∅ e = mret w ∧ val_to_expr v = val_to_expr w) | ||
| 434 | ↔ e -->* val_to_expr v. | ||
| 435 | Proof. | ||
| 436 | split. | ||
| 437 | - by intros (n & w & (e' & ? & ->)%interp_sound & ->). | ||
| 438 | - intros Hsteps. apply interp_complete in Hsteps as ([] & ? & ? & ?); | ||
| 439 | unfold nf, red; | ||
| 440 | naive_solver eauto using final_val_to_expr, step_not_val_to_expr. | ||
| 441 | Qed. | ||
| 442 | |||
| 443 | Theorem interp_sound_complete_ret_string e s : | ||
| 444 | (∃ n, interp n ∅ e = mret (VString s)) ↔ e -->* EString s. | ||
| 445 | Proof. | ||
| 446 | split. | ||
| 447 | - by intros [n (e' & ? & ->)%interp_sound]. | ||
| 448 | - intros Hsteps. apply interp_complete_ret in Hsteps as ([] & ? & ? & ?); | ||
| 449 | simplify_eq/=; eauto. | ||
| 450 | Qed. | ||
| 451 | |||
| 452 | Theorem interp_sound_complete_fail e : | ||
| 453 | (∃ n, interp n ∅ e = mfail) ↔ ∃ e', e -->* e' ∧ stuck e'. | ||
| 454 | Proof. | ||
| 455 | split. | ||
| 456 | - by intros [n ?%interp_sound]. | ||
| 457 | - intros (e' & Hsteps & Hnf & Hforced). by eapply interp_complete_fail. | ||
| 458 | Qed. | ||
| 459 | |||
| 460 | Theorem interp_sound_complete_no_fuel e : | ||
| 461 | (∀ n, interp n ∅ e = NoFuel) ↔ all_loop step e. | ||
| 462 | Proof. | ||
| 463 | rewrite all_loop_alt. split. | ||
| 464 | - intros Hnofuel e' Hsteps. | ||
| 465 | destruct (red_final_interp e') as [|[|He']]; [done|..]. | ||
| 466 | + apply interp_complete_ret in Hsteps as (w & m & Hinterp & _); last done. | ||
| 467 | by rewrite Hnofuel in Hinterp. | ||
| 468 | + apply interp_sound_complete_fail in He' as (e'' & ? & [Hnf _]). | ||
| 469 | destruct (interp_complete e e'') as (mv & n & Hinterp & _); [by etrans|done|]. | ||
| 470 | by rewrite Hnofuel in Hinterp. | ||
| 471 | - intros Hred n. destruct (interp n ∅ e) as [mv|] eqn:Hinterp; [|done]. | ||
| 472 | apply interp_sound in Hinterp as (e' & Hsteps%Hred & Hstuck). | ||
| 473 | destruct mv as [v|]; simplify_eq/=. | ||
| 474 | + apply final_nf in Hsteps as []. apply final_val_to_expr. | ||
| 475 | + by destruct Hstuck as [[] ?]. | ||
| 476 | Qed. | ||
| 477 | |||
| 478 | End evallang. | ||
diff --git a/theories/evallang/operational.v b/theories/evallang/operational.v new file mode 100644 index 0000000..79174dd --- /dev/null +++ b/theories/evallang/operational.v | |||
| @@ -0,0 +1,140 @@ | |||
| 1 | From Coq Require Import Ascii. | ||
| 2 | From mininix Require Export utils. | ||
| 3 | From stdpp Require Import options. | ||
| 4 | |||
| 5 | Module Import evallang. | ||
| 6 | |||
| 7 | Inductive expr := | ||
| 8 | | EString (s : string) | ||
| 9 | | EId (ds : option expr) (x : string) | ||
| 10 | | EEval (ds : gmap string expr) (ee : expr) | ||
| 11 | | EAbs (ex e : expr) | ||
| 12 | | EApp (e1 e2 : expr). | ||
| 13 | |||
| 14 | Module parser. | ||
| 15 | Inductive token := | ||
| 16 | | TId (s : string) | ||
| 17 | | TString (s : string) | ||
| 18 | | TColon | ||
| 19 | | TExclamation | ||
| 20 | | TParenL | ||
| 21 | | TParenR. | ||
| 22 | |||
| 23 | Inductive token_state := | ||
| 24 | TSString (s : string) | TSId (s : string) | TSOther. | ||
| 25 | |||
| 26 | Definition token_state_push (st : token_state) (k : list token) : list token := | ||
| 27 | match st with | ||
| 28 | | TSId s => TId (String.rev s) :: k | ||
| 29 | | _ => k | ||
| 30 | end. | ||
| 31 | |||
| 32 | Fixpoint tokenize_go (sin : string) (st : token_state) | ||
| 33 | (k : list token) : option (list token) := | ||
| 34 | match sin, st with | ||
| 35 | | "", TSString _ => None (* no closing "" *) | ||
| 36 | | "", _ => Some (reverse (token_state_push st k)) | ||
| 37 | | String "\" (String """" sin), TSString s => | ||
| 38 | tokenize_go sin (TSString (String """" s)) k | ||
| 39 | | String """" sin, TSString s => | ||
| 40 | tokenize_go sin TSOther (TString (String.rev s) :: k) | ||
| 41 | | String a sin, TSString s => tokenize_go sin (TSString (String a s)) k | ||
| 42 | | String ":" sin, _ => tokenize_go sin TSOther (TColon :: token_state_push st k) | ||
| 43 | | String "!" sin, _ => tokenize_go sin TSOther (TExclamation :: token_state_push st k) | ||
| 44 | | String "(" sin, _ => tokenize_go sin TSOther (TParenL :: token_state_push st k) | ||
| 45 | | String ")" sin, _ => tokenize_go sin TSOther (TParenR :: token_state_push st k) | ||
| 46 | | String """" sin, _ => tokenize_go sin (TSString "") k | ||
| 47 | | String a sin, TSOther => | ||
| 48 | if Ascii.is_space a then tokenize_go sin TSOther k | ||
| 49 | else tokenize_go sin (TSId (String a EmptyString)) k | ||
| 50 | | String a sin, TSId s => | ||
| 51 | if Ascii.is_space a then tokenize_go sin TSOther (TId (String.rev s) :: k) | ||
| 52 | else tokenize_go sin (TSId (String a s)) k | ||
| 53 | end. | ||
| 54 | Definition tokenize (sin : string) : option (list token) := | ||
| 55 | tokenize_go sin TSOther []. | ||
| 56 | |||
| 57 | Inductive stack_item := | ||
| 58 | | SExpr (e : expr) | ||
| 59 | | SAbsR (e : expr) | ||
| 60 | | SEval | ||
| 61 | | SParenL. | ||
| 62 | |||
| 63 | Definition stack_push (e : expr) (k : list stack_item) : list stack_item := | ||
| 64 | match k with | ||
| 65 | | SExpr e1 :: k => SExpr (EApp e1 e) :: k | ||
| 66 | | SEval :: k => SExpr (EEval ∅ e) :: k | ||
| 67 | | _ => SExpr e :: k | ||
| 68 | end. | ||
| 69 | |||
| 70 | Fixpoint stack_pop_go (e : expr) | ||
| 71 | (k : list stack_item) : option (expr * list stack_item) := | ||
| 72 | match k with | ||
| 73 | | SAbsR e1 :: k => stack_pop_go (EAbs e1 e) k | ||
| 74 | | _ => Some (e, k) | ||
| 75 | end. | ||
| 76 | |||
| 77 | Definition stack_pop (k : list stack_item) : option (expr * list stack_item) := | ||
| 78 | match k with | ||
| 79 | | SExpr e :: k => stack_pop_go e k | ||
| 80 | | _ => None | ||
| 81 | end. | ||
| 82 | |||
| 83 | Fixpoint parse_go (ts : list token) (k : list stack_item) : option expr := | ||
| 84 | match ts with | ||
| 85 | | [] => '(e, k) ← stack_pop k; guard (k = []);; Some e | ||
| 86 | | TString x :: ts => parse_go ts (stack_push (EString x) k) | ||
| 87 | | TId "eval" :: TExclamation :: ts => parse_go ts (SEval :: k) | ||
| 88 | | TId x :: TColon :: ts => parse_go ts (SAbsR (EString x) :: k) | ||
| 89 | | TId x :: ts => parse_go ts (stack_push (EId None x) k) | ||
| 90 | | TColon :: ts => | ||
| 91 | '(e, k) ← stack_pop k; | ||
| 92 | parse_go ts (SAbsR e :: k) | ||
| 93 | | TParenL :: ts => parse_go ts (SParenL :: k) | ||
| 94 | | TParenR :: ts => | ||
| 95 | '(e, k) ← stack_pop k; | ||
| 96 | match k with | ||
| 97 | | SParenL :: k => parse_go ts (stack_push e k) | ||
| 98 | | _ => None | ||
| 99 | end | ||
| 100 | | _ => None | ||
| 101 | end. | ||
| 102 | |||
| 103 | Definition parse (sin : string) : option expr := | ||
| 104 | ts ← tokenize sin; parse_go ts []. | ||
| 105 | End parser. | ||
| 106 | |||
| 107 | Definition parse := parser.parse. | ||
| 108 | |||
| 109 | Fixpoint subst (ds : gmap string expr) (e : expr) : expr := | ||
| 110 | match e with | ||
| 111 | | EString s => EString s | ||
| 112 | | EId ds' x => EId (ds !! x ∪ ds') x | ||
| 113 | | EEval ds' ee => EEval (ds ∪ ds') (subst ds ee) | ||
| 114 | | EAbs ex e => EAbs (subst ds ex) (subst ds e) | ||
| 115 | | EApp e1 e2 => EApp (subst ds e1) (subst ds e2) | ||
| 116 | end. | ||
| 117 | |||
| 118 | Reserved Infix "-->" (right associativity, at level 55). | ||
| 119 | Inductive step : expr → expr → Prop := | ||
| 120 | | Sβ x e1 e2 : EApp (EAbs (EString x) e1) e2 --> subst {[x:=e2]} e1 | ||
| 121 | | SId e x : EId (Some e) x --> e | ||
| 122 | | SEvalString ds s e : parse s = Some e → EEval ds (EString s) --> subst ds e | ||
| 123 | | SAbsL ex1 ex1' e : ex1 --> ex1' → EAbs ex1 e --> EAbs ex1' e | ||
| 124 | | SAppL e1 e1' e2 : e1 --> e1' → EApp e1 e2 --> EApp e1' e2 | ||
| 125 | | SEval ds e1 e1' : e1 --> e1' → EEval ds e1 --> EEval ds e1' | ||
| 126 | where "e1 --> e2" := (step e1 e2). | ||
| 127 | |||
| 128 | Infix "-->*" := (rtc step) (right associativity, at level 55). | ||
| 129 | |||
| 130 | Definition final (e : expr) : Prop := | ||
| 131 | match e with | ||
| 132 | | EString _ => True | ||
| 133 | | EAbs (EString _) _ => True | ||
| 134 | | _ => False | ||
| 135 | end. | ||
| 136 | |||
| 137 | Definition stuck (e : expr) : Prop := | ||
| 138 | nf step e ∧ ¬final e. | ||
| 139 | |||
| 140 | End evallang. | ||
diff --git a/theories/evallang/operational_props.v b/theories/evallang/operational_props.v new file mode 100644 index 0000000..31724c0 --- /dev/null +++ b/theories/evallang/operational_props.v | |||
| @@ -0,0 +1,33 @@ | |||
| 1 | From mininix Require Export evallang.operational. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | Module Import evallang. | ||
| 5 | Export evallang. | ||
| 6 | |||
| 7 | (** Properties of operational semantics *) | ||
| 8 | Lemma step_not_final e1 e2 : e1 --> e2 → ¬final e1. | ||
| 9 | Proof. induction 1; simpl; repeat case_match; naive_solver. Qed. | ||
| 10 | Lemma final_nf e : final e → nf step e. | ||
| 11 | Proof. by intros ? [??%step_not_final]. Qed. | ||
| 12 | |||
| 13 | Lemma SAbsL_rtc ex1 ex1' e : ex1 -->* ex1' → EAbs ex1 e -->* EAbs ex1' e. | ||
| 14 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
| 15 | Lemma SAppL_rtc e1 e1' e2 : e1 -->* e1' → EApp e1 e2 -->* EApp e1' e2. | ||
| 16 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
| 17 | Lemma SEval_rtc ds e1 e1' : e1 -->* e1' → EEval ds e1 -->* EEval ds e1'. | ||
| 18 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
| 19 | |||
| 20 | Ltac inv_step := repeat | ||
| 21 | match goal with H : ?e --> _ |- _ => assert_fails (is_var e); inv H end. | ||
| 22 | |||
| 23 | Lemma step_det e d1 d2 : | ||
| 24 | e --> d1 → | ||
| 25 | e --> d2 → | ||
| 26 | d1 = d2. | ||
| 27 | Proof. | ||
| 28 | intros Hred1. revert d2. | ||
| 29 | induction Hred1; intros d2 Hred2; inv Hred2; try by inv_step; | ||
| 30 | f_equal; by apply IHHred1. | ||
| 31 | Qed. | ||
| 32 | |||
| 33 | End evallang. | ||
diff --git a/theories/evallang/tests.v b/theories/evallang/tests.v new file mode 100644 index 0000000..eaba8a0 --- /dev/null +++ b/theories/evallang/tests.v | |||
| @@ -0,0 +1,33 @@ | |||
| 1 | From mininix Require Export evallang.interp. | ||
| 2 | From stdpp Require Import options. | ||
| 3 | |||
| 4 | Import evallang. | ||
| 5 | |||
| 6 | Definition interp' (n : nat) (s : string) : res val := | ||
| 7 | interp n ∅ (EEval ∅ (EString s)). | ||
| 8 | |||
| 9 | Lemma test_1_a : interp' 1000 ("(x: x) ""s""") = mret (VString "s"). | ||
| 10 | Proof. by vm_compute. Qed. | ||
| 11 | Lemma test_1_b : interp' 1000 ("(""x"": x) ""s""") = mret (VString "s"). | ||
| 12 | Proof. by vm_compute. Qed. | ||
| 13 | Lemma test_1_c : interp' 1000 ("((y:y) ""x"": x) ""s""") = mret (VString "s"). | ||
| 14 | Proof. by vm_compute. Qed. | ||
| 15 | Lemma test_1_d : interp' 1000 ("(((y:y) ""x""): x) ""s""") = mret (VString "s"). | ||
| 16 | Proof. by vm_compute. Qed. | ||
| 17 | |||
| 18 | Lemma test_2 : interp' 1000 ("(x: y: eval! y) ""s"" ""x""") = mret (VString "s"). | ||
| 19 | Proof. by vm_compute. Qed. | ||
| 20 | |||
| 21 | Lemma test_3 : interp' 1000 ("eval! ""x: x"" ""s""") = mret (VString "s"). | ||
| 22 | Proof. by vm_compute. Qed. | ||
| 23 | |||
| 24 | Lemma test_4_a : | ||
| 25 | interp' 1000 ("(x: y: eval! y) ""s"" ""x""") = mret (VString "s"). | ||
| 26 | Proof. by vm_compute. Qed. | ||
| 27 | Lemma test_4_b : | ||
| 28 | interp' 1000 ("eval! ""(x: y: eval! y) \""s\"" \""x\""""") = mret (VString "s"). | ||
| 29 | Proof. by vm_compute. Qed. | ||
| 30 | |||
| 31 | Lemma test_5 : | ||
| 32 | interp' 1000 ("(x: y: eval! ""x: x"" (eval! y)) ""s"" ""x""") = mret (VString "s"). | ||
| 33 | Proof. by vm_compute. Qed. | ||