diff options
| author | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 | 
|---|---|---|
| committer | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 | 
| commit | 73df1945b31c0beee88cf4476df4ccd09d31403b (patch) | |
| tree | ed00db26b711442e643f38b66888a3df56e33ebd /matching.v | |
| download | mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.tar.gz mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.zip | |
Import Coq project
Diffstat (limited to 'matching.v')
| -rw-r--r-- | matching.v | 609 | 
1 files changed, 609 insertions, 0 deletions
| diff --git a/matching.v b/matching.v new file mode 100644 index 0000000..494bb92 --- /dev/null +++ b/matching.v | |||
| @@ -0,0 +1,609 @@ | |||
| 1 | Require Import Coq.Strings.Ascii. | ||
| 2 | Require Import Coq.Strings.String. | ||
| 3 | Require Import Coq.NArith.BinNat. | ||
| 4 | From stdpp Require Import fin_sets gmap option. | ||
| 5 | From mininix Require Import expr maptools. | ||
| 6 | |||
| 7 | Open Scope string_scope. | ||
| 8 | Open Scope N_scope. | ||
| 9 | Open Scope Z_scope. | ||
| 10 | Open Scope nat_scope. | ||
| 11 | |||
| 12 | Reserved Notation "bs '~' m '~>' brs" (at level 55). | ||
| 13 | |||
| 14 | Inductive matches_r : gmap string expr → matcher → gmap string b_rhs -> Prop := | ||
| 15 | | E_MatchEmpty strict : ∅ ~ M_Matcher ∅ strict ~> ∅ | ||
| 16 | | E_MatchAny bs : bs ~ M_Matcher ∅ false ~> ∅ | ||
| 17 | | E_MatchMandatory bs x e m bs' strict : | ||
| 18 | bs !! x = None → | ||
| 19 | m !! x = None → | ||
| 20 | delete x bs ~ M_Matcher m strict ~> bs' → | ||
| 21 | <[x := e]>bs ~ M_Matcher (<[x := M_Mandatory]>m) strict | ||
| 22 | ~> <[x := B_Nonrec e]>bs' | ||
| 23 | | E_MatchOptAvail bs x d e m bs' strict : | ||
| 24 | bs !! x = None → | ||
| 25 | m !! x = None → | ||
| 26 | delete x bs ~ M_Matcher m strict ~> bs' → | ||
| 27 | <[x := d]>bs ~ M_Matcher (<[x := M_Optional e]>m) strict | ||
| 28 | ~> <[x := B_Nonrec d]>bs' | ||
| 29 | | E_MatchOptDefault bs x e m bs' strict : | ||
| 30 | bs !! x = None → | ||
| 31 | m !! x = None → | ||
| 32 | bs ~ M_Matcher m strict ~> bs' → | ||
| 33 | bs ~ M_Matcher (<[x := M_Optional e]>m) strict ~> <[x := B_Rec e]>bs' | ||
| 34 | where "bs ~ m ~> brs" := (matches_r bs m brs). | ||
| 35 | |||
| 36 | Definition map_foldM `{Countable K} `{FinMap K M} `{MBind F} `{MRet F} {A B} | ||
| 37 | (f : K → A → B → F B) (init : B) (m : M A) : F B := | ||
| 38 | map_fold (λ i x acc, acc ≫= f i x) (mret init) m. | ||
| 39 | |||
| 40 | Definition matches_aux_f (x : string) (rhs : m_rhs) | ||
| 41 | (acc : option (gmap string expr * gmap string b_rhs)) := | ||
| 42 | acc ← acc; | ||
| 43 | let (bs, brs) := (acc : gmap string expr * gmap string b_rhs) in | ||
| 44 | match rhs with | ||
| 45 | | M_Mandatory => | ||
| 46 | e ← bs !! x; | ||
| 47 | Some (bs, <[x := B_Nonrec e]>brs) | ||
| 48 | | M_Optional e => | ||
| 49 | match bs !! x with | ||
| 50 | | Some e' => Some (bs, <[x := B_Nonrec e']>brs) | ||
| 51 | | None => Some (bs, <[x := B_Rec e]>brs) | ||
| 52 | end | ||
| 53 | end. | ||
| 54 | |||
| 55 | Definition matches_aux (ms : gmap string m_rhs) (bs : gmap string expr) : | ||
| 56 | option (gmap string expr * gmap string b_rhs) := | ||
| 57 | map_fold matches_aux_f (Some (bs, ∅)) ms. | ||
| 58 | |||
| 59 | Definition matches (m : matcher) (bs : gmap string expr) : | ||
| 60 | option (gmap string b_rhs) := | ||
| 61 | match m with | ||
| 62 | | M_Matcher ms strict => | ||
| 63 | guard (strict → dom bs ⊆ matcher_keys m);; | ||
| 64 | snd <$> matches_aux ms bs | ||
| 65 | end. | ||
| 66 | |||
| 67 | (* Copied from stdpp and changed so that the value types | ||
| 68 | of m1 and m2 are decoupled *) | ||
| 69 | Lemma dom_subseteq_size' `{FinMapDom K M D} {A B} (m1 : M A) (m2 : M B) : | ||
| 70 | dom m2 ⊆ dom m1 → size m2 ≤ size m1. | ||
| 71 | Proof. | ||
| 72 | revert m1. induction m2 as [|i x m2 ? IH] using map_ind; intros m1 Hdom. | ||
| 73 | { rewrite map_size_empty. lia. } | ||
| 74 | rewrite dom_insert in Hdom. | ||
| 75 | assert (i ∉ dom m2) by (by apply not_elem_of_dom). | ||
| 76 | assert (i ∈ dom m1) as [x' Hx']%elem_of_dom by set_solver. | ||
| 77 | rewrite <-(insert_delete m1 i x') by done. | ||
| 78 | rewrite !map_size_insert_None, <-Nat.succ_le_mono | ||
| 79 | by (by rewrite ?lookup_delete). | ||
| 80 | apply IH. rewrite dom_delete. set_solver. | ||
| 81 | Qed. | ||
| 82 | |||
| 83 | Lemma dom_empty_subset `{Countable K} `{FinMapDom K M D} {A B} (m : M A) : | ||
| 84 | dom m ⊆ dom (∅ : M B) → m = ∅. | ||
| 85 | Proof. | ||
| 86 | intros Hdom. | ||
| 87 | apply dom_subseteq_size' in Hdom. | ||
| 88 | rewrite map_size_empty in Hdom. | ||
| 89 | inv Hdom. | ||
| 90 | by apply map_size_empty_inv. | ||
| 91 | Qed. | ||
| 92 | |||
| 93 | Lemma matches_aux_empty bs : matches_aux ∅ bs = Some (bs, ∅). | ||
| 94 | Proof. done. Qed. | ||
| 95 | |||
| 96 | Lemma matches_aux_f_comm (x : string) (rhs : m_rhs) (m : gmap string m_rhs) | ||
| 97 | (y z : string) (rhs1 rhs2 : m_rhs) | ||
| 98 | (acc : option (gmap string expr * gmap string b_rhs)) : | ||
| 99 | m !! x = None → y ≠ z → | ||
| 100 | <[x:=rhs]> m !! y = Some rhs1 → | ||
| 101 | <[x:=rhs]> m !! z = Some rhs2 → | ||
| 102 | matches_aux_f y rhs1 (matches_aux_f z rhs2 acc) = | ||
| 103 | matches_aux_f z rhs2 (matches_aux_f y rhs1 acc). | ||
| 104 | Proof. | ||
| 105 | intros Hxm Hyz Hym Hzm. | ||
| 106 | unfold matches_aux_f. | ||
| 107 | destruct acc. | ||
| 108 | repeat (simplify_option_eq || done || | ||
| 109 | destruct (g !! y) eqn:Egy || destruct (g !! z) eqn:Egz || | ||
| 110 | case_match || rewrite insert_commute). done. | ||
| 111 | Qed. | ||
| 112 | |||
| 113 | Lemma matches_aux_insert m bs x rhs : | ||
| 114 | m !! x = None → | ||
| 115 | matches_aux (<[x := rhs]>m) bs = matches_aux_f x rhs (matches_aux m bs). | ||
| 116 | Proof. | ||
| 117 | intros Hnotin. unfold matches_aux. | ||
| 118 | rewrite map_fold_insert_L; try done. | ||
| 119 | clear bs. | ||
| 120 | intros y z rhs1 rhs2 acc Hyz Hym Hzm. | ||
| 121 | by eapply matches_aux_f_comm. | ||
| 122 | Qed. | ||
| 123 | |||
| 124 | Lemma matches_aux_bs m bs bs' brs : | ||
| 125 | matches_aux m bs = Some (bs', brs) → bs = bs'. | ||
| 126 | Proof. | ||
| 127 | revert bs bs' brs. | ||
| 128 | induction m using map_ind; intros bs bs' brs Hmatches. | ||
| 129 | - rewrite matches_aux_empty in Hmatches. congruence. | ||
| 130 | - rewrite matches_aux_insert in Hmatches; try done. | ||
| 131 | unfold matches_aux_f in Hmatches. | ||
| 132 | simplify_option_eq. | ||
| 133 | destruct H0. | ||
| 134 | case_match; try case_match; | ||
| 135 | simplify_option_eq; by apply IHm with (brs := g0). | ||
| 136 | Qed. | ||
| 137 | |||
| 138 | Lemma matches_aux_impl m bs brs x e : | ||
| 139 | m !! x = None → | ||
| 140 | bs !! x = None → | ||
| 141 | matches_aux m (<[x := e]>bs) = Some (<[x := e]>bs, brs) → | ||
| 142 | matches_aux m bs = Some (bs, brs). | ||
| 143 | Proof. | ||
| 144 | intros Hxm Hxbs Hmatches. | ||
| 145 | revert bs brs Hxm Hxbs Hmatches. | ||
| 146 | induction m using map_ind; intros bs brs Hxm Hxbs Hmatches. | ||
| 147 | - rewrite matches_aux_empty. | ||
| 148 | rewrite matches_aux_empty in Hmatches. | ||
| 149 | by simplify_option_eq. | ||
| 150 | - rewrite matches_aux_insert in Hmatches by done. | ||
| 151 | rewrite matches_aux_insert by done. | ||
| 152 | apply lookup_insert_None in Hxm as [Hxm1 Hxm2]. | ||
| 153 | unfold matches_aux_f in Hmatches. simplify_option_eq. | ||
| 154 | destruct H0. case_match. | ||
| 155 | + simplify_option_eq. | ||
| 156 | rewrite lookup_insert_ne in Heqo0 by done. | ||
| 157 | pose proof (IHm bs g0 Hxm1 Hxbs Heqo). | ||
| 158 | rewrite H1. by simplify_option_eq. | ||
| 159 | + case_match; simplify_option_eq; | ||
| 160 | rewrite lookup_insert_ne in H1 by done; | ||
| 161 | pose proof (IHm bs g0 Hxm1 Hxbs Heqo); | ||
| 162 | rewrite H0; by simplify_option_eq. | ||
| 163 | Qed. | ||
| 164 | |||
| 165 | Lemma matches_aux_impl_2 m bs brs x e : | ||
| 166 | m !! x = None → | ||
| 167 | bs !! x = None → | ||
| 168 | matches_aux m bs = Some (bs, brs) → | ||
| 169 | matches_aux m (<[x := e]>bs) = Some (<[x := e]>bs, brs). | ||
| 170 | Proof. | ||
| 171 | intros Hxm Hxbs Hmatches. | ||
| 172 | revert bs brs Hxm Hxbs Hmatches. | ||
| 173 | induction m using map_ind; intros bs brs Hxm Hxbs Hmatches. | ||
| 174 | - rewrite matches_aux_empty. | ||
| 175 | rewrite matches_aux_empty in Hmatches. | ||
| 176 | by simplify_option_eq. | ||
| 177 | - rewrite matches_aux_insert in Hmatches by done. | ||
| 178 | rewrite matches_aux_insert by done. | ||
| 179 | apply lookup_insert_None in Hxm as [Hxm1 Hxm2]. | ||
| 180 | unfold matches_aux_f in Hmatches. simplify_option_eq. | ||
| 181 | destruct H0. case_match. | ||
| 182 | + simplify_option_eq. | ||
| 183 | rewrite <-lookup_insert_ne with (i := x) (x := e) in Heqo0 by done. | ||
| 184 | pose proof (IHm bs g0 Hxm1 Hxbs Heqo). | ||
| 185 | rewrite H1. by simplify_option_eq. | ||
| 186 | + case_match; simplify_option_eq; | ||
| 187 | rewrite <-lookup_insert_ne with (i := x) (x := e) in H1 by done; | ||
| 188 | pose proof (IHm bs g0 Hxm1 Hxbs Heqo); | ||
| 189 | rewrite H0; by simplify_option_eq. | ||
| 190 | Qed. | ||
| 191 | |||
| 192 | Lemma matches_aux_dom m bs brs : | ||
| 193 | matches_aux m bs = Some (bs, brs) → dom m = dom brs. | ||
| 194 | Proof. | ||
| 195 | revert bs brs. | ||
| 196 | induction m using map_ind; intros bs brs Hmatches. | ||
| 197 | - rewrite matches_aux_empty in Hmatches. by simplify_eq. | ||
| 198 | - rewrite matches_aux_insert in Hmatches by done. | ||
| 199 | unfold matches_aux_f in Hmatches. simplify_option_eq. destruct H0. | ||
| 200 | case_match; try case_match; | ||
| 201 | simplify_option_eq; apply IHm in Heqo; set_solver. | ||
| 202 | Qed. | ||
| 203 | |||
| 204 | Lemma matches_aux_inv m bs brs x e : | ||
| 205 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 206 | matches_aux (<[x := M_Mandatory]>m) (<[x := e]>bs) = | ||
| 207 | Some (<[x := e]>bs, <[x := B_Nonrec e]>brs) → | ||
| 208 | matches_aux m bs = Some (bs, brs). | ||
| 209 | Proof. | ||
| 210 | intros Hxm Hxbs Hxbrs Hmatches. | ||
| 211 | rewrite matches_aux_insert in Hmatches by done. | ||
| 212 | unfold matches_aux_f in Hmatches. simplify_option_eq. | ||
| 213 | destruct H. simplify_option_eq. | ||
| 214 | rewrite lookup_insert in Heqo0. simplify_option_eq. | ||
| 215 | apply matches_aux_impl in Heqo; try done. | ||
| 216 | simplify_option_eq. | ||
| 217 | apply matches_aux_dom in Heqo as Hdom. | ||
| 218 | rewrite Heqo. do 2 f_equal. | ||
| 219 | assert (g0 !! x = None). | ||
| 220 | { apply not_elem_of_dom in Hxm. | ||
| 221 | rewrite Hdom in Hxm. | ||
| 222 | by apply not_elem_of_dom. } | ||
| 223 | replace g0 with (delete x (<[x:=B_Nonrec H]> g0)); | ||
| 224 | try by rewrite delete_insert. | ||
| 225 | replace brs with (delete x (<[x:=B_Nonrec H]> brs)); | ||
| 226 | try by rewrite delete_insert. | ||
| 227 | by rewrite H1. | ||
| 228 | Qed. | ||
| 229 | |||
| 230 | Lemma disjoint_union_subseteq_l `{FinSet A C} `{!LeibnizEquiv C} (X Y Z : C) : | ||
| 231 | X ## Y → X ## Z → X ∪ Y ⊆ X ∪ Z → Y ⊆ Z. | ||
| 232 | Proof. | ||
| 233 | revert Y Z. | ||
| 234 | induction X using set_ind_L; intros Y Z Hdisj1 Hdisj2 Hsubs. | ||
| 235 | - by do 2 rewrite union_empty_l_L in Hsubs. | ||
| 236 | - do 2 rewrite <-union_assoc_L in Hsubs. | ||
| 237 | apply IHX; set_solver. | ||
| 238 | Qed. | ||
| 239 | |||
| 240 | Lemma singleton_notin_union_disjoint `{FinMapDom K M D} {A} (m : M A) (i : K) : | ||
| 241 | m !! i = None → | ||
| 242 | {[i]} ## dom m. | ||
| 243 | Proof. | ||
| 244 | intros Hlookup. apply disjoint_singleton_l. | ||
| 245 | by apply not_elem_of_dom in Hlookup. | ||
| 246 | Qed. | ||
| 247 | |||
| 248 | Lemma matches_step bs brs m strict x : | ||
| 249 | matches (M_Matcher (<[x := M_Mandatory]>m) strict) bs = Some brs → | ||
| 250 | ∃ e, bs !! x = Some e ∧ brs !! x = Some (B_Nonrec e). | ||
| 251 | Proof. | ||
| 252 | intros Hmatches. | ||
| 253 | destruct (decide (is_Some (bs !! x))). | ||
| 254 | - destruct i. exists x0. split; [done|]. | ||
| 255 | unfold matches in Hmatches. | ||
| 256 | destruct strict; simplify_option_eq. | ||
| 257 | + replace (<[x := M_Mandatory]>m) with (<[x := M_Mandatory]>(delete x m)) | ||
| 258 | in Heqo0; try by rewrite insert_delete_insert. | ||
| 259 | rewrite matches_aux_insert in Heqo0 by apply lookup_delete. | ||
| 260 | unfold matches_aux_f in Heqo0. simplify_option_eq. | ||
| 261 | destruct H3. simplify_option_eq. | ||
| 262 | apply matches_aux_bs in Heqo as Hdom. simplify_option_eq. | ||
| 263 | apply lookup_insert. | ||
| 264 | + replace (<[x := M_Mandatory]>m) with (<[x := M_Mandatory]>(delete x m)) | ||
| 265 | in Heqo; try by rewrite insert_delete_insert. | ||
| 266 | rewrite matches_aux_insert in Heqo by apply lookup_delete. | ||
| 267 | unfold matches_aux_f in Heqo. simplify_option_eq. | ||
| 268 | destruct H1. simplify_option_eq. | ||
| 269 | apply matches_aux_bs in Heqo0 as Hdom. simplify_option_eq. | ||
| 270 | apply lookup_insert. | ||
| 271 | - unfold matches in Hmatches. | ||
| 272 | destruct strict; simplify_option_eq. | ||
| 273 | + replace (<[x := M_Mandatory]>m) with (<[x := M_Mandatory]>(delete x m)) | ||
| 274 | in Heqo0; try by rewrite insert_delete_insert. | ||
| 275 | rewrite matches_aux_insert in Heqo0 by apply lookup_delete. | ||
| 276 | unfold matches_aux_f in Heqo0. simplify_option_eq. | ||
| 277 | destruct H2. simplify_option_eq. | ||
| 278 | apply matches_aux_bs in Heqo as Hdom. simplify_option_eq. | ||
| 279 | rewrite Heqo1 in n. | ||
| 280 | exfalso. by apply n. | ||
| 281 | + replace (<[x := M_Mandatory]>m) with (<[x := M_Mandatory]>(delete x m)) | ||
| 282 | in Heqo; try by rewrite insert_delete_insert. | ||
| 283 | rewrite matches_aux_insert in Heqo by apply lookup_delete. | ||
| 284 | unfold matches_aux_f in Heqo. simplify_option_eq. | ||
| 285 | destruct H0. simplify_option_eq. | ||
| 286 | apply matches_aux_bs in Heqo0 as Hdom. simplify_option_eq. | ||
| 287 | rewrite Heqo1 in n. | ||
| 288 | exfalso. by apply n. | ||
| 289 | Qed. | ||
| 290 | |||
| 291 | Lemma matches_step' bs brs m strict x : | ||
| 292 | matches (M_Matcher (<[x := M_Mandatory]>m) strict) bs = Some brs → | ||
| 293 | ∃ e bs' brs', | ||
| 294 | bs' !! x = None ∧ bs = <[x := e]>bs' ∧ | ||
| 295 | brs' !! x = None ∧ brs = <[x := B_Nonrec e]>brs'. | ||
| 296 | Proof. | ||
| 297 | intros Hmatches. | ||
| 298 | apply matches_step in Hmatches as He. | ||
| 299 | destruct He as [e [He1 He2]]. | ||
| 300 | exists e, (delete x bs), (delete x brs). | ||
| 301 | split_and!; by apply lookup_delete || rewrite insert_delete. | ||
| 302 | Qed. | ||
| 303 | |||
| 304 | Lemma matches_opt_step bs brs m strict x d : | ||
| 305 | matches (M_Matcher (<[x := M_Optional d]>m) strict) bs = Some brs → | ||
| 306 | (∃ e, bs !! x = Some e ∧ brs !! x = Some (B_Nonrec e)) ∨ | ||
| 307 | bs !! x = None ∧ brs !! x = Some (B_Rec d). | ||
| 308 | Proof. | ||
| 309 | intros Hmatches. | ||
| 310 | destruct (decide (is_Some (bs !! x))). | ||
| 311 | - destruct i. left. exists x0. split; [done|]. | ||
| 312 | unfold matches in Hmatches. | ||
| 313 | destruct strict; simplify_option_eq. | ||
| 314 | + replace (<[x := M_Optional d]>m) with (<[x := M_Optional d]>(delete x m)) | ||
| 315 | in Heqo0; try by rewrite insert_delete_insert. | ||
| 316 | rewrite matches_aux_insert in Heqo0 by apply lookup_delete. | ||
| 317 | unfold matches_aux_f in Heqo0. simplify_option_eq. | ||
| 318 | destruct H3. simplify_option_eq. | ||
| 319 | apply matches_aux_bs in Heqo as Hdom. simplify_option_eq. | ||
| 320 | apply lookup_insert. | ||
| 321 | + replace (<[x := M_Optional d]>m) with (<[x := M_Optional d]>(delete x m)) | ||
| 322 | in Heqo; try by rewrite insert_delete_insert. | ||
| 323 | rewrite matches_aux_insert in Heqo by apply lookup_delete. | ||
| 324 | unfold matches_aux_f in Heqo. simplify_option_eq. | ||
| 325 | destruct H1. simplify_option_eq. | ||
| 326 | apply matches_aux_bs in Heqo0 as Hdom. simplify_option_eq. | ||
| 327 | apply lookup_insert. | ||
| 328 | - unfold matches in Hmatches. | ||
| 329 | destruct strict; simplify_option_eq. | ||
| 330 | + right. apply eq_None_not_Some in n. split; [done|]. | ||
| 331 | destruct H0. simplify_option_eq. | ||
| 332 | apply matches_aux_bs in Heqo0 as Hbs. subst. | ||
| 333 | replace (<[x := M_Optional d]>m) with (<[x := M_Optional d]>(delete x m)) | ||
| 334 | in Heqo0; try by rewrite insert_delete_insert. | ||
| 335 | rewrite matches_aux_insert in Heqo0 by apply lookup_delete. | ||
| 336 | unfold matches_aux_f in Heqo0. simplify_option_eq. | ||
| 337 | destruct H0. simplify_option_eq. | ||
| 338 | apply matches_aux_bs in Heqo as Hdom. simplify_option_eq. | ||
| 339 | apply lookup_insert. | ||
| 340 | + right. apply eq_None_not_Some in n. split; [done|]. | ||
| 341 | destruct H. simplify_option_eq. | ||
| 342 | apply matches_aux_bs in Heqo as Hbs. subst. | ||
| 343 | replace (<[x := M_Optional d]>m) with (<[x := M_Optional d]>(delete x m)) | ||
| 344 | in Heqo; try by rewrite insert_delete_insert. | ||
| 345 | rewrite matches_aux_insert in Heqo by apply lookup_delete. | ||
| 346 | unfold matches_aux_f in Heqo. simplify_option_eq. | ||
| 347 | destruct H. simplify_option_eq. | ||
| 348 | apply matches_aux_bs in Heqo0 as Hdom. simplify_option_eq. | ||
| 349 | apply lookup_insert. | ||
| 350 | Qed. | ||
| 351 | |||
| 352 | Lemma matches_opt_step' bs brs m strict x d : | ||
| 353 | matches (M_Matcher (<[x := M_Optional d]>m) strict) bs = Some brs → | ||
| 354 | (∃ e bs' brs', | ||
| 355 | bs' !! x = None ∧ bs = <[x := e]>bs' ∧ | ||
| 356 | brs' !! x = None ∧ brs = <[x := B_Nonrec e]>brs') ∨ | ||
| 357 | (∃ brs', | ||
| 358 | bs !! x = None ∧ brs' !! x = None ∧ | ||
| 359 | brs = <[x := B_Rec d]>brs'). | ||
| 360 | Proof. | ||
| 361 | intros Hmatches. | ||
| 362 | apply matches_opt_step in Hmatches as He. | ||
| 363 | destruct He as [He|He]. | ||
| 364 | - destruct He as [e [He1 He2]]. left. | ||
| 365 | exists e, (delete x bs), (delete x brs). | ||
| 366 | split; try split; try split. | ||
| 367 | + apply lookup_delete. | ||
| 368 | + by rewrite insert_delete. | ||
| 369 | + apply lookup_delete. | ||
| 370 | + by rewrite insert_delete. | ||
| 371 | - destruct He as [He1 He2]. right. | ||
| 372 | exists (delete x brs). split; try split; try done. | ||
| 373 | + apply lookup_delete. | ||
| 374 | + by rewrite insert_delete. | ||
| 375 | Qed. | ||
| 376 | |||
| 377 | Lemma matches_inv m bs brs strict x e : | ||
| 378 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 379 | matches (M_Matcher (<[x := M_Mandatory]>m) strict) (<[x := e]>bs) = | ||
| 380 | Some (<[x := B_Nonrec e]>brs) → | ||
| 381 | matches (M_Matcher m strict) bs = Some brs. | ||
| 382 | Proof. | ||
| 383 | intros Hxm Hxbs Hxbrs Hmatch. | ||
| 384 | destruct strict. | ||
| 385 | - simplify_option_eq. | ||
| 386 | + destruct H0. apply matches_aux_bs in Heqo0 as Hbs. | ||
| 387 | simplify_option_eq. by erewrite matches_aux_inv. | ||
| 388 | + exfalso. apply H2. | ||
| 389 | repeat rewrite dom_insert in H1. | ||
| 390 | assert ({[x]} ## dom bs). { by apply singleton_notin_union_disjoint. } | ||
| 391 | assert ({[x]} ## dom m). { by apply singleton_notin_union_disjoint. } | ||
| 392 | by apply disjoint_union_subseteq_l with (X := {[x]}). | ||
| 393 | - simplify_option_eq. | ||
| 394 | destruct H. apply matches_aux_bs in Heqo as Hbs. | ||
| 395 | simplify_option_eq. by erewrite matches_aux_inv. | ||
| 396 | Qed. | ||
| 397 | |||
| 398 | Lemma matches_aux_avail_inv m bs brs x d e : | ||
| 399 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 400 | matches_aux (<[x := M_Optional d]>m) (<[x := e]>bs) = | ||
| 401 | Some (<[x := e]>bs, <[x := B_Nonrec e]>brs) → | ||
| 402 | matches_aux m bs = Some (bs, brs). | ||
| 403 | Proof. | ||
| 404 | intros Hxm Hxbs Hxbrs Hmatches. | ||
| 405 | rewrite matches_aux_insert in Hmatches by done. | ||
| 406 | unfold matches_aux_f in Hmatches. simplify_option_eq. | ||
| 407 | destruct H. simplify_option_eq. | ||
| 408 | apply matches_aux_bs in Heqo as Hbs. subst. | ||
| 409 | rewrite lookup_insert in Hmatches. simplify_option_eq. | ||
| 410 | apply matches_aux_impl in Heqo; try done. | ||
| 411 | simplify_option_eq. | ||
| 412 | apply matches_aux_dom in Heqo as Hdom. | ||
| 413 | rewrite Heqo. do 2 f_equal. | ||
| 414 | assert (g0 !! x = None). | ||
| 415 | { apply not_elem_of_dom in Hxm. | ||
| 416 | rewrite Hdom in Hxm. | ||
| 417 | by apply not_elem_of_dom. } | ||
| 418 | replace g0 with (delete x (<[x:=B_Nonrec e]> g0)); | ||
| 419 | try by rewrite delete_insert. | ||
| 420 | replace brs with (delete x (<[x:=B_Nonrec e]> brs)); | ||
| 421 | try by rewrite delete_insert. | ||
| 422 | by rewrite Hmatches. | ||
| 423 | Qed. | ||
| 424 | |||
| 425 | Lemma matches_avail_inv m bs brs strict x d e : | ||
| 426 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 427 | matches (M_Matcher (<[x := M_Optional d]>m) strict) (<[x := e]>bs) = | ||
| 428 | Some (<[x := B_Nonrec e]>brs) → | ||
| 429 | matches (M_Matcher m strict) bs = Some brs. | ||
| 430 | Proof. | ||
| 431 | intros Hxm Hxbs Hxbrs Hmatch. | ||
| 432 | destruct strict. | ||
| 433 | - simplify_option_eq. | ||
| 434 | + destruct H0. apply matches_aux_bs in Heqo0 as Hbs. | ||
| 435 | simplify_option_eq. by erewrite matches_aux_avail_inv. | ||
| 436 | + exfalso. apply H2. | ||
| 437 | repeat rewrite dom_insert in H1. | ||
| 438 | assert ({[x]} ## dom bs). { by apply singleton_notin_union_disjoint. } | ||
| 439 | assert ({[x]} ## dom m). { by apply singleton_notin_union_disjoint. } | ||
| 440 | by apply disjoint_union_subseteq_l with (X := {[x]}). | ||
| 441 | - simplify_option_eq. | ||
| 442 | destruct H. apply matches_aux_bs in Heqo as Hbs. | ||
| 443 | simplify_option_eq. by erewrite matches_aux_avail_inv. | ||
| 444 | Qed. | ||
| 445 | |||
| 446 | Lemma matches_aux_default_inv m bs brs x e : | ||
| 447 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 448 | matches_aux (<[x := M_Optional e]>m) bs = | ||
| 449 | Some (bs, <[x := B_Rec e]>brs) → | ||
| 450 | matches_aux m bs = Some (bs, brs). | ||
| 451 | Proof. | ||
| 452 | intros Hxm Hxbs Hxbrs Hmatches. | ||
| 453 | rewrite matches_aux_insert in Hmatches by done. | ||
| 454 | unfold matches_aux_f in Hmatches. simplify_option_eq. | ||
| 455 | destruct H. simplify_option_eq. | ||
| 456 | apply matches_aux_bs in Heqo as Hbs. subst. | ||
| 457 | rewrite Hxbs in Hmatches. simplify_option_eq. | ||
| 458 | apply matches_aux_dom in Heqo as Hdom. | ||
| 459 | do 2 f_equal. | ||
| 460 | assert (g0 !! x = None). | ||
| 461 | { apply not_elem_of_dom in Hxm. | ||
| 462 | rewrite Hdom in Hxm. | ||
| 463 | by apply not_elem_of_dom. } | ||
| 464 | replace g0 with (delete x (<[x:=B_Rec e]> g0)); | ||
| 465 | try by rewrite delete_insert. | ||
| 466 | replace brs with (delete x (<[x:=B_Rec e]> brs)); | ||
| 467 | try by rewrite delete_insert. | ||
| 468 | by rewrite Hmatches. | ||
| 469 | Qed. | ||
| 470 | |||
| 471 | Lemma matches_default_inv m bs brs strict x e : | ||
| 472 | m !! x = None → bs !! x = None → brs !! x = None → | ||
| 473 | matches (M_Matcher (<[x := M_Optional e]>m) strict) bs = | ||
| 474 | Some (<[x := B_Rec e]>brs) → | ||
| 475 | matches (M_Matcher m strict) bs = Some brs. | ||
| 476 | Proof. | ||
| 477 | intros Hxm Hxbs Hxbrs Hmatch. | ||
| 478 | destruct strict. | ||
| 479 | - simplify_option_eq. | ||
| 480 | + destruct H0. apply matches_aux_bs in Heqo0 as Hbs. | ||
| 481 | simplify_option_eq. by erewrite matches_aux_default_inv. | ||
| 482 | + exfalso. apply H2. | ||
| 483 | rewrite dom_insert in H1. | ||
| 484 | assert ({[x]} ## dom bs). { by apply singleton_notin_union_disjoint. } | ||
| 485 | assert ({[x]} ## dom m). { by apply singleton_notin_union_disjoint. } | ||
| 486 | apply disjoint_union_subseteq_l with (X := {[x]}); set_solver. | ||
| 487 | - simplify_option_eq. | ||
| 488 | destruct H. apply matches_aux_bs in Heqo as Hbs. | ||
| 489 | simplify_option_eq. by erewrite matches_aux_default_inv. | ||
| 490 | Qed. | ||
| 491 | |||
| 492 | Theorem matches_sound m bs brs : matches m bs = Some brs → bs ~ m ~> brs. | ||
| 493 | Proof. | ||
| 494 | intros Hmatches. | ||
| 495 | destruct m. | ||
| 496 | revert strict bs brs Hmatches. | ||
| 497 | induction ms using map_ind; intros strict bs brs Hmatches. | ||
| 498 | - destruct strict; simplify_option_eq. | ||
| 499 | + apply dom_empty_subset in H0. simplify_eq. | ||
| 500 | constructor. | ||
| 501 | + constructor. | ||
| 502 | - destruct x. | ||
| 503 | + apply matches_step' in Hmatches as He. | ||
| 504 | destruct He as [e [bs' [brs' [Hbs'1 [Hbs'2 [Hbrs'1 Hbrs'2]]]]]]. | ||
| 505 | subst. constructor; try done. | ||
| 506 | rewrite delete_notin by done. | ||
| 507 | apply IHms. | ||
| 508 | by apply matches_inv in Hmatches. | ||
| 509 | + apply matches_opt_step' in Hmatches as He. destruct He as [He|He]. | ||
| 510 | * destruct He as [d [bs' [brs' [Hibs' [Hbs' [Hibrs' Hbrs']]]]]]. | ||
| 511 | subst. constructor; try done. | ||
| 512 | rewrite delete_notin by done. | ||
| 513 | apply IHms. | ||
| 514 | by apply matches_avail_inv in Hmatches. | ||
| 515 | * destruct He as [brs' [Hibs [Hibrs' Hbrs']]]. | ||
| 516 | subst. constructor; try done. | ||
| 517 | apply IHms. | ||
| 518 | by apply matches_default_inv in Hmatches. | ||
| 519 | Qed. | ||
| 520 | |||
| 521 | Theorem matches_complete m bs brs : bs ~ m ~> brs → matches m bs = Some brs. | ||
| 522 | Proof. | ||
| 523 | intros Hbs. | ||
| 524 | induction Hbs. | ||
| 525 | - unfold matches. by simplify_option_eq. | ||
| 526 | - unfold matches. by simplify_option_eq. | ||
| 527 | - unfold matches in *. destruct strict. | ||
| 528 | + simplify_option_eq. | ||
| 529 | * simplify_option_eq. destruct H2. simplify_option_eq. | ||
| 530 | apply fmap_Some. exists (<[x := e]>bs, <[x := B_Nonrec e]>g0). | ||
| 531 | split; [|done]. | ||
| 532 | rewrite matches_aux_insert by done. | ||
| 533 | apply matches_aux_bs in Heqo0 as Hbs'. simplify_option_eq. | ||
| 534 | apply matches_aux_impl_2 with (x := x) (e := e) in Heqo0; | ||
| 535 | [| done | apply lookup_delete]. | ||
| 536 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 537 | unfold matches_aux_f. | ||
| 538 | simplify_option_eq. apply bind_Some. exists e. | ||
| 539 | split; [apply lookup_insert | done]. | ||
| 540 | * exfalso. apply H4. | ||
| 541 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 542 | apply map_dom_delete_insert_subseteq_L. | ||
| 543 | set_solver. | ||
| 544 | + simplify_option_eq. destruct H1. simplify_option_eq. | ||
| 545 | apply fmap_Some. exists (<[x := e]>bs, <[x := B_Nonrec e]>g0). | ||
| 546 | split; [|done]. | ||
| 547 | rewrite matches_aux_insert by done. | ||
| 548 | apply matches_aux_bs in Heqo as Hbs'. simplify_option_eq. | ||
| 549 | apply matches_aux_impl_2 with (x := x) (e := e) in Heqo; | ||
| 550 | [| done | apply lookup_delete]. | ||
| 551 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 552 | unfold matches_aux_f. | ||
| 553 | simplify_option_eq. apply bind_Some. exists e. | ||
| 554 | split; [apply lookup_insert | done]. | ||
| 555 | - unfold matches in *. destruct strict. | ||
| 556 | + simplify_option_eq. | ||
| 557 | * simplify_option_eq. destruct H2. simplify_option_eq. | ||
| 558 | apply fmap_Some. exists (<[x := d]>bs, <[x := B_Nonrec d]>g0). | ||
| 559 | split; [|done]. | ||
| 560 | rewrite matches_aux_insert by done. | ||
| 561 | apply matches_aux_bs in Heqo0 as Hbs'. simplify_option_eq. | ||
| 562 | apply matches_aux_impl_2 with (x := x) (e := d) in Heqo0; | ||
| 563 | [| done | apply lookup_delete]. | ||
| 564 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 565 | unfold matches_aux_f. | ||
| 566 | simplify_option_eq. case_match. | ||
| 567 | -- rewrite lookup_insert in H2. congruence. | ||
| 568 | -- by rewrite lookup_insert in H2. | ||
| 569 | * exfalso. apply H4. | ||
| 570 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 571 | apply map_dom_delete_insert_subseteq_L. | ||
| 572 | set_solver. | ||
| 573 | + simplify_option_eq. destruct H1. simplify_option_eq. | ||
| 574 | apply fmap_Some. exists (<[x := d]>bs, <[x := B_Nonrec d]>g0). | ||
| 575 | split; [|done]. | ||
| 576 | rewrite matches_aux_insert by done. | ||
| 577 | apply matches_aux_bs in Heqo as Hbs'. simplify_option_eq. | ||
| 578 | apply matches_aux_impl_2 with (x := x) (e := d) in Heqo; | ||
| 579 | [| done | apply lookup_delete]. | ||
| 580 | replace bs with (delete x bs); try by apply delete_notin. | ||
| 581 | unfold matches_aux_f. | ||
| 582 | simplify_option_eq. case_match. | ||
| 583 | * rewrite lookup_insert in H1. congruence. | ||
| 584 | * by rewrite lookup_insert in H1. | ||
| 585 | - unfold matches in *. destruct strict. | ||
| 586 | + simplify_option_eq. | ||
| 587 | * simplify_option_eq. destruct H2. simplify_option_eq. | ||
| 588 | apply fmap_Some. exists (bs, <[x := B_Rec e]>g0). | ||
| 589 | split; [|done]. | ||
| 590 | rewrite matches_aux_insert by done. | ||
| 591 | apply matches_aux_bs in Heqo0 as Hbs'. simplify_option_eq. | ||
| 592 | unfold matches_aux_f. | ||
| 593 | by simplify_option_eq. | ||
| 594 | * exfalso. apply H4. set_solver. | ||
| 595 | + simplify_option_eq. destruct H1. simplify_option_eq. | ||
| 596 | apply fmap_Some. exists (bs, <[x := B_Rec e]>g0). | ||
| 597 | split; [|done]. | ||
| 598 | rewrite matches_aux_insert by done. | ||
| 599 | apply matches_aux_bs in Heqo as Hbs'. simplify_option_eq. | ||
| 600 | unfold matches_aux_f. | ||
| 601 | by simplify_option_eq. | ||
| 602 | Qed. | ||
| 603 | |||
| 604 | Theorem matches_correct m bs brs : matches m bs = Some brs ↔ bs ~ m ~> brs. | ||
| 605 | Proof. split; [apply matches_sound | apply matches_complete]. Qed. | ||
| 606 | |||
| 607 | Theorem matches_deterministic m bs brs1 brs2 : | ||
| 608 | bs ~ m ~> brs1 → bs ~ m ~> brs2 → brs1 = brs2. | ||
| 609 | Proof. intros H1 H2. apply matches_correct in H1, H2. congruence. Qed. | ||