├── .gitignore ├── CHANGES.md ├── Cardinal.v ├── FSetList.v ├── Lib_FinSet.v ├── Lib_FinSetImpl.v ├── Lib_ListFacts.v ├── Lib_ListFactsMore.v ├── Lib_MyFSetInterface.v ├── Lib_Tactic.v ├── ML_SP_Definitions.v ├── ML_SP_Domain.v ├── ML_SP_Eval.v ├── ML_SP_Extraction.v ├── ML_SP_Inference.v ├── ML_SP_Inference_wf.v ├── ML_SP_Infrastructure.v ├── ML_SP_Rename.v ├── ML_SP_SoundDomain.v ├── ML_SP_Soundness.v ├── ML_SP_Unify.v ├── ML_SP_Unify_wf.v ├── Metatheory.v ├── Metatheory_Env.v ├── Metatheory_Fresh.v ├── Metatheory_SP.v ├── Metatheory_Var.v ├── Omega.v ├── README.md ├── _CoqProject ├── build.sh ├── mkmakefile.sh ├── times.txt ├── typinf.ml ├── typinf.mli └── use_typinf.ml /.gitignore: -------------------------------------------------------------------------------- 1 | *.vo 2 | *.glob 3 | .*.aux 4 | .coqdeps.d 5 | Makefile 6 | Makefile.conf 7 | -------------------------------------------------------------------------------- /CHANGES.md: -------------------------------------------------------------------------------- 1 | # Certified Interpreter April 2020 2 | 3 | 2020-04-08 [Jacques] 4 | * update for Coq 8.11 5 | 6 | 2019-11-18 [Xuanrui] 7 | * remove some warnings 8 | 9 | # Certified Interpreter November 2019 10 | 11 | 2019-11-15 [Jacques] 12 | * update for Coq 8.10 13 | 14 | 2013-01-17 [Jacques] 15 | * update for Coq 8.4 16 | 17 | 2012-06-18 [Jacques] 18 | * update for Coq 8.3 19 | 20 | 2010-10-25 [Jacques] 21 | * release for APLAS paper 22 | -------------------------------------------------------------------------------- /Cardinal.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Cardinality lemmas * 3 | * Jacques Garrigue, July 2008 * 4 | ***************************************************************************) 5 | 6 | Require Import List SetoidList Arith Omega Metatheory. 7 | Set Implicit Arguments. 8 | 9 | Lemma elements_empty : forall L, 10 | S.elements L = nil -> L = {}. 11 | Proof. 12 | intros. 13 | apply eq_ext. 14 | intro. split; intro. 15 | use (S.elements_1 H0). 16 | rewrite H in H1. 17 | inversion H1. 18 | elim (in_empty H0). 19 | Qed. 20 | 21 | Lemma diff_empty_r : forall L, S.diff L {} = L. 22 | intros. 23 | apply eq_ext; intro; split; intro. 24 | apply* S.diff_1. 25 | apply* S.diff_3. 26 | Qed. 27 | 28 | Definition sort_lt_all := InfA_alt _ _ _. 29 | (* Check sort_lt_all. *) 30 | 31 | Lemma sort_lt_notin : forall a l0, 32 | sort S.E.lt l0 -> 33 | lelistA S.E.lt a l0 -> 34 | ~ InA S.E.eq a l0. 35 | Proof. 36 | intros. 37 | intro. 38 | use (proj1 (sort_lt_all a H) H0 _ H1). 39 | elim (S.E.lt_not_eq _ _ H2). reflexivity. 40 | Qed. 41 | 42 | Definition sort_lt_nodup (l : list var) := SortA_NoDupA _ _ _ (l:=l). 43 | (* Check sort_lt_nodup. *) 44 | 45 | Lemma sort_lt_ext : forall l1 l2, 46 | sort S.E.lt l1 -> sort S.E.lt l2 -> 47 | (forall a, InA S.E.eq a l1 <-> InA S.E.eq a l2) -> l1 = l2. 48 | Proof. 49 | intros. 50 | gen l2; induction H; intros. 51 | destruct* l2. 52 | use (proj2 (H1 t) (InA_cons_hd l2 (refl_equal t))). 53 | inversions H. 54 | inversions H1. 55 | use (proj1 (H2 a) (InA_cons_hd l (refl_equal a))). 56 | inversion H3. 57 | destruct (S.E.compare a a0). 58 | elim (sort_lt_notin H1 (cons_leA _ _ _ _ l1)). 59 | apply* (proj1 (H2 a)). auto with ordered_type. 60 | rewrite <- e in *. clear e a0. 61 | rewrite* (IHSorted l0). 62 | intro; split; intro. 63 | destruct (a0 == a). 64 | subst. 65 | elim (sort_lt_notin H H0 H5). 66 | use (proj1 (H2 a0) (InA_cons_tl a H5)). 67 | inversions* H6. 68 | destruct (a0 == a). 69 | subst. 70 | elim (sort_lt_notin H3 H4 H5). 71 | use (proj2 (H2 a0) (InA_cons_tl a H5)). 72 | inversions* H6. 73 | elim (sort_lt_notin (cons_sort H H0) (cons_leA _ _ _ _ l1)). 74 | apply* (proj2 (H2 a0)). 75 | Qed. 76 | 77 | Lemma remove_union : forall a L1 L2, 78 | S.remove a (L1 \u L2) = S.remove a L1 \u S.remove a L2. 79 | Proof. 80 | intros; apply eq_ext; intro; split; intro; sets_solve. 81 | Qed. 82 | 83 | Lemma nodup_notin_split : forall a l2 l1, 84 | NoDupA S.E.eq (l1 ++ a :: l2) -> ~InA S.E.eq a l1 /\ ~InA S.E.eq a l2. 85 | Proof. 86 | induction l1; simpl; intro; inversions H. 87 | split2*. intro. inversion H0. 88 | destruct (IHl1 H3). 89 | split2*. 90 | intro. inversions H4. 91 | elim H2. 92 | apply (InA_eqA _ H6 (l:=l1++a::l2)). 93 | apply (In_InA _). 94 | apply in_or_app. simpl*. 95 | elim (H0 H6). 96 | Qed. 97 | 98 | Lemma diff_remove : forall a L1 L2, 99 | a \in L2 -> S.diff (S.remove a L1) (S.remove a L2) = S.diff L1 L2. 100 | Proof. 101 | intros. 102 | apply eq_ext; intros; split; intro; sets_solve. 103 | apply* S.diff_3. 104 | destruct* (a == a0). 105 | intro. 106 | elim Hn. 107 | apply* S.remove_2. 108 | apply S.diff_3. 109 | apply* S.remove_2. 110 | intros Haa0. 111 | rewrite Haa0 in H. 112 | now elim Hn. 113 | intro; elim Hn. 114 | apply* S.remove_3. 115 | Qed. 116 | 117 | Lemma sort_tl : forall a l, sort S.E.lt (a::l) -> sort S.E.lt l. 118 | Proof. 119 | intros. inversions* H. 120 | Qed. 121 | 122 | Lemma sort_lelistA : forall a l, 123 | sort S.E.lt (a::l) -> lelistA S.E.lt a l. 124 | Proof. 125 | intros. inversions* H. 126 | Qed. 127 | 128 | Hint Resolve sort_tl sort_lelistA : core. 129 | 130 | Lemma sort_split : forall y l2 l1, 131 | sort S.E.lt (l1 ++ y :: l2) -> sort S.E.lt (l1 ++ l2). 132 | Proof. 133 | induction l1; simpl; intros. eauto. 134 | constructor. eauto. 135 | destruct l1; simpl in *. 136 | inversions H. 137 | inversions H3. 138 | apply* (InfA_ltA _). 139 | inversions* H. 140 | inversions* H3. 141 | Qed. 142 | 143 | Lemma elements_tl : forall a elts L, 144 | S.elements L = a :: elts -> S.elements (S.remove a L) = elts. 145 | Proof. 146 | intros. 147 | apply sort_lt_ext. 148 | apply S.elements_3. 149 | use (S.elements_3 L). 150 | rewrite H in H0. 151 | inversions* H0. 152 | intro; split; intro. 153 | use (S.elements_2 H0). 154 | use (S.remove_3 H1). 155 | use (S.elements_1 H2). 156 | rewrite H in H3. 157 | inversions* H3. 158 | elim (S.remove_1 (sym_eq H5) H1). 159 | apply S.elements_1. 160 | apply S.remove_2. 161 | intro. 162 | rewrite H1 in H. 163 | use (sort_lt_nodup (S.elements_3 L)). 164 | rewrite H in H2. 165 | inversions* H2. 166 | apply S.elements_2. 167 | rewrite* H. 168 | Qed. 169 | 170 | Lemma remove_remove : forall L a b, 171 | S.remove a (S.remove b L) = S.remove b (S.remove a L). 172 | Proof. 173 | intro. 174 | assert (forall a b x, x \in S.remove a (S.remove b L) -> 175 | x \in S.remove b (S.remove a L)). 176 | intros; sets_solve. 177 | intros. 178 | apply eq_ext; intro; split; intro; sets_solve. 179 | Qed. 180 | 181 | Lemma remove_idem : forall a L, 182 | S.remove a (S.remove a L) = S.remove a L. 183 | Proof. 184 | intros; apply eq_ext; intro; split; intro; sets_solve. 185 | Qed. 186 | 187 | Lemma elements_remove : forall a L, 188 | S.elements (S.remove a L) = removeA eq_var_dec a (S.elements L). 189 | Proof. 190 | intros. 191 | remember (S.elements L) as elts. 192 | gen L; induction elts; simpl; intros. 193 | apply sort_lt_ext. apply S.elements_3. 194 | auto. 195 | intro; split; intro. 196 | use (S.elements_1 (S.remove_3 (S.elements_2 H))). 197 | rewrite <- Heqelts in H0. 198 | inversion H0. 199 | inversion H. 200 | destruct (a == a0). 201 | subst. 202 | rewrite <- (IHelts (S.remove a0 L)). 203 | rewrite* remove_idem. 204 | apply sym_eq. 205 | apply* elements_tl. 206 | rewrite <- (IHelts (S.remove a0 L)); clear IHelts. 207 | apply sort_lt_ext. apply S.elements_3. 208 | constructor. apply S.elements_3. 209 | apply (InA_InfA _). 210 | intros. 211 | use (S.remove_3 (S.elements_2 H)). 212 | use (elements_tl (sym_equal Heqelts)). 213 | use (S.elements_1 H0). 214 | rewrite H1 in H2. 215 | use (S.elements_3 L). 216 | rewrite <- Heqelts in H3. 217 | inversions H3. 218 | use (sort_lt_all a0 H6). 219 | rewrite remove_remove. 220 | intro; split; intro. 221 | destruct* (a1 == a0). 222 | apply InA_cons_tl. 223 | apply S.elements_1. 224 | apply* S.remove_2. 225 | apply* S.elements_2. 226 | destruct (a0 == a1). 227 | subst. 228 | apply S.elements_1. 229 | apply* S.remove_2. 230 | apply S.elements_2. 231 | rewrite* <- Heqelts. auto with ordered_type. 232 | inversions H. elim n0; auto. 233 | apply S.elements_1. 234 | now apply (S.remove_3 (S.elements_2 H1)). 235 | rewrite* <- (@elements_tl a0 elts L). 236 | Qed. 237 | 238 | Lemma cardinal_union : forall L1 L2, 239 | S.cardinal (L1 \u L2) = S.cardinal L2 + S.cardinal (S.diff L1 L2). 240 | Proof. 241 | intros. 242 | repeat rewrite S.cardinal_1. 243 | remember (S.elements (L1 \u L2)) as elts. 244 | remember (S.elements L2) as elts2. 245 | remember (S.elements (S.diff L1 L2)) as elts1. 246 | gen L1 L2 elts1 elts. 247 | induction elts2; intros. 248 | use (elements_empty (sym_equal Heqelts2)). 249 | subst. 250 | rewrite diff_empty_r. 251 | rewrite* union_empty_r. 252 | use (elements_tl (sym_equal Heqelts2)). 253 | assert (a \in L2). 254 | apply S.elements_2. 255 | rewrite* <- Heqelts2. 256 | auto with ordered_type. 257 | assert (InA S.E.eq a elts). 258 | subst. 259 | now apply S.elements_1, S.union_3. 260 | subst. 261 | destruct (InA_split H1) as [l1 [y [l2 [Hl1 Hl]]]]. 262 | rewrite Hl. 263 | rewrite length_app. 264 | simpl. rewrite <- plus_n_Sm. 265 | rewrite <- length_app. 266 | apply (f_equal S). 267 | apply* (IHelts2 (S.remove a L1) (S.remove a L2)); clear IHelts2. 268 | apply (f_equal S.elements). 269 | rewrite* diff_remove. 270 | apply sort_lt_ext. 271 | use (S.elements_3 (L1 \u L2)). 272 | rewrite Hl in H. 273 | apply* sort_split. 274 | apply S.elements_3. 275 | intro; split; intro. 276 | apply S.elements_1. 277 | rewrite <- remove_union. 278 | rewrite <- Hl1 in *; clear Hl1 y. 279 | use (S.elements_3 (L1 \u L2)). 280 | destruct (a == a0). 281 | rewrite Hl in H2. 282 | use (sort_lt_nodup H2). 283 | destruct (nodup_notin_split _ _ _ H3). 284 | subst; destruct* (InA_app H). 285 | apply* S.remove_2. 286 | apply S.elements_2. 287 | rewrite Hl. 288 | clear -H. 289 | induction l1; simpl in *. auto. 290 | inversions* H. 291 | rewrite <- Hl1 in *; clear Hl1 y. 292 | use (S.elements_2 H). 293 | rewrite <- remove_union in H2. 294 | use (S.remove_3 H2). 295 | use (S.elements_1 H3). 296 | rewrite Hl in H4. 297 | clear -H2 H4. 298 | induction l1; simpl in *; inversions* H4. 299 | elim (S.remove_1 (sym_equal H0) H2). 300 | Qed. 301 | 302 | Lemma cardinal_remove : forall a L, 303 | a \in L -> 304 | S (S.cardinal (S.remove a L)) = S.cardinal L. 305 | Proof. 306 | intros. 307 | repeat rewrite S.cardinal_1. 308 | rewrite elements_remove. 309 | use (sort_lt_nodup (S.elements_3 L)). 310 | use (S.elements_1 H). 311 | clear H; induction H1. 312 | rewrite H. 313 | simpl. 314 | destruct* (y == y). 315 | inversions H0. 316 | clear -H3; induction l. auto. 317 | simpl. 318 | destruct (y==a). elim H3; auto. 319 | simpl; rewrite* IHl. 320 | simpl. 321 | inversions H0. 322 | destruct* (a==y). 323 | subst*. 324 | Qed. 325 | 326 | Lemma remove_subset : forall x L1 L2, L1 << L2 -> 327 | forall y, y \in S.remove x L1 -> y \in S.remove x L2. 328 | Proof. 329 | intros. 330 | apply S.remove_2. 331 | intro Hx; elim (S.remove_1 Hx H0). 332 | apply H; apply* S.remove_3. 333 | Qed. 334 | 335 | Lemma cardinal_subset : forall L1 L2, 336 | L1 << L2 -> S.cardinal L1 <= S.cardinal L2. 337 | Proof. 338 | intro. 339 | repeat rewrite S.cardinal_1. 340 | remember (S.elements L1) as elts1. 341 | gen L1; induction elts1; simpl; intros. 342 | omega. 343 | puts (elements_tl (sym_eq Heqelts1)). 344 | puts (IHelts1 _ (sym_eq H0) (S.remove a L2)). 345 | puts (H1 (remove_subset H)). 346 | assert (a \in L2). 347 | apply H. 348 | apply S.elements_2. 349 | rewrite* <- Heqelts1. 350 | auto with ordered_type. 351 | rewrite <- (cardinal_remove H3). 352 | omega. 353 | Qed. 354 | 355 | Lemma cardinal_empty : S.cardinal {} = 0. 356 | Proof. 357 | rewrite S.cardinal_1. 358 | case_eq (S.elements {}); intros. simpl*. 359 | assert (In e (e::l)) by simpl*. 360 | rewrite <- H in H0. 361 | assert (e \in {}). 362 | now apply S.elements_2, Var_as_OT_Facts.ListIn_In. 363 | elim (in_empty H1). 364 | Qed. 365 | 366 | Lemma cardinal_0 : forall L, 367 | S.cardinal L = 0 -> L = {}. 368 | Proof. 369 | intros. 370 | rewrite S.cardinal_1 in H. 371 | case_rewrite R1 (S.elements L). 372 | apply eq_ext; intros; split; intro; intros; sets_solve. 373 | use (S.elements_1 H0). 374 | rewrite R1 in H1. 375 | inversion H1. 376 | Qed. 377 | -------------------------------------------------------------------------------- /Lib_FinSet.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * A Library for Finite Sets with Leibnitz Equality * 3 | * Brian Aydemir & Arthur Charguéraud, March 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import OrderedTypeEx. 8 | Require Import Lib_Tactic Lib_MyFSetInterface. 9 | 10 | (* ********************************************************************** *) 11 | (** * Interface for Finite Sets *) 12 | 13 | Declare Scope set_scope. 14 | 15 | Module Type FinSet. 16 | Declare Module E : UsualOrderedType. 17 | 18 | (** These finite sets are compatible with Coq's FSet library, except with 19 | respect to [Hint]s. [MyFSetInterface.S] is a redeclaration of 20 | [FSetInterface.S] with all the [Hint]s placed in the [sets] database 21 | instead of [core]. This is mainly to avoid polluting the [core] database 22 | with lemmas that are only usable by [eauto]. *) 23 | 24 | Declare Module S : Lib_MyFSetInterface.S with Module E := E. 25 | 26 | (** Define aliases. *) 27 | 28 | Definition fset := S.t. 29 | Definition elt := S.elt. 30 | 31 | (** Equality on these sets is extensional. *) 32 | 33 | Parameter eq_ext : 34 | forall s s' : fset, (forall a : elt, S.In a s <-> S.In a s') -> s = s'. 35 | 36 | Parameter eq_if_Equal : 37 | forall s s' : fset, S.Equal s s' -> s = s'. 38 | 39 | (** Notations. *) 40 | 41 | Notation "{}" := (S.empty) : set_scope. 42 | 43 | Notation "{{ x }}" := (S.singleton x) : set_scope. 44 | 45 | Notation "E \u F" := (S.union E F) 46 | (at level 68, left associativity) : set_scope. 47 | 48 | Notation "x \in E" := (S.In x E) (at level 69) : set_scope. 49 | 50 | Notation "x \notin E" := (~(S.In x E)) (at level 69) : set_scope. 51 | 52 | Notation "E << F" := (S.Subset E F) (at level 70) : set_scope. 53 | 54 | Bind Scope set_scope with S.t. 55 | Bind Scope set_scope with fset. 56 | 57 | End FinSet. 58 | 59 | (* ********************************************************************** *) 60 | (** * Facts *) 61 | 62 | Module FinSetFacts (M : FinSet). 63 | 64 | Import M. 65 | Local Open Scope set_scope. 66 | 67 | (** Interaction of in with constructions *) 68 | 69 | Lemma in_empty : forall x, 70 | x \in {} -> False. 71 | Proof. 72 | intros. 73 | assert (S.Empty {}) by auto using S.empty_1. unfold S.Empty in *. 74 | assert (x \notin {}) by auto. 75 | intuition. 76 | Qed. 77 | 78 | Lemma in_singleton : forall x y, 79 | x \in {{y}} <-> x = y. 80 | Proof. 81 | intros; split. 82 | intro H. symmetry. apply S.singleton_1. trivial. 83 | intro H. apply S.singleton_2. unfold S.E.eq. auto. 84 | Qed. 85 | 86 | Lemma in_union : forall x E F, 87 | x \in (E \u F) <-> (x \in E) \/ (x \in F). 88 | Proof. 89 | intros; split. 90 | auto using S.union_1. 91 | intro H; destruct H as [ H | H ]; auto using S.union_2, S.union_3. 92 | Qed. 93 | 94 | (** Properties of union *) 95 | 96 | Lemma union_same : forall E, 97 | E \u E = E. 98 | Proof. 99 | intros. apply eq_if_Equal. 100 | split; repeat rewrite in_union; intuition. 101 | Qed. 102 | 103 | Lemma union_comm : forall E F, 104 | E \u F = F \u E. 105 | Proof. 106 | intros. apply eq_if_Equal. 107 | split; repeat rewrite in_union; intuition. 108 | Qed. 109 | 110 | Lemma union_assoc : forall E F G, 111 | E \u (F \u G) = (E \u F) \u G. 112 | Proof. 113 | intros. apply eq_if_Equal. 114 | split; repeat rewrite in_union; intuition. 115 | Qed. 116 | 117 | Lemma union_empty_l : forall E, 118 | {} \u E = E. 119 | Proof. 120 | intros. apply eq_if_Equal. 121 | split; repeat rewrite in_union; intuition. 122 | contradictions. apply* in_empty. 123 | Qed. 124 | 125 | (** More properties of in *) 126 | 127 | Lemma in_same : forall x, 128 | x \in {{x}}. 129 | intros. rewrite* in_singleton. 130 | Qed. 131 | 132 | (** More properties of union *) 133 | 134 | Lemma union_empty_r : forall E, 135 | E \u {} = E. 136 | intros. rewrite union_comm. 137 | apply union_empty_l. 138 | Qed. 139 | 140 | Lemma union_comm_assoc : forall E F G, 141 | E \u (F \u G) = F \u (E \u G). 142 | intros. rewrite union_assoc. 143 | rewrite (union_comm E F). 144 | rewrite* <- union_assoc. 145 | Qed. 146 | 147 | (** Subset relation properties *) 148 | 149 | Lemma subset_refl : forall E, 150 | E << E. 151 | introz. auto. 152 | Qed. 153 | 154 | Lemma subset_trans : forall F E G, 155 | E << F -> F << G -> E << G. 156 | introz. auto. 157 | Qed. 158 | 159 | (** Interaction of subset with constructions *) 160 | 161 | Lemma subset_empty_l : forall E, 162 | {} << E. 163 | introz. contradictions. apply* in_empty. 164 | Qed. 165 | 166 | Lemma subset_singleton : forall x E, 167 | x \in E <-> {{x}} << E. 168 | unfold S.Subset. split; intros. 169 | rewrite in_singleton in H0. subst*. 170 | apply (H x). apply in_same. 171 | Qed. 172 | 173 | Lemma subset_union_weak_l : forall E F, 174 | E << (E \u F). 175 | introz. rewrite* in_union. 176 | Qed. 177 | 178 | Lemma subset_union_weak_r : forall E F, 179 | F << (E \u F). 180 | introz. rewrite* in_union. 181 | Qed. 182 | 183 | Lemma subset_union_l : forall E F G, 184 | (E \u F) << G <-> (E << G) /\ (F << G). 185 | unfold S.Subset. splits; introz. 186 | apply H. apply* subset_union_weak_l. 187 | apply H. apply* subset_union_weak_r. 188 | rewrite in_union in H0. intuition auto. 189 | Qed. 190 | 191 | (** Interaction of notin with constructions *) 192 | 193 | Lemma notin_empty : forall x, 194 | x \notin {}. 195 | introz. apply* in_empty. 196 | Qed. 197 | 198 | Lemma notin_singleton : forall x y, 199 | x \notin {{y}} <-> x <> y. 200 | split; introz. 201 | apply H. rewrite* in_singleton. 202 | apply H. rewrite* <- in_singleton. 203 | Qed. 204 | 205 | Lemma notin_same : forall x, 206 | x \notin {{x}} -> False. 207 | intros. use in_same. 208 | Qed. 209 | 210 | Lemma notin_union : forall x E F, 211 | x \notin (E \u F) <-> (x \notin E) /\ (x \notin F). 212 | splits; intros. 213 | rewrite in_union in H. auto*. 214 | rewrite in_union in H. auto*. 215 | rewrite* in_union. 216 | Qed. 217 | 218 | End FinSetFacts. 219 | -------------------------------------------------------------------------------- /Lib_FinSetImpl.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Variation on FSet from library * 3 | * Brian Aydemir, March 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | 7 | (***********************************************************************) 8 | (* v * The Coq Proof Assistant / The Coq Development Team *) 9 | (* Raw.t; 45 | sorted : sort_bool this}. 46 | 47 | Definition from_sorted : forall (xs : Raw.t), sort E.lt xs -> sort_bool xs. 48 | Proof. 49 | intros xs H. unfold sort_bool. 50 | case (sort_dec E.lt OTFacts.lt_dec xs); tauto. 51 | Defined. 52 | 53 | Definition to_sorted : forall xs, sort_bool xs -> sort E.lt xs. 54 | Proof. 55 | unfold sort_bool. intros xs. 56 | case (sort_dec E.lt OTFacts.lt_dec xs); auto. 57 | intros. discriminate. 58 | Defined. 59 | 60 | Coercion to_sorted : sort_bool >-> sort. 61 | 62 | Definition Build_slist' (xs : Raw.t) (pf : sort E.lt xs) := 63 | Build_slist (from_sorted pf). 64 | 65 | (* XXX new code ends here *) 66 | 67 | (* XXX lots of minor minor tweaks *) 68 | 69 | Definition t := slist. 70 | Definition elt := E.t. 71 | 72 | Definition In (x : elt) (s : t) : Prop := InA E.eq x s.(this). 73 | Definition Equal (s s':t) : Prop := forall a : elt, In a s <-> In a s'. 74 | Definition Subset (s s':t) : Prop := forall a : elt, In a s -> In a s'. 75 | Definition Empty (s:t) : Prop := forall a : elt, ~ In a s. 76 | Definition For_all (P : elt -> Prop)(s:t) : Prop := forall x, In x s -> P x. 77 | Definition Exists (P : elt -> Prop)(s:t) : Prop := exists x, In x s /\ P x. 78 | 79 | Definition mem (x : elt) (s : t) : bool := Raw.mem x s. 80 | Definition add (x : elt)(s : t) : t := Build_slist' (Raw.add_sort (sorted s) x). 81 | Definition remove (x : elt)(s : t) : t := Build_slist' (Raw.remove_sort (sorted s) x). 82 | Definition singleton (x : elt) : t := Build_slist' (Raw.singleton_sort x). 83 | Definition union (s s' : t) : t := 84 | Build_slist' (Raw.union_sort (sorted s) (sorted s')). 85 | Definition inter (s s' : t) : t := 86 | Build_slist' (Raw.inter_sort (sorted s) (sorted s')). 87 | Definition diff (s s' : t) : t := 88 | Build_slist' (Raw.diff_sort (sorted s) (sorted s')). 89 | Definition equal (s s' : t) : bool := Raw.equal s s'. 90 | Definition subset (s s' : t) : bool := Raw.subset s s'. 91 | Definition empty : t := Build_slist' Raw.empty_sort. 92 | Definition is_empty (s : t) : bool := Raw.is_empty s. 93 | Definition elements (s : t) : list elt := Raw.elements s. 94 | Definition min_elt (s : t) : option elt := Raw.min_elt s. 95 | Definition max_elt (s : t) : option elt := Raw.max_elt s. 96 | Definition choose (s : t) : option elt := Raw.choose s. 97 | Definition fold (B : Type) (f : elt -> B -> B) (s : t) : B -> B := Raw.fold (B:=B) f s. 98 | Definition cardinal (s : t) : nat := Raw.cardinal s. 99 | Definition filter (f : elt -> bool) (s : t) : t := 100 | Build_slist' (Raw.filter_sort (sorted s) f). 101 | Definition for_all (f : elt -> bool) (s : t) : bool := Raw.for_all f s. 102 | Definition exists_ (f : elt -> bool) (s : t) : bool := Raw.exists_ f s. 103 | Definition partition (f : elt -> bool) (s : t) : t * t := 104 | let p := Raw.partition f s in 105 | (@Build_slist' (fst p) (Raw.partition_sort_1 (sorted s) f), 106 | @Build_slist' (snd p) (Raw.partition_sort_2 (sorted s) f)). 107 | Definition eq (s s' : t) : Prop := Raw.eq s s'. 108 | Definition lt (s s' : t) : Prop := Raw.lt s s'. 109 | 110 | Section Spec. 111 | Variable s s' s'': t. 112 | Variable x y : elt. 113 | 114 | Lemma In_1 : E.eq x y -> In x s -> In y s. 115 | Proof. exact (fun H H' => Raw.MX.In_eq H H'). Qed. 116 | 117 | Lemma mem_1 : In x s -> mem x s = true. 118 | Proof. exact (fun H => Raw.mem_1 s.(sorted) H). Qed. 119 | Lemma mem_2 : mem x s = true -> In x s. 120 | Proof. exact (fun H => Raw.mem_2 H). Qed. 121 | 122 | Lemma equal_1 : Equal s s' -> equal s s' = true. 123 | Proof. exact (Raw.equal_1 s.(sorted) s'.(sorted)). Qed. 124 | Lemma equal_2 : equal s s' = true -> Equal s s'. 125 | Proof. exact (fun H => Raw.equal_2 H). Qed. 126 | 127 | Lemma subset_1 : Subset s s' -> subset s s' = true. 128 | Proof. exact (Raw.subset_1 s.(sorted) s'.(sorted)). Qed. 129 | Lemma subset_2 : subset s s' = true -> Subset s s'. 130 | Proof. exact (fun H => Raw.subset_2 H). Qed. 131 | 132 | Lemma empty_1 : Empty empty. 133 | Proof. exact Raw.empty_1. Qed. 134 | 135 | Lemma is_empty_1 : Empty s -> is_empty s = true. 136 | Proof. exact (fun H => Raw.is_empty_1 H). Qed. 137 | Lemma is_empty_2 : is_empty s = true -> Empty s. 138 | Proof. exact (fun H => Raw.is_empty_2 H). Qed. 139 | 140 | Lemma add_1 : E.eq x y -> In y (add x s). 141 | Proof. exact (fun H => Raw.add_1 s.(sorted) H). Qed. 142 | Lemma add_2 : In y s -> In y (add x s). 143 | Proof. exact (fun H => Raw.add_2 s.(sorted) x H). Qed. 144 | Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s. 145 | Proof. exact (fun H => Raw.add_3 s.(sorted) H). Qed. 146 | 147 | Lemma remove_1 : E.eq x y -> ~ In y (remove x s). 148 | Proof. exact (fun H => Raw.remove_1 s.(sorted) H). Qed. 149 | Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s). 150 | Proof. exact (fun H H' => Raw.remove_2 s.(sorted) H H'). Qed. 151 | Lemma remove_3 : In y (remove x s) -> In y s. 152 | Proof. exact (fun H => Raw.remove_3 s.(sorted) H). Qed. 153 | 154 | Lemma singleton_1 : In y (singleton x) -> E.eq x y. 155 | Proof. exact (fun H => Raw.singleton_1 H). Qed. 156 | Lemma singleton_2 : E.eq x y -> In y (singleton x). 157 | Proof. exact (fun H => Raw.singleton_2 H). Qed. 158 | 159 | Lemma union_1 : In x (union s s') -> In x s \/ In x s'. 160 | Proof. exact (fun H => Raw.union_1 s.(sorted) s'.(sorted) H). Qed. 161 | Lemma union_2 : In x s -> In x (union s s'). 162 | Proof. exact (fun H => Raw.union_2 s.(sorted) s'.(sorted) H). Qed. 163 | Lemma union_3 : In x s' -> In x (union s s'). 164 | Proof. exact (fun H => Raw.union_3 s.(sorted) s'.(sorted) H). Qed. 165 | 166 | Lemma inter_1 : In x (inter s s') -> In x s. 167 | Proof. exact (fun H => Raw.inter_1 s.(sorted) s'.(sorted) H). Qed. 168 | Lemma inter_2 : In x (inter s s') -> In x s'. 169 | Proof. exact (fun H => Raw.inter_2 s.(sorted) s'.(sorted) H). Qed. 170 | Lemma inter_3 : In x s -> In x s' -> In x (inter s s'). 171 | Proof. exact (fun H => Raw.inter_3 s.(sorted) s'.(sorted) H). Qed. 172 | 173 | Lemma diff_1 : In x (diff s s') -> In x s. 174 | Proof. exact (fun H => Raw.diff_1 s.(sorted) s'.(sorted) H). Qed. 175 | Lemma diff_2 : In x (diff s s') -> ~ In x s'. 176 | Proof. exact (fun H => Raw.diff_2 s.(sorted) s'.(sorted) H). Qed. 177 | Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s'). 178 | Proof. exact (fun H => Raw.diff_3 s.(sorted) s'.(sorted) H). Qed. 179 | 180 | Lemma fold_1 : forall (A : Type) (i : A) (f : elt -> A -> A), 181 | fold f s i = fold_left (fun a e => f e a) (elements s) i. 182 | Proof. exact (Raw.fold_1 s.(sorted)). Qed. 183 | 184 | Lemma cardinal_1 : cardinal s = length (elements s). 185 | Proof. exact (Raw.cardinal_1 s.(sorted)). Qed. 186 | 187 | Section Filter. 188 | 189 | Variable f : elt -> bool. 190 | 191 | Lemma filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s. 192 | Proof. exact (@Raw.filter_1 s x f). Qed. 193 | Lemma filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true. 194 | Proof. exact (@Raw.filter_2 s x f). Qed. 195 | Lemma filter_3 : 196 | compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s). 197 | Proof. exact (@Raw.filter_3 s x f). Qed. 198 | 199 | Lemma for_all_1 : 200 | compat_bool E.eq f -> 201 | For_all (fun x => f x = true) s -> for_all f s = true. 202 | Proof. exact (@Raw.for_all_1 s f). Qed. 203 | Lemma for_all_2 : 204 | compat_bool E.eq f -> 205 | for_all f s = true -> For_all (fun x => f x = true) s. 206 | Proof. exact (@Raw.for_all_2 s f). Qed. 207 | 208 | Lemma exists_1 : 209 | compat_bool E.eq f -> 210 | Exists (fun x => f x = true) s -> exists_ f s = true. 211 | Proof. exact (@Raw.exists_1 s f). Qed. 212 | Lemma exists_2 : 213 | compat_bool E.eq f -> 214 | exists_ f s = true -> Exists (fun x => f x = true) s. 215 | Proof. exact (@Raw.exists_2 s f). Qed. 216 | 217 | Lemma partition_1 : 218 | compat_bool E.eq f -> Equal (fst (partition f s)) (filter f s). 219 | Proof. exact (@Raw.partition_1 s f). Qed. 220 | Lemma partition_2 : 221 | compat_bool E.eq f -> 222 | Equal (snd (partition f s)) (filter (fun x => negb (f x)) s). 223 | Proof. exact (@Raw.partition_2 s f). Qed. 224 | 225 | End Filter. 226 | 227 | Lemma elements_1 : In x s -> InA E.eq x (elements s). 228 | Proof. exact (fun H => Raw.elements_1 H). Qed. 229 | Lemma elements_2 : InA E.eq x (elements s) -> In x s. 230 | Proof. exact (fun H => Raw.elements_2 H). Qed. 231 | Lemma elements_3 : sort E.lt (elements s). 232 | Proof. exact (Raw.elements_3 s.(sorted)). Qed. 233 | Lemma elements_3w : NoDupA E.eq (elements s). 234 | Proof. exact (Raw.elements_3w s.(sorted)). Qed. 235 | 236 | Lemma min_elt_1 : min_elt s = Some x -> In x s. 237 | Proof. exact (fun H => Raw.min_elt_1 H). Qed. 238 | Lemma min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x. 239 | Proof. exact (fun H => Raw.min_elt_2 s.(sorted) H). Qed. 240 | Lemma min_elt_3 : min_elt s = None -> Empty s. 241 | Proof. exact (fun H => Raw.min_elt_3 H). Qed. 242 | 243 | Lemma max_elt_1 : max_elt s = Some x -> In x s. 244 | Proof. exact (fun H => Raw.max_elt_1 H). Qed. 245 | Lemma max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y. 246 | Proof. exact (fun H => Raw.max_elt_2 s.(sorted) H). Qed. 247 | Lemma max_elt_3 : max_elt s = None -> Empty s. 248 | Proof. exact (fun H => Raw.max_elt_3 H). Qed. 249 | 250 | Lemma choose_1 : choose s = Some x -> In x s. 251 | Proof. exact (fun H => Raw.choose_1 H). Qed. 252 | Lemma choose_2 : choose s = None -> Empty s. 253 | Proof. exact (fun H => Raw.choose_2 H). Qed. 254 | Lemma choose_3 : choose s = Some x -> choose s' = Some y -> 255 | Equal s s' -> E.eq x y. 256 | Proof. exact (@Raw.choose_3 _ _ s.(sorted) s'.(sorted) x y). Qed. 257 | 258 | Lemma eq_refl : eq s s. 259 | Proof. exact (Raw.eq_refl s). Qed. 260 | Lemma eq_sym : eq s s' -> eq s' s. 261 | Proof. exact (@Raw.eq_sym s s'). Qed. 262 | Lemma eq_trans : eq s s' -> eq s' s'' -> eq s s''. 263 | Proof. exact (@Raw.eq_trans s s' s''). Qed. 264 | 265 | Lemma lt_trans : lt s s' -> lt s' s'' -> lt s s''. 266 | Proof. exact (@Raw.lt_trans s s' s''). Qed. 267 | Lemma lt_not_eq : lt s s' -> ~ eq s s'. 268 | Proof. exact (Raw.lt_not_eq s.(sorted) s'.(sorted)). Qed. 269 | 270 | Definition compare : Compare lt eq s s'. 271 | Proof. 272 | elim (Raw.compare s.(sorted) s'.(sorted)); 273 | [ constructor 1 | constructor 2 | constructor 3 ]; 274 | auto. 275 | Defined. 276 | 277 | Definition eq_dec : { eq s s' } + { ~ eq s s' }. 278 | Proof. 279 | change eq with Equal. 280 | case_eq (equal s s'); intro H; [left | right]. 281 | apply equal_2; auto. 282 | intro H'; rewrite equal_1 in H; auto; discriminate. 283 | Defined. 284 | 285 | End Spec. 286 | 287 | End MakeRaw. 288 | 289 | (* XXX begin new code *) 290 | 291 | (** Now, implement the [FinSet] interface. *) 292 | 293 | Module Make (X : UsualOrderedType) <: FinSet with Module E := X. 294 | 295 | Module Import E := X. 296 | Module Import S := MakeRaw X. 297 | 298 | (** Define aliases. *) 299 | 300 | Definition fset := S.t. 301 | Definition elt := S.elt. 302 | 303 | (** Equality on these sets is extensional. *) 304 | 305 | Lemma bool_dec : forall x y : bool, 306 | x = y \/ x <> y. 307 | Proof. 308 | induction x; induction y; intuition. 309 | Qed. 310 | 311 | Theorem eq_ext : forall s s' : t, (forall a, In a s <-> In a s') -> s = s'. 312 | Proof. 313 | intros [s H] [s' J] K. 314 | assert (s = s'). 315 | unfold Raw.t in *. eapply Sort_InA_eq_ext; eauto using to_sorted. 316 | eexact E.lt_trans. 317 | eexact E.lt_not_eq. 318 | subst. 319 | rewrite (eq_proofs_unicity bool_dec H J). 320 | reflexivity. 321 | Qed. 322 | 323 | Theorem eq_if_Equal : forall s s' : t, Equal s s' -> s = s'. 324 | Proof. 325 | unfold Equal. intros s s'. 326 | auto using eq_ext. 327 | Qed. 328 | 329 | End Make. 330 | -------------------------------------------------------------------------------- /Lib_ListFacts.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Some facts about lists to complete standard library * 3 | * Brian E. Aydemir * 4 | ***************************************************************************) 5 | 6 | Require Import List. 7 | Require Import SetoidList. 8 | Require Import Sorting. 9 | Require Import Relations. 10 | Set Implicit Arguments. 11 | 12 | (* ********************************************************************** *) 13 | (** * List membership *) 14 | 15 | Lemma not_in_cons : 16 | forall (A : Type) (ys : list A) x y, 17 | x <> y -> ~ In x ys -> ~ In x (y :: ys). 18 | Proof. 19 | induction ys; simpl; intuition. 20 | Qed. 21 | 22 | Hint Resolve not_in_cons : core. 23 | 24 | Lemma not_In_app : 25 | forall (A : Type) (xs ys : list A) x, 26 | ~ In x xs -> ~ In x ys -> ~ In x (xs ++ ys). 27 | Proof. 28 | intros A xs ys x H J K. 29 | case (in_app_or _ _ _ K); auto. 30 | Qed. 31 | 32 | Hint Resolve not_In_app : core. 33 | 34 | Lemma elim_not_In_cons : 35 | forall (A : Type) (y : A) (ys : list A) (x : A), 36 | ~ In x (y :: ys) -> x <> y /\ ~ In x ys. 37 | Proof. 38 | intros. simpl in *. auto. 39 | Qed. 40 | 41 | Lemma elim_not_In_app : 42 | forall (A : Type) (xs ys : list A) (x : A), 43 | ~ In x (xs ++ ys) -> ~ In x xs /\ ~ In x ys. 44 | Proof. 45 | split; auto using in_or_app. 46 | Qed. 47 | 48 | (* ********************************************************************** *) 49 | (** * List inclusion *) 50 | 51 | Lemma incl_nil : 52 | forall (A : Type) (xs : list A), incl nil xs. 53 | Proof. 54 | unfold incl. 55 | intros A xs a H; inversion H. 56 | Qed. 57 | 58 | Hint Resolve incl_nil : core. 59 | 60 | Lemma incl_trans : 61 | forall (A : Type) (xs ys zs : list A), 62 | incl xs ys -> incl ys zs -> incl xs zs. 63 | Proof. 64 | unfold incl; firstorder. 65 | Qed. 66 | 67 | Hint Immediate incl_trans : core. 68 | 69 | Lemma In_incl : 70 | forall (A : Type) (x : A) (ys zs : list A), 71 | In x ys -> incl ys zs -> In x zs. 72 | Proof. 73 | unfold incl; auto. 74 | Qed. 75 | 76 | Hint Immediate In_incl : core. 77 | 78 | Lemma elim_incl_cons : 79 | forall (A : Type) (x : A) (xs zs : list A), 80 | incl (x :: xs) zs -> In x zs /\ incl xs zs. 81 | Proof. 82 | unfold incl; intros; split; auto with datatypes. 83 | Qed. 84 | 85 | Lemma elim_incl_app : 86 | forall (A : Type) (xs ys zs : list A), 87 | incl (xs ++ ys) zs -> incl xs zs /\ incl ys zs. 88 | Proof. 89 | unfold incl; intros; split; auto with datatypes. 90 | Qed. 91 | 92 | (* ********************************************************************** *) 93 | (** * Automation *) 94 | 95 | (** 96 | The following are placed in the [datatypes] library by the List theory. 97 | It's convenient to also have them in [core]. 98 | *) 99 | 100 | Hint Resolve in_eq : core. 101 | Hint Resolve in_cons : core. 102 | Hint Resolve incl_refl : core. 103 | Hint Resolve incl_nil : core. 104 | Hint Resolve incl_cons : core. 105 | Hint Resolve incl_tl : core. 106 | Hint Resolve incl_app : core. 107 | Hint Immediate incl_trans : core. 108 | 109 | (** 110 | The following tactics can be used to simply hypotheses concerning lists. 111 | *) 112 | 113 | Ltac simpl_list_hyp H := 114 | let LH1 := fresh "LH" in 115 | let LH2 := fresh "LH" in 116 | match type of H with 117 | | incl (?J :: ?K) ?L => 118 | destruct (elim_incl_cons H) as [LH1 LH2]; clear H; 119 | try simpl_list_hyp LH1; try simpl_list_hyp LH2 120 | | incl (?J ++ ?K) ?L => 121 | destruct (elim_incl_app J K H) as [LH1 LH2]; clear H; 122 | try simpl_list_hyp LH1; try simpl_list_hyp LH2 123 | | incl nil _ => 124 | clear H 125 | | In ?x (?y :: ?ys) => 126 | destruct (in_inv H) as [LH1 | LH1]; clear H; 127 | try simpl_list_hyp LH1 128 | | In ?x (?ys ++ ?zs) => 129 | destruct (in_app_or ys zs x H) as [LH1 | LH1]; clear H; 130 | try simpl_list_hyp LH1 131 | | In _ nil => 132 | simpl in H; intuition 133 | | ~ In _ nil => 134 | clear H 135 | | ~ In _ (_ :: _) => 136 | destruct (elim_not_In_cons H) as [LH1 LH2]; clear H; 137 | try simpl_list_hyp LH1; try simpl_list_hyp LH2 138 | | ~ In ?x (?K ++ ?L) => 139 | destruct (elim_not_In_app K L x H) as [LH1 LH2]; clear H; 140 | try simpl_list_hyp LH1; try simpl_list_hyp LH2 141 | | In _ _ => 142 | progress (simpl in H) 143 | | incl _ _ => 144 | progress (simpl in H) 145 | | ~ In _ _ => 146 | progress (simpl in H) 147 | end. 148 | 149 | Ltac simpl_list_hyps := 150 | match goal with 151 | | H : _ |- _ => simpl_list_hyp H; simpl_list_hyps 152 | | H : ~ (?a = ?b \/ False), J : ?b = ?a |- _ => subst b; intuition 153 | | H : ~ (?a = ?b \/ False), J : ?a = ?b |- _ => subst a; intuition 154 | | _ => idtac 155 | end. 156 | 157 | Hint Extern 4 (In ?x ?L) => simpl; simpl_list_hyps : core. 158 | Hint Extern 4 (~ In ?x ?L) => simpl; simpl_list_hyps : core. 159 | Hint Extern 4 (incl ?L1 ?L2) => simpl; simpl_list_hyps : core. 160 | 161 | (* ********************************************************************** *) 162 | (** * Setoid facts *) 163 | 164 | Lemma InA_iff_In : 165 | forall (A : Set) x xs, InA (@eq A) x xs <-> In x xs. 166 | Proof. 167 | induction xs as [ | y ys IH ]. 168 | split; intros H; inversion H. 169 | split; intros H; inversion_clear H. 170 | subst x; auto with datatypes. 171 | assert (In x ys) by intuition; auto with datatypes. 172 | subst y; auto with datatypes. 173 | assert (InA (@eq A) x ys) by intuition; auto with datatypes. 174 | Qed. 175 | 176 | (* ********************************************************************** *) 177 | (** * Decidable sorting *) 178 | 179 | (** 180 | It is decidable to tell whether a list a sorted according to 181 | some relation. 182 | *) 183 | 184 | Section DecidableSorting. 185 | 186 | Variable A : Type. 187 | Variable leA : relation A. 188 | Hypothesis leA_dec : forall x y, {leA x y} + {~ leA x y}. 189 | 190 | Theorem lelistA_dec : 191 | forall a xs, {lelistA leA a xs} + {~ lelistA leA a xs}. 192 | Proof. 193 | induction xs as [ | x xs IH ]; auto with datatypes. 194 | case (leA_dec a x); auto with datatypes. 195 | intros H. right. intros J. inversion J. auto. 196 | Qed. 197 | 198 | Theorem sort_dec : 199 | forall xs, {sort leA xs} + {~ sort leA xs}. 200 | Proof. 201 | induction xs as [ | x xs IH ]; auto with datatypes. 202 | case IH; case (lelistA_dec x xs); auto with datatypes. 203 | intros H J. right. intros K. inversion K. auto. 204 | intros H J. right. intros K. inversion K. auto. 205 | intros H J. right. intros K. inversion K. auto. 206 | Qed. 207 | 208 | End DecidableSorting. 209 | 210 | (* ********************************************************************** *) 211 | (** * Equality on sorted lists *) 212 | 213 | (** 214 | Two sorted lists are equal if they contain the same elements. 215 | *) 216 | 217 | Section Equality_ext. 218 | 219 | Variable A : Type. 220 | Variable ltA : relation A. 221 | Hypothesis ltA_trans : forall x y z, ltA x y -> ltA y z -> ltA x z. 222 | Hypothesis ltA_not_eqA : forall x y, ltA x y -> x <> y. 223 | Hypothesis ltA_eqA : forall x y z, ltA x y -> y = z -> ltA x z. 224 | Hypothesis eqA_ltA : forall x y z, x = y -> ltA y z -> ltA x z. 225 | 226 | Hint Resolve ltA_trans : core. 227 | Hint Immediate ltA_eqA eqA_ltA : core. 228 | 229 | Notation Inf := (lelistA ltA). 230 | Notation Sort := (sort ltA). 231 | 232 | Lemma strictOrder_ltA : StrictOrder ltA. 233 | Proof. 234 | try repeat (constructor; intuition). 235 | intro; intros; intro. 236 | apply (ltA_not_eqA H). 237 | auto. 238 | Qed. 239 | Hint Resolve strictOrder_ltA : core. 240 | 241 | Lemma not_InA_if_Sort_Inf : 242 | forall xs a, Sort xs -> Inf a xs -> ~ InA (@eq A) a xs. 243 | Proof. 244 | induction xs as [ | x xs IH ]; intros a Hsort Hinf H. 245 | inversion H. 246 | inversion H; subst. 247 | inversion Hinf; subst. 248 | assert (x <> x) by auto; intuition. 249 | inversion Hsort; inversion Hinf; subst. 250 | assert (Inf a xs) by eauto using InfA_ltA. 251 | assert (~ InA (@eq A) a xs) by auto. 252 | intuition. 253 | Qed. 254 | 255 | Lemma Sort_eq_head : 256 | forall x xs y ys, 257 | Sort (x :: xs) -> 258 | Sort (y :: ys) -> 259 | (forall a, InA (@eq A) a (x :: xs) <-> InA (@eq A) a (y :: ys)) -> 260 | x = y. 261 | Proof. 262 | intros x xs y ys SortXS SortYS H. 263 | inversion SortXS; inversion SortYS; subst. 264 | assert (Q3 : InA (@eq A) x (y :: ys)) by firstorder. 265 | assert (Q4 : InA (@eq A) y (x :: xs)) by firstorder. 266 | inversion Q3; subst; auto. 267 | inversion Q4; subst; auto. 268 | assert (ltA y x) by (refine (SortA_InfA_InA _ _ _ H6 H7 H1); intuition). 269 | assert (ltA x y) by (refine (SortA_InfA_InA _ _ _ H2 H3 H4); intuition). 270 | assert (y <> y) by eauto. 271 | intuition. 272 | Qed. 273 | 274 | Lemma Sort_InA_eq_ext : 275 | forall xs ys, 276 | Sort xs -> 277 | Sort ys -> 278 | (forall a, InA (@eq A) a xs <-> InA (@eq A) a ys) -> 279 | xs = ys. 280 | Proof. 281 | induction xs as [ | x xs IHxs ]; induction ys as [ | y ys IHys ]; 282 | intros SortXS SortYS H; auto. 283 | (* xs -> nil, ys -> y :: ys *) 284 | assert (Q : InA (@eq A) y nil) by firstorder. 285 | inversion Q. 286 | (* xs -> x :: xs, ys -> nil *) 287 | assert (Q : InA (@eq A) x nil) by firstorder. 288 | inversion Q. 289 | (* xs -> x :: xs, ys -> y :: ys *) 290 | inversion SortXS; inversion SortYS; subst. 291 | assert (x = y) by eauto using Sort_eq_head. 292 | cut (forall a, InA (@eq A) a xs <-> InA (@eq A) a ys). 293 | intros. assert (xs = ys) by auto. subst. auto. 294 | intros a; split; intros L. 295 | assert (Q2 : InA (@eq A) a (y :: ys)) by firstorder. 296 | inversion Q2; subst; auto. 297 | assert (Q5 : ~ InA (@eq A) y xs) by auto using not_InA_if_Sort_Inf. 298 | intuition. 299 | assert (Q2 : InA (@eq A) a (x :: xs)) by firstorder. 300 | inversion Q2; subst; auto. 301 | assert (Q5 : ~ InA (@eq A) y ys) by auto using not_InA_if_Sort_Inf. 302 | intuition. 303 | Qed. 304 | 305 | End Equality_ext. 306 | -------------------------------------------------------------------------------- /Lib_ListFactsMore.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * More properties about Lists * 3 | * Arthur Chargueraud, July 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import Lib_Tactic List. 8 | 9 | (** Results on List.map *) 10 | 11 | Lemma list_map_id : forall (A : Set) l, 12 | List.map (fun t : A => t) l = l. 13 | Proof. 14 | induction l; simpl; f_equal*. 15 | Qed. 16 | 17 | Lemma list_concat_right : forall (A : Set) (x : A) l1 l2, 18 | l1 ++ (x :: l2) = (l1 ++ (x :: nil)) ++ l2. 19 | Proof. 20 | intros. induction l1; simpl; f_equal*. 21 | Qed. 22 | 23 | Lemma list_map_nth : forall (A : Set) (f : A -> A) d l n, 24 | f d = d -> 25 | f (List.nth n l d) = List.nth n (List.map f l) d. 26 | Proof. 27 | induction l; simpl; introv Fix; destruct* n. 28 | Qed. 29 | 30 | (** Results on List.forall *) 31 | 32 | Inductive list_forall (A : Set) (P : A -> Prop) : list A -> Prop := 33 | | list_forall_nil : list_forall P nil 34 | | list_forall_cons : forall L x, 35 | list_forall P L -> P x -> list_forall P (x::L). 36 | 37 | Hint Constructors list_forall : core. 38 | 39 | Lemma list_forall_concat : forall (A : Set) (P : A -> Prop) L1 L2, 40 | list_forall P L1 -> 41 | list_forall P L2 -> 42 | list_forall P (L1 ++ L2). 43 | Proof. 44 | induction L1; simpl; intros. auto. inversions* H. 45 | Qed. 46 | 47 | (** Results on List.for_n *) 48 | 49 | Definition list_for_n (A : Set) (P : A -> Prop) (n : nat) (L : list A) := 50 | n = length L /\ list_forall P L. 51 | 52 | Lemma list_for_n_concat : forall (A : Set) (P : A -> Prop) n1 n2 L1 L2, 53 | list_for_n P n1 L1 -> 54 | list_for_n P n2 L2 -> 55 | list_for_n P (n1+n2) (L1 ++ L2). 56 | Proof. 57 | unfold list_for_n. intros. split. 58 | rewrite* length_app. apply* list_forall_concat. 59 | Qed. 60 | 61 | Hint Extern 1 (?n = length ?xs) => 62 | match goal with H: list_for_n _ ?n ?xs |- _ => 63 | apply (proj1 H) end : more. 64 | 65 | Hint Extern 1 (length ?xs = ?n) => 66 | match goal with H: list_for_n _ ?n ?xs |- _ => 67 | apply (sym_eq (proj1 H)) end : more. 68 | 69 | 70 | 71 | -------------------------------------------------------------------------------- /Lib_MyFSetInterface.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Variation on FSet from library * 3 | * Modified by Brian E. Aydemir (baydemir [at] cis.upenn.edu) March, 2007. * 4 | ***************************************************************************) 5 | 6 | (** 7 | This module is nothing more than a redeclaration of [FSetInterface.S], 8 | except all the [Hints] in the interface are put into the [sets] database 9 | instead of [core]. 10 | *) 11 | 12 | (***********************************************************************) 13 | (* v * The Coq Proof Assistant / The Coq Development Team *) 14 | (* t -> Prop. 44 | Definition Equal s s' := forall a : elt, In a s <-> In a s'. 45 | Definition Subset s s' := forall a : elt, In a s -> In a s'. 46 | Definition Empty s := forall a : elt, ~ In a s. 47 | Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. 48 | Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. 49 | 50 | Notation "s [=] t" := (Equal s t) (at level 70, no associativity). 51 | Notation "s [<=] t" := (Subset s t) (at level 70, no associativity). 52 | 53 | Parameter empty : t. 54 | (** The empty set. *) 55 | 56 | Parameter is_empty : t -> bool. 57 | (** Test whether a set is empty or not. *) 58 | 59 | Parameter mem : elt -> t -> bool. 60 | (** [mem x s] tests whether [x] belongs to the set [s]. *) 61 | 62 | Parameter add : elt -> t -> t. 63 | (** [add x s] returns a set containing all elements of [s], 64 | plus [x]. If [x] was already in [s], [s] is returned unchanged. *) 65 | 66 | Parameter singleton : elt -> t. 67 | (** [singleton x] returns the one-element set containing only [x]. *) 68 | 69 | Parameter remove : elt -> t -> t. 70 | (** [remove x s] returns a set containing all elements of [s], 71 | except [x]. If [x] was not in [s], [s] is returned unchanged. *) 72 | 73 | Parameter union : t -> t -> t. 74 | (** Set union. *) 75 | 76 | Parameter inter : t -> t -> t. 77 | (** Set intersection. *) 78 | 79 | Parameter diff : t -> t -> t. 80 | (** Set difference. *) 81 | 82 | Definition eq : t -> t -> Prop := Equal. 83 | Parameter lt : t -> t -> Prop. 84 | Parameter compare : forall s s' : t, Compare lt eq s s'. 85 | (** Total ordering between sets. Can be used as the ordering function 86 | for doing sets of sets. *) 87 | 88 | Parameter equal : t -> t -> bool. 89 | (** [equal s1 s2] tests whether the sets [s1] and [s2] are 90 | equal, that is, contain equal elements. *) 91 | 92 | Parameter subset : t -> t -> bool. 93 | (** [subset s1 s2] tests whether the set [s1] is a subset of 94 | the set [s2]. *) 95 | 96 | (** Coq comment: [iter] is useless in a purely functional world *) 97 | (** iter: (elt -> unit) -> set -> unit. i*) 98 | (** [iter f s] applies [f] in turn to all elements of [s]. 99 | The order in which the elements of [s] are presented to [f] 100 | is unspecified. *) 101 | 102 | Parameter fold : forall A : Type, (elt -> A -> A) -> t -> A -> A. 103 | (** [fold f s a] computes [(f xN ... (f x2 (f x1 a))...)], 104 | where [x1 ... xN] are the elements of [s], in increasing order. *) 105 | 106 | Parameter for_all : (elt -> bool) -> t -> bool. 107 | (** [for_all p s] checks if all elements of the set 108 | satisfy the predicate [p]. *) 109 | 110 | Parameter exists_ : (elt -> bool) -> t -> bool. 111 | (** [exists p s] checks if at least one element of 112 | the set satisfies the predicate [p]. *) 113 | 114 | Parameter filter : (elt -> bool) -> t -> t. 115 | (** [filter p s] returns the set of all elements in [s] 116 | that satisfy predicate [p]. *) 117 | 118 | Parameter partition : (elt -> bool) -> t -> t * t. 119 | (** [partition p s] returns a pair of sets [(s1, s2)], where 120 | [s1] is the set of all the elements of [s] that satisfy the 121 | predicate [p], and [s2] is the set of all the elements of 122 | [s] that do not satisfy [p]. *) 123 | 124 | Parameter cardinal : t -> nat. 125 | (** Return the number of elements of a set. *) 126 | (** Coq comment: nat instead of int ... *) 127 | 128 | Parameter elements : t -> list elt. 129 | (** Return the list of all elements of the given set. 130 | The returned list is sorted in increasing order with respect 131 | to the ordering [Ord.compare], where [Ord] is the argument 132 | given to {!Set.Make}. *) 133 | 134 | Parameter min_elt : t -> option elt. 135 | (** Return the smallest element of the given set 136 | (with respect to the [Ord.compare] ordering), or raise 137 | [Not_found] if the set is empty. *) 138 | (** Coq comment: [Not_found] is represented by the option type *) 139 | 140 | Parameter max_elt : t -> option elt. 141 | (** Same as {!Set.S.min_elt}, but returns the largest element of the 142 | given set. *) 143 | (** Coq comment: [Not_found] is represented by the option type *) 144 | 145 | Parameter choose : t -> option elt. 146 | (** Return one element of the given set, or raise [Not_found] if 147 | the set is empty. Which element is chosen is unspecified, 148 | but equal elements will be chosen for equal sets. *) 149 | (** Coq comment: [Not_found] is represented by the option type *) 150 | 151 | Section Spec. 152 | 153 | Variable s s' s'' : t. 154 | Variable x y : elt. 155 | 156 | (** Specification of [In] *) 157 | Parameter In_1 : E.eq x y -> In x s -> In y s. 158 | 159 | (** Specification of [eq] *) 160 | Parameter eq_refl : eq s s. 161 | Parameter eq_sym : eq s s' -> eq s' s. 162 | Parameter eq_trans : eq s s' -> eq s' s'' -> eq s s''. 163 | 164 | (** Specification of [lt] *) 165 | Parameter lt_trans : lt s s' -> lt s' s'' -> lt s s''. 166 | Parameter lt_not_eq : lt s s' -> ~ eq s s'. 167 | 168 | (** Specification of [mem] *) 169 | Parameter mem_1 : In x s -> mem x s = true. 170 | Parameter mem_2 : mem x s = true -> In x s. 171 | 172 | (** Specification of [equal] *) 173 | Parameter equal_1 : s[=]s' -> equal s s' = true. 174 | Parameter equal_2 : equal s s' = true ->s[=]s'. 175 | 176 | (** Specification of [subset] *) 177 | Parameter subset_1 : s[<=]s' -> subset s s' = true. 178 | Parameter subset_2 : subset s s' = true -> s[<=]s'. 179 | 180 | (** Specification of [empty] *) 181 | Parameter empty_1 : Empty empty. 182 | 183 | (** Specification of [is_empty] *) 184 | Parameter is_empty_1 : Empty s -> is_empty s = true. 185 | Parameter is_empty_2 : is_empty s = true -> Empty s. 186 | 187 | (** Specification of [add] *) 188 | Parameter add_1 : E.eq x y -> In y (add x s). 189 | Parameter add_2 : In y s -> In y (add x s). 190 | Parameter add_3 : ~ E.eq x y -> In y (add x s) -> In y s. 191 | 192 | (** Specification of [remove] *) 193 | Parameter remove_1 : E.eq x y -> ~ In y (remove x s). 194 | Parameter remove_2 : ~ E.eq x y -> In y s -> In y (remove x s). 195 | Parameter remove_3 : In y (remove x s) -> In y s. 196 | 197 | (** Specification of [singleton] *) 198 | Parameter singleton_1 : In y (singleton x) -> E.eq x y. 199 | Parameter singleton_2 : E.eq x y -> In y (singleton x). 200 | 201 | (** Specification of [union] *) 202 | Parameter union_1 : In x (union s s') -> In x s \/ In x s'. 203 | Parameter union_2 : In x s -> In x (union s s'). 204 | Parameter union_3 : In x s' -> In x (union s s'). 205 | 206 | (** Specification of [inter] *) 207 | Parameter inter_1 : In x (inter s s') -> In x s. 208 | Parameter inter_2 : In x (inter s s') -> In x s'. 209 | Parameter inter_3 : In x s -> In x s' -> In x (inter s s'). 210 | 211 | (** Specification of [diff] *) 212 | Parameter diff_1 : In x (diff s s') -> In x s. 213 | Parameter diff_2 : In x (diff s s') -> ~ In x s'. 214 | Parameter diff_3 : In x s -> ~ In x s' -> In x (diff s s'). 215 | 216 | (** Specification of [fold] *) 217 | Parameter fold_1 : forall (A : Type) (i : A) (f : elt -> A -> A), 218 | fold f s i = fold_left (fun a e => f e a) (elements s) i. 219 | 220 | (** Specification of [cardinal] *) 221 | Parameter cardinal_1 : cardinal s = length (elements s). 222 | 223 | Section Filter. 224 | 225 | Variable f : elt -> bool. 226 | 227 | (** Specification of [filter] *) 228 | Parameter filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s. 229 | Parameter filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true. 230 | Parameter filter_3 : 231 | compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s). 232 | 233 | (** Specification of [for_all] *) 234 | Parameter for_all_1 : 235 | compat_bool E.eq f -> 236 | For_all (fun x => f x = true) s -> for_all f s = true. 237 | Parameter for_all_2 : 238 | compat_bool E.eq f -> 239 | for_all f s = true -> For_all (fun x => f x = true) s. 240 | 241 | (** Specification of [exists] *) 242 | Parameter exists_1 : 243 | compat_bool E.eq f -> 244 | Exists (fun x => f x = true) s -> exists_ f s = true. 245 | Parameter exists_2 : 246 | compat_bool E.eq f -> 247 | exists_ f s = true -> Exists (fun x => f x = true) s. 248 | 249 | (** Specification of [partition] *) 250 | Parameter partition_1 : compat_bool E.eq f -> 251 | fst (partition f s) [=] filter f s. 252 | Parameter partition_2 : compat_bool E.eq f -> 253 | snd (partition f s) [=] filter (fun x => negb (f x)) s. 254 | 255 | End Filter. 256 | 257 | (** Specification of [elements] *) 258 | Parameter elements_1 : In x s -> InA E.eq x (elements s). 259 | Parameter elements_2 : InA E.eq x (elements s) -> In x s. 260 | Parameter elements_3 : sort E.lt (elements s). 261 | Parameter elements_3w : NoDupA E.eq (elements s). 262 | 263 | (** Specification of [min_elt] *) 264 | Parameter min_elt_1 : min_elt s = Some x -> In x s. 265 | Parameter min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x. 266 | Parameter min_elt_3 : min_elt s = None -> Empty s. 267 | 268 | (** Specification of [max_elt] *) 269 | Parameter max_elt_1 : max_elt s = Some x -> In x s. 270 | Parameter max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y. 271 | Parameter max_elt_3 : max_elt s = None -> Empty s. 272 | 273 | (** Specification of [choose] *) 274 | Parameter choose_1 : choose s = Some x -> In x s. 275 | Parameter choose_2 : choose s = None -> Empty s. 276 | (* Parameter choose_equal: 277 | (equal s s')=true -> E.eq (choose s) (choose s'). *) 278 | 279 | End Spec. 280 | 281 | (* begin hide *) 282 | Hint Immediate In_1 : sets. 283 | 284 | Hint Resolve mem_1 mem_2 equal_1 equal_2 subset_1 subset_2 empty_1 285 | is_empty_1 is_empty_2 choose_1 choose_2 add_1 add_2 add_3 remove_1 286 | remove_2 remove_3 singleton_1 singleton_2 union_1 union_2 union_3 inter_1 287 | inter_2 inter_3 diff_1 diff_2 diff_3 filter_1 filter_2 filter_3 for_all_1 288 | for_all_2 exists_1 exists_2 partition_1 partition_2 elements_1 elements_2 289 | elements_3 min_elt_1 min_elt_2 min_elt_3 max_elt_1 max_elt_2 max_elt_3 290 | : sets. 291 | (* end hide *) 292 | 293 | End S. 294 | -------------------------------------------------------------------------------- /Lib_Tactic.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * General useful tactics for Coq * 3 | * Brian Aydemir & Arthur Charguéraud, March 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | (* ********************************************************************** *) 7 | (** * Simple variations on existing tactics *) 8 | 9 | (** [contradictions] replace the goal by False and prove it if False is 10 | derivable from the context or if [discriminate] applies. *) 11 | 12 | Ltac contradictions := 13 | assert False; [ try discriminate | contradiction ]. 14 | 15 | (** [cuts] does [cut] then [intro] in the first subgoal. *) 16 | 17 | Ltac cuts H E := 18 | cut (E); [ intro H | idtac ]. 19 | 20 | (** [inversions H] is a shortcut for [inversion H] followed by [subst]. *) 21 | 22 | Ltac inversions H := 23 | inversion H; subst. 24 | 25 | (** [poses H E] adds an hypothesis with name H and with type the type of E. *) 26 | 27 | Ltac poses H E := 28 | pose (H := E); clearbody H. 29 | 30 | (** [puts] is a version of [poses] where Coq chooses the name introduced. *) 31 | 32 | Ltac puts E := 33 | let H := fresh in poses H E. 34 | 35 | (** [asserts H E] is a synonymous for [assert (X : E)] provided for 36 | uniformity with the rest of the syntax. *) 37 | 38 | Ltac asserts H E := 39 | assert (H : E). 40 | 41 | (** [sets X E] replaces all occurences of E by a name X, and forgets the 42 | fact that X is equal to X -- it makes the goal more general *) 43 | 44 | Ltac sets X E := 45 | set (X := E) in *; clearbody X. 46 | 47 | (** [introz] repeats [intro] as long as possible. Contrary to [intros], 48 | it unfolds any definition on the way. *) 49 | 50 | Ltac introz := 51 | intro; try introz. 52 | 53 | (* ********************************************************************** *) 54 | (** * Unfolding *) 55 | 56 | (** [folds] is a shortcut for [fold in *] *) 57 | 58 | Tactic Notation "folds" constr(H) := 59 | fold H in *. 60 | Tactic Notation "folds" constr(H1) constr(H2) := 61 | folds H1; folds H2. 62 | Tactic Notation "folds" constr(H1) constr(H2) constr(H3) := 63 | folds H1; folds H2; folds H3. 64 | 65 | (** [unfolds] is a shortcut for [unfold in *] *) 66 | 67 | Tactic Notation "unfolds" reference(F1) := 68 | unfold F1 in *. 69 | Tactic Notation "unfolds" reference(F1) reference(F2) := 70 | unfold F1 in *; unfold F2 in *. 71 | Tactic Notation "unfolds" reference(F1) reference(F2) reference(F3) := 72 | unfold F1 in *; unfold F2 in *; unfold F3 in *. 73 | 74 | (** [unfold_hd] unfolds the definition at the head of the goal. *) 75 | 76 | Tactic Notation "unfold_hd" := 77 | match goal with 78 | | |- ?P => unfold P 79 | | |- ?P _ => unfold P 80 | | |- ?P _ _ => unfold P 81 | | |- ?P _ _ _ => unfold P 82 | | |- ?P _ _ _ _ => unfold P 83 | end. 84 | 85 | 86 | (* ********************************************************************** *) 87 | (** * Simplification *) 88 | 89 | (** [simpls] is a shortcut for [simpl in *] *) 90 | 91 | Tactic Notation "simpls" := 92 | simpl in *. 93 | Tactic Notation "simpls" reference(F1) := 94 | simpl F1 in *. 95 | Tactic Notation "simpls" reference(F1) reference(F2) := 96 | simpl F1 in *; simpl F2 in *. 97 | Tactic Notation "simpls" reference(F1) reference(F2) reference(F3) := 98 | simpl F1 in *; simpl F2 in *; simpl F3 in *. 99 | 100 | (** [unsimpl E] replaces all occurence of X by E, where X is the result 101 | which tactic [simpl] would give when applied to E. *) 102 | 103 | Tactic Notation "unsimpl" constr(E) := 104 | let F := (eval simpl in E) in change F with E. 105 | 106 | Tactic Notation "unsimpl" constr(E) "in" hyp(H) := 107 | let F := (eval simpl in E) in change F with E in H. 108 | 109 | 110 | (* ********************************************************************** *) 111 | (** * Rewriting *) 112 | 113 | (** [rewrites] is an iterated version of [rewrite]. Beware of loops! *) 114 | 115 | Tactic Notation "rewrites" constr(E) := 116 | repeat rewrite E. 117 | Tactic Notation "rewrites" "<-" constr(E) := 118 | repeat rewrite <- E. 119 | Tactic Notation "rewrites" constr(E) "in" ident(H) := 120 | repeat rewrite E in H. 121 | Tactic Notation "rewrites" "<-" constr(E) "in" ident(H) := 122 | repeat rewrite <- E in H. 123 | 124 | (** [asserts_rew] can be used to assert an equality holds and rewrite it 125 | straight away in the current goal *) 126 | 127 | Tactic Notation "asserts_rew" constr(E) := 128 | let EQ := fresh in (assert (EQ : E); 129 | [ idtac | rewrite EQ; clear EQ ]). 130 | Tactic Notation "asserts_rew" "<-" constr(E) := 131 | let EQ := fresh in (assert (EQ : E); 132 | [ idtac | rewrite <- EQ; clear EQ ]). 133 | Tactic Notation "asserts_rew" constr(E) "in" hyp(H) := 134 | let EQ := fresh in (assert (EQ : E); 135 | [ idtac | rewrite EQ in H; clear EQ ]). 136 | Tactic Notation "asserts_rew" "<-" constr(E) "in" hyp(H) := 137 | let EQ := fresh in (assert (EQ : E); 138 | [ idtac | rewrite <- EQ in H; clear EQ ]). 139 | 140 | (** [do_rew] is used to perform the sequence: 141 | rewrite the goal, execute a tactic, rewrite the goal back *) 142 | 143 | Tactic Notation "do_rew" constr(E) tactic(T) := 144 | rewrite E; T; try rewrite <- E. 145 | 146 | Tactic Notation "do_rew" "<-" constr(E) tactic(T) := 147 | rewrite <- E; T; try rewrite E. 148 | 149 | (** [do_rew_2] is the same as [do_rew] but it does rewrite twice *) 150 | 151 | Tactic Notation "do_rew_2" constr(E) tactic(T) := 152 | do 2 rewrite E; T; try do 2 rewrite <- E. 153 | 154 | Tactic Notation "do_rew_2" "<-" constr(E) tactic(T) := 155 | do 2 rewrite <- E; T; try do 2 rewrite E. 156 | 157 | 158 | (* ********************************************************************** *) 159 | (** * Generalization *) 160 | 161 | (** 162 | [gen_eq c as x H] takes all occurrences of [c] in the current goal's 163 | conclusion, replaces them by the variable [x], and introduces the equality 164 | [x = c] as the hypothesis H. Useful if one needs to generalize the goal 165 | prior to applying an induction tactic. 166 | *) 167 | 168 | Tactic Notation "gen_eq" constr(c) "as" ident(x) ident(H) := 169 | set (x := c); assert (H : x = c) by reflexivity; clearbody x. 170 | 171 | (** 172 | A variation on the above in which all occurrences of [c] in the goal are 173 | replaced, not only those in the conclusion. 174 | *) 175 | 176 | Tactic Notation "gen_eq" constr(c) "as" ident(x) := 177 | set (x := c) in *; 178 | let H := fresh in (assert (H : x = c) by reflexivity; clearbody x; revert H). 179 | 180 | (** [gen] is a shortname for the [generalize dependent] tactic. *) 181 | 182 | Tactic Notation "gen" ident(X1) := 183 | generalize dependent X1. 184 | Tactic Notation "gen" ident(X1) ident(X2) := 185 | gen X2; gen X1. 186 | Tactic Notation "gen" ident(X1) ident(X2) ident(X3) := 187 | gen X3; gen X2; gen X1. 188 | Tactic Notation "gen" ident(X1) ident(X2) ident(X3) ident(X4) := 189 | gen X4; gen X3; gen X2; gen X1. 190 | Tactic Notation "gen" ident(X1) ident(X2) ident(X3) ident(X4) ident(X5) := 191 | gen X5; gen X4; gen X3; gen X2; gen X1. 192 | 193 | 194 | (* ********************************************************************** *) 195 | (** * Splitting N-ary Conjonctions *) 196 | 197 | (** [split3] and [split4] respectively destruct a triple and a quadruple 198 | of propositions. *) 199 | 200 | Tactic Notation "split3" := 201 | split; [ idtac | split ]. 202 | Tactic Notation "split4" := 203 | split; [ idtac | split3 ]. 204 | 205 | (** [splits] calls [split] recursively as long as possible. *) 206 | 207 | Tactic Notation "splits" := 208 | repeat split. 209 | 210 | (** [esplitN] are iterated version of [esplit], used to introduce 211 | uninstanciated variables in goal of the form [exists x, P x]. *) 212 | 213 | Tactic Notation "esplit2" := 214 | esplit; esplit. 215 | Tactic Notation "esplit3" := 216 | esplit; esplit; esplit. 217 | Tactic Notation "esplit4" := 218 | esplit; esplit; esplit; esplit. 219 | 220 | 221 | (* ********************************************************************** *) 222 | (** * Branching N-ary Disjunction *) 223 | 224 | (** Short-hand tactics for branching when the goal is of the form 225 | [P1 \/ P2 \/ P3]. *) 226 | 227 | Tactic Notation "or_31" := left. 228 | Tactic Notation "or_32" := right; left. 229 | Tactic Notation "or_33" := right; right. 230 | 231 | 232 | (* ********************************************************************** *) 233 | (** * Destructing conjonctions behind implications *) 234 | 235 | (** [destructi T] is to be used on a [T] of the form 236 | [A1 -> .. -> AN -> X /\ Y]. It generates the [Ai] as subgoals 237 | and adds two hypotheses X and Y to the current goal. *) 238 | 239 | Tactic Notation "destructi" constr(T) := 240 | let rec doit H := 241 | match type of H with 242 | | ?P -> ?Q => let A := fresh "A" in 243 | (assert (A : P); [ idtac | doit (H A); clear A ]) 244 | | _ => first [destruct H | puts H] 245 | end in doit T. 246 | 247 | Tactic Notation "destructi" constr(T) "as" simple_intropattern(I) := 248 | let rec doit H := 249 | match type of H with 250 | | ?P -> ?Q => let A := fresh "A" in 251 | (assert (A : P); [ idtac | doit (H A); clear A ]) 252 | | _ => first [destruct H as I | poses I H] 253 | end in doit T. 254 | 255 | (** [destructs T] calls [destruct] recursively on [T] as long as possible *) 256 | 257 | Ltac destructs H := 258 | let X := fresh in let Y := fresh in 259 | first [ destruct H as [X Y]; destructs X; destructs Y | idtac ]. 260 | 261 | 262 | (* ********************************************************************** *) 263 | (** * Introduction *) 264 | 265 | (** [introv] is used to repeat intro on all dependent variables; basically 266 | it introduces all the names which are mentionned further in the goal. *) 267 | 268 | Tactic Notation "introv" := 269 | let rec go _ := match goal with 270 | | |- ?P -> ?Q => idtac 271 | | |- forall _, _ => intro; try go tt 272 | end in first [ go tt | intro; go tt ]. 273 | 274 | Tactic Notation "introv" simple_intropattern(I) := 275 | introv; intros I. 276 | Tactic Notation "introv" simple_intropattern(I1) ident(I2) := 277 | introv; intros I1 I2. 278 | Tactic Notation "introv" simple_intropattern(I1) ident(I2) ident(I3) := 279 | introv; intros I1 I2 I3. 280 | Tactic Notation "introv" simple_intropattern(I1) ident(I2) ident(I3) ident(I4) := 281 | introv; intros I1 I2 I3 I4. 282 | Tactic Notation "introv" simple_intropattern(I1) ident(I2) ident(I3) ident(I4) ident(I5) := 283 | introv; intros I1 I2 I3 I4 I5. 284 | 285 | 286 | (* ********************************************************************** *) 287 | (** * Exists *) 288 | 289 | (** [exists T1 ... TN] is a shorthand for [exists T1; ...; exists TN]. *) 290 | 291 | Tactic Notation "exists" constr(T1) := 292 | exists T1. 293 | Tactic Notation "exists" constr(T1) constr(T2) := 294 | exists T1; exists T2. 295 | Tactic Notation "exists" constr(T1) constr(T2) constr(T3) := 296 | exists T1; exists T2; exists T3. 297 | Tactic Notation "exists" constr(T1) constr(T2) constr(T3) constr(T4) := 298 | exists T1; exists T2; exists T3; exists T4. 299 | 300 | 301 | (* ********************************************************************** *) 302 | (** * Forward Chaining - Adapted from a suggestion by Xavier Leroy *) 303 | 304 | Lemma modus_ponens : forall (P Q : Prop), 305 | P -> (P -> Q) -> Q. 306 | Proof. auto. Qed. 307 | 308 | Arguments modus_ponens [P Q]. 309 | 310 | Tactic Notation "forward" constr(x) "as" simple_intropattern(H) := 311 | (refine (modus_ponens (x _ _ _ _ _ _ _ _ _) _); [ | | | | | | | | | intros H ]) 312 | || (refine (modus_ponens (x _ _ _ _ _ _ _ _) _); [ | | | | | | | | intros H ]) 313 | || (refine (modus_ponens (x _ _ _ _ _ _ _) _); [ | | | | | | | intros H ]) 314 | || (refine (modus_ponens (x _ _ _ _ _ _) _); [ | | | | | | intros H ]) 315 | || (refine (modus_ponens (x _ _ _ _ _) _); [ | | | | | intros H ]) 316 | || (refine (modus_ponens (x _ _ _ _) _); [ | | | | intros H ]) 317 | || (refine (modus_ponens (x _ _ _) _); [ | | | intros H ]) 318 | || (refine (modus_ponens (x _ _) _); [ | | intros H ]) 319 | || (refine (modus_ponens (x _) _); [ | intros H ]) 320 | || (refine (modus_ponens (x _ _ _ _ _ _ _ _ _) _); [ | | | | | | | | intros H ]) 321 | || (refine (modus_ponens (x _ _ _ _ _ _ _ _) _); [ | | | | | | | intros H ]) 322 | || (refine (modus_ponens (x _ _ _ _ _ _ _) _); [ | | | | | | intros H ]) 323 | || (refine (modus_ponens (x _ _ _ _ _ _) _); [ | | | | | intros H ]) 324 | || (refine (modus_ponens (x _ _ _ _ _) _); [ | | | | intros H ]) 325 | || (refine (modus_ponens (x _ _ _ _) _); [ | | | intros H ]) 326 | || (refine (modus_ponens (x _ _ _) _); [ | | intros H ]) 327 | || (refine (modus_ponens (x _ _) _); [ | intros H ]) 328 | || (refine (modus_ponens (x _) _); [ intros H ]). 329 | 330 | Tactic Notation "forward" constr(x) := 331 | refine (modus_ponens (x _ _ _ _ _ _ _ _ _ _) _) 332 | || refine (modus_ponens (x _ _ _ _ _ _ _ _ _) _) 333 | || refine (modus_ponens (x _ _ _ _ _ _ _ _) _) 334 | || refine (modus_ponens (x _ _ _ _ _ _ _) _) 335 | || refine (modus_ponens (x _ _ _ _ _ _) _) 336 | || refine (modus_ponens (x _ _ _ _ _) _) 337 | || refine (modus_ponens (x _ _ _ _) _) 338 | || refine (modus_ponens (x _ _ _) _) 339 | || refine (modus_ponens (x _ _) _) 340 | || refine (modus_ponens (x _) _). 341 | 342 | 343 | (* ********************************************************************** *) 344 | (** * Tactics with Automation *) 345 | 346 | (** The name of a tactic followed by a star means: apply the tactic, then 347 | applies [auto*] on the generated subgoals. [auto*] is a tactic 348 | which tries to solve the goal with either auto or intuition eauto. 349 | It leaves the goal unchanged if it can't solve the goal. *) 350 | 351 | (** Exceptions to the naming convention are: [take] which stands for [exists*] 352 | and [use] which stands for [puts*]. Exceptions to the behaviour for 353 | [asserts*] which only calls [auto*] in the new subgoal, and [apply*] 354 | which first tries [apply] and if it fails it tries [eapply] and then 355 | in both cases calls [auto*]. *) 356 | 357 | Tactic Notation "auto" "*" := 358 | try solve [ auto | intuition (eauto with *) ]. 359 | Tactic Notation "auto" "*" int_or_var(n) := 360 | try solve [ auto | intuition (eauto n with *) ]. 361 | Tactic Notation "asserts" "*" ident(H) constr(E) := 362 | assert (H : E); [ auto* | idtac ]. 363 | Tactic Notation "apply" "*" constr(H) := 364 | first [ apply H | eapply H ]; auto*. 365 | Tactic Notation "apply" "*" constr(H) := 366 | first [ apply H | eapply H ]; auto*. 367 | Tactic Notation "contradictions" "*" := 368 | contradictions; auto*. 369 | Tactic Notation "destruct" "*" constr(H) := 370 | destruct H; auto*. 371 | Tactic Notation "destruct" "*" constr(H) "as" simple_intropattern(I) := 372 | destruct H as I; auto*. 373 | Tactic Notation "f_equal" "*" := 374 | f_equal; auto*. 375 | Tactic Notation "induction" "*" constr(H) := 376 | induction H; auto*. 377 | Tactic Notation "inversion" "*" constr(H) := 378 | inversion H; auto*. 379 | Tactic Notation "inversions" "*" constr(H) := 380 | inversions H; auto*. 381 | Tactic Notation "rewrite" "*" constr(H) := 382 | rewrite H; auto*. 383 | Tactic Notation "rewrite" "*" "<-" constr(H) := 384 | rewrite <- H; auto*. 385 | Tactic Notation "do_rew" "*" constr(E) tactic(T) := 386 | (do_rew E T); auto*. 387 | Tactic Notation "do_rew" "*" "<-" constr(E) tactic(T) := 388 | (do_rew <- E T); auto*. 389 | Tactic Notation "do_rew_2" "*" constr(E) tactic(T) := 390 | (do_rew_2 E T); auto*. 391 | Tactic Notation "do_rew_2" "*" "<-" constr(E) tactic(T) := 392 | (do_rew_2 <- E T); auto*. 393 | Tactic Notation "simpl" "*" := 394 | simpl; auto*. 395 | Tactic Notation "simpls" "*" := 396 | simpls; auto*. 397 | Tactic Notation "unsimpl" "*" constr(E) := 398 | unsimpl E; auto*. 399 | Tactic Notation "unsimpl" "*" constr(E) "in" hyp(H) := 400 | unsimpl E in H; auto*. 401 | Tactic Notation "split2" "*" := 402 | split; auto*. 403 | Tactic Notation "split3" "*" := 404 | split3; auto*. 405 | Tactic Notation "split4" "*" := 406 | split4; auto*. 407 | Tactic Notation "splits" "*" := 408 | splits; auto*. 409 | Tactic Notation "esplit2" "*" := 410 | esplit2; auto*. 411 | Tactic Notation "esplit3" "*" := 412 | esplit3; auto*. 413 | Tactic Notation "esplit4" "*" := 414 | esplit4; auto*. 415 | Tactic Notation "right" "*" := 416 | right; auto*. 417 | Tactic Notation "left" "*" := 418 | left; auto*. 419 | Tactic Notation "or_31" "*" := 420 | or_31; auto*. 421 | Tactic Notation "or_32" "*" := 422 | or_32; auto*. 423 | Tactic Notation "or_33" "*" := 424 | or_33; auto*. 425 | Tactic Notation "destructi" "*" constr(H) := 426 | destructi H; auto*. 427 | Tactic Notation "subst" "*" := 428 | subst; auto*. 429 | Tactic Notation "use" constr(expr) := 430 | puts expr; auto*. 431 | Tactic Notation "use" constr(expr1) constr(expr2) := 432 | puts expr1; use expr2. 433 | Tactic Notation "use" constr(expr1) constr(expr2) constr(expr3) := 434 | puts expr1; use expr2 expr3. 435 | Tactic Notation "exists" "*" constr(T1) := 436 | exists T1; auto*. 437 | Tactic Notation "exists" "*" constr(T1) constr(T2) := 438 | exists T1 T2; auto*. 439 | Tactic Notation "exists" "*" constr(T1) constr(T2) constr(T3) := 440 | exists T1 T2 T3. 441 | Tactic Notation "exists" "*" constr(T1) constr(T2) constr(T3) constr(T4) := 442 | exists T1 T2 T3 T4. 443 | Tactic Notation "forward" "*" constr(x) "as" simple_intropattern(H) := 444 | forward x; auto*. 445 | Tactic Notation "forward" "*" constr(x) := 446 | forward x; auto*. 447 | 448 | 449 | (* ********************************************************************** *) 450 | (** * Tactics with Limited Automation *) 451 | 452 | Tactic Notation "rewrite" "~" constr(H) := 453 | rewrite H; auto. 454 | Tactic Notation "rewrite" "~" "<-" constr(H) := 455 | rewrite <- H; auto. 456 | Tactic Notation "apply" "~" constr(H) := 457 | first [ apply H | eapply H ]; auto. 458 | Tactic Notation "destructi" "~" constr(H) := 459 | destructi H; auto. 460 | Tactic Notation "destruct" "~" constr(H) := 461 | destruct H; auto. 462 | Tactic Notation "destruct" "~" constr(H) "as" simple_intropattern(I) := 463 | destruct H as I; auto. 464 | Tactic Notation "destructi" "~" constr(H) "as" simple_intropattern(I) := 465 | destructi H as I; auto. 466 | Tactic Notation "split2" "~" := 467 | split; auto. 468 | Tactic Notation "split3" "~" := 469 | split3; auto. 470 | Tactic Notation "split4" "~" := 471 | split4; auto. 472 | Tactic Notation "splits" "~" := 473 | splits; auto. 474 | Tactic Notation "forward" "~" constr(x) "as" simple_intropattern(H) := 475 | forward x as H; auto. 476 | Tactic Notation "forward" "~" constr(x) := 477 | forward x; auto. 478 | 479 | 480 | (* ********************************************************************** *) 481 | (** * Projections *) 482 | 483 | (** Short-hand notations for projections from triples. *) 484 | 485 | Notation "'proj31' P" := (proj1 P) (at level 69). 486 | Notation "'proj32' P" := (proj1 (proj2 P)) (at level 69). 487 | Notation "'proj33' P" := (proj2 (proj2 P)) (at level 69). 488 | 489 | Notation "'proj41' P" := (proj1 P) (at level 69). 490 | Notation "'proj42' P" := (proj1 (proj2 P)) (at level 69). 491 | Notation "'proj43' P" := (proj1 (proj2 (proj2 P))) (at level 69). 492 | Notation "'proj44' P" := (proj2 (proj2 (proj2 P))) (at level 69). 493 | 494 | 495 | 496 | 497 | -------------------------------------------------------------------------------- /ML_SP_Definitions.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Preservation and Progress for mini-ML (CBV) - Definitions * 3 | * Arthur Chargueraud, March 2007, Coq v8.1 * 4 | * Extension to structural polymorphism * 5 | * Jacques Garrigue, October 2007 - May 2008 * 6 | ***************************************************************************) 7 | 8 | Set Implicit Arguments. 9 | Require Import Metatheory List Arith. 10 | 11 | (* Constraint domain specification *) 12 | 13 | Module Type CstrIntf. 14 | Parameter cstr attr : Set. 15 | Parameter valid : cstr -> Prop. 16 | Parameter valid_dec : forall c, sumbool (valid c) (~valid c). 17 | Parameter eq_dec : forall x y:attr, {x=y}+{x<>y}. 18 | Parameter unique : cstr -> attr -> bool. 19 | Parameter lub : cstr -> cstr -> cstr. 20 | Parameter entails : cstr -> cstr -> Prop. 21 | Parameter entails_refl : forall c, entails c c. 22 | Hint Resolve entails_refl : core. 23 | Parameter entails_trans : forall c1 c2 c3, 24 | entails c1 c2 -> entails c2 c3 -> entails c1 c3. 25 | Parameter entails_lub : forall c1 c2 c, 26 | (entails c c1 /\ entails c c2) <-> entails c (lub c1 c2). 27 | Parameter entails_unique : forall c1 c2 v, 28 | entails c1 c2 -> unique c2 v = true -> unique c1 v = true. 29 | Parameter entails_valid : forall c1 c2, 30 | entails c1 c2 -> valid c1 -> valid c2. 31 | End CstrIntf. 32 | 33 | Module Type CstIntf. 34 | Parameter const : Set. 35 | Parameter arity : const -> nat. 36 | End CstIntf. 37 | 38 | Declare Scope typ_scope. 39 | 40 | (* Parameterized definitions *) 41 | 42 | Module MkDefs(Cstr:CstrIntf)(Const:CstIntf). 43 | 44 | (* ********************************************************************** *) 45 | (** ** Description of types *) 46 | 47 | (** Grammar of types. *) 48 | 49 | Inductive typ : Set := 50 | | typ_bvar : nat -> typ 51 | | typ_fvar : var -> typ 52 | | typ_arrow : typ -> typ -> typ. 53 | 54 | (** Types are inhabited, giving us a default value. *) 55 | 56 | Definition typ_def := typ_bvar 0. 57 | 58 | (** Constraint domain *) 59 | 60 | Definition coherent kc (kr:list(Cstr.attr*typ)) := forall x T U, 61 | Cstr.unique kc x = true -> In (x,T) kr -> In (x,U) kr -> T = U. 62 | 63 | Record ckind : Set := Kind { 64 | kind_cstr : Cstr.cstr; 65 | kind_valid : Cstr.valid kind_cstr; 66 | kind_rel : list (Cstr.attr*typ); 67 | kind_coherent : coherent kind_cstr kind_rel }. 68 | 69 | Definition kind := option ckind. 70 | 71 | Definition entails K K' := 72 | Cstr.entails (kind_cstr K) (kind_cstr K') /\ 73 | forall T:Cstr.attr*typ, In T (kind_rel K') -> In T (kind_rel K). 74 | 75 | (** Type schemes. *) 76 | 77 | Record sch : Set := Sch { 78 | sch_type : typ ; 79 | sch_kinds : list kind }. 80 | 81 | Notation "'sch_arity' M" := 82 | (length (sch_kinds M)) (at level 20, no associativity). 83 | 84 | (** Opening body of type schemes. *) 85 | 86 | Fixpoint typ_open (T : typ) (Vs : list typ) {struct T} : typ := 87 | match T with 88 | | typ_bvar i => nth i Vs typ_def 89 | | typ_fvar x => typ_fvar x 90 | | typ_arrow T1 T2 => typ_arrow (typ_open T1 Vs) (typ_open T2 Vs) 91 | end. 92 | 93 | (** Opening body of type schemes with variables *) 94 | 95 | Definition typ_fvars := 96 | List.map typ_fvar. 97 | 98 | Definition typ_open_vars T Xs := 99 | typ_open T (typ_fvars Xs). 100 | 101 | (** Instanciation of a type scheme *) 102 | 103 | Definition sch_open M := 104 | typ_open (sch_type M). 105 | 106 | Definition sch_open_vars M := 107 | typ_open_vars (sch_type M). 108 | 109 | Notation "M ^^ Vs" := (sch_open M Vs) 110 | (at level 67, only parsing) : typ_scope. 111 | Notation "M ^ Xs" := 112 | (sch_open_vars M Xs) (only parsing) : typ_scope. 113 | 114 | Bind Scope typ_scope with typ. 115 | Open Scope typ_scope. 116 | 117 | (** Locally closed types *) 118 | 119 | Inductive type : typ -> Prop := 120 | | type_fvar : forall X, 121 | type (typ_fvar X) 122 | | type_arrow : forall T1 T2, 123 | type T1 -> 124 | type T2 -> 125 | type (typ_arrow T1 T2). 126 | 127 | (** List of n locally closed types *) 128 | 129 | Definition types := list_for_n type. 130 | 131 | (** Iterating and opening kinds *) 132 | 133 | Definition kind_types (K:kind) := 134 | match K with 135 | | None => nil 136 | | Some k => list_snd (kind_rel k) 137 | end. 138 | 139 | Definition All_kind_types (P:typ->Prop) K := 140 | list_forall P (kind_types K). 141 | 142 | Lemma map_coherent : forall f kc kr, 143 | coherent kc kr -> 144 | coherent kc (map_snd f kr). 145 | Proof. 146 | intros. intro; intros. 147 | puts (H x). 148 | destruct (map_snd_inv _ _ _ _ H1) as [T' [Heq Hin]]. 149 | destruct (map_snd_inv _ _ _ _ H2) as [U' [Heq' Hin']]. 150 | subst. 151 | rewrite* (H3 T' U'). 152 | Qed. 153 | 154 | Definition ckind_map_spec (f:typ->typ) (k:ckind): 155 | {k' | kind_cstr k = kind_cstr k' /\ 156 | kind_rel k' = map_snd f (kind_rel k)}. 157 | Proof. 158 | intros. 159 | destruct k as [kc kv kr kh]. 160 | exists (Kind kv (map_coherent f kh)). 161 | simpl. auto. 162 | Defined. 163 | 164 | Definition ckind_map f k := proj1_sig (ckind_map_spec f k). 165 | 166 | Definition kind_map f (K:kind) : kind := 167 | match K with 168 | | None => None 169 | | Some k => Some (ckind_map f k) 170 | end. 171 | 172 | Definition kind_open K Vs := kind_map (fun T => typ_open T Vs) K. 173 | 174 | (** Body of a scheme *) 175 | 176 | Definition typ_body T Ks := 177 | forall Xs, length Ks = length Xs -> 178 | type (typ_open_vars T Xs) /\ 179 | list_forall (All_kind_types (fun T' => type (typ_open_vars T' Xs))) Ks. 180 | 181 | (** Definition of a well-formed scheme *) 182 | 183 | Definition scheme M := 184 | typ_body (sch_type M) (sch_kinds M). 185 | 186 | (* ********************************************************************** *) 187 | (** ** Description of terms *) 188 | 189 | (** Grammar of terms. *) 190 | 191 | Inductive trm : Set := 192 | | trm_bvar : nat -> trm 193 | | trm_fvar : var -> trm 194 | | trm_abs : trm -> trm 195 | | trm_let : trm -> trm -> trm 196 | | trm_app : trm -> trm -> trm 197 | | trm_cst : Const.const -> trm. 198 | 199 | (** Opening term binders. *) 200 | 201 | Fixpoint trm_open_rec (k : nat) (u : trm) (t : trm) {struct t} : trm := 202 | match t with 203 | | trm_bvar i => if k === i then u else (trm_bvar i) 204 | | trm_fvar x => trm_fvar x 205 | | trm_abs t1 => trm_abs (trm_open_rec (S k) u t1) 206 | | trm_let t1 t2 => trm_let (trm_open_rec k u t1) (trm_open_rec (S k) u t2) 207 | | trm_app t1 t2 => trm_app (trm_open_rec k u t1) (trm_open_rec k u t2) 208 | | trm_cst c => trm_cst c 209 | end. 210 | 211 | Definition trm_open t u := trm_open_rec 0 u t. 212 | 213 | (* Notation "{ k ~> u } t" := (trm_open_rec k u t) (at level 67). *) 214 | Notation "t ^^ u" := (trm_open t u) (at level 67). 215 | Notation "t ^ x" := (trm_open t (trm_fvar x)). 216 | 217 | (** Locally closed termessions *) 218 | 219 | Inductive term : trm -> Prop := 220 | | term_var : forall x, 221 | term (trm_fvar x) 222 | | term_abs : forall L t1, 223 | (forall x, x \notin L -> term (t1 ^ x)) -> 224 | term (trm_abs t1) 225 | | term_let : forall L t1 t2, 226 | term t1 -> 227 | (forall x, x \notin L -> term (t2 ^ x)) -> 228 | term (trm_let t1 t2) 229 | | term_app : forall t1 t2, 230 | term t1 -> 231 | term t2 -> 232 | term (trm_app t1 t2) 233 | | term_cst : forall c, 234 | term (trm_cst c). 235 | 236 | (** Definition of the body of an abstraction *) 237 | 238 | Definition term_body t := 239 | exists L, forall x, x \notin L -> term (t ^ x). 240 | 241 | (** Term instanciation *) 242 | 243 | Definition trm_def := trm_bvar 0. 244 | 245 | Fixpoint trm_inst_rec (k : nat) (tl : list trm) (t : trm) {struct t} : trm := 246 | match t with 247 | | trm_bvar i => if le_lt_dec k i then nth (i-k) tl t else trm_bvar i 248 | | trm_fvar x => trm_fvar x 249 | | trm_abs t1 => trm_abs (trm_inst_rec (S k) tl t1) 250 | | trm_let t1 t2 => trm_let (trm_inst_rec k tl t1) (trm_inst_rec (S k) tl t2) 251 | | trm_app t1 t2 => trm_app (trm_inst_rec k tl t1) (trm_inst_rec k tl t2) 252 | | trm_cst c => trm_cst c 253 | end. 254 | 255 | Definition trm_inst t tl := trm_inst_rec 0 tl t. 256 | 257 | (** Applying a constant *) 258 | 259 | Definition const_app c vl := fold_left trm_app vl (trm_cst c). 260 | 261 | (* ********************************************************************** *) 262 | (** ** Description of typing *) 263 | 264 | (** Definition of kinding environments *) 265 | 266 | Definition kenv := env kind. 267 | 268 | Definition kenv_ok K := 269 | ok K /\ env_prop (All_kind_types type) K. 270 | 271 | (** Proper instanciation *) 272 | 273 | Inductive well_kinded : kenv -> kind -> typ -> Prop := 274 | | wk_any : forall K T, 275 | well_kinded K None T 276 | | wk_kind : forall k' K k x, 277 | binds x (Some k') K -> 278 | entails k' k -> 279 | well_kinded K (Some k) (typ_fvar x). 280 | 281 | Hint Constructors well_kinded : core. 282 | 283 | Definition kinds_open Ks Us := 284 | List.map (fun k:kind => kind_open k Us) Ks. 285 | 286 | Definition proper_instance K Ks Us := 287 | types (length Ks) Us /\ 288 | list_forall2 (well_kinded K) (kinds_open Ks Us) Us. 289 | 290 | Definition kinds_open_vars Ks Xs := 291 | List.combine Xs (kinds_open Ks (typ_fvars Xs)). 292 | 293 | (** Definition of typing environments *) 294 | 295 | Definition env := env sch. 296 | 297 | Definition env_ok (E:env) := ok E /\ env_prop scheme E. 298 | 299 | (** Computing free variables of a type. *) 300 | 301 | Fixpoint typ_fv (T : typ) {struct T} : vars := 302 | match T with 303 | | typ_bvar i => {} 304 | | typ_fvar x => {{x}} 305 | | typ_arrow T1 T2 => (typ_fv T1) \u (typ_fv T2) 306 | end. 307 | 308 | (** Computing free variables of a list of terms. *) 309 | 310 | Definition typ_fv_list := 311 | List.fold_right (fun t acc => typ_fv t \u acc) {}. 312 | 313 | (** Computing free variables of a kind. *) 314 | 315 | Definition kind_fv k := 316 | typ_fv_list (kind_types k). 317 | 318 | Definition kind_fv_list := 319 | List.fold_right (fun t acc => kind_fv t \u acc) {}. 320 | 321 | (** Computing free variables of a type scheme. *) 322 | 323 | Definition sch_fv M := 324 | typ_fv (sch_type M) \u kind_fv_list (sch_kinds M). 325 | 326 | (** Computing free type variables of the values of an environment. *) 327 | 328 | Definition env_fv := 329 | fv_in sch_fv. 330 | 331 | (** Grammar of values *) 332 | 333 | Inductive valu : nat -> trm -> Prop := 334 | | value_abs : forall t1, term (trm_abs t1) -> valu 0 (trm_abs t1) 335 | | value_cst : forall c, valu (Const.arity c) (trm_cst c) 336 | | value_app : forall n t1 n2 t2, 337 | valu (S n) t1 -> 338 | valu n2 t2 -> 339 | valu n (trm_app t1 t2). 340 | 341 | Definition value t := exists n, valu n t. 342 | 343 | (** Another functor for delta-rules *) 344 | 345 | Module Type DeltaIntf. 346 | Parameter type : Const.const -> sch. 347 | Parameter closed : forall c, sch_fv (type c) = {}. 348 | Parameter scheme : forall c, scheme (type c). 349 | Parameter reduce : forall c tl, 350 | list_for_n value (S(Const.arity c)) tl -> trm. 351 | Parameter term : forall c tl vl, 352 | term (@reduce c tl vl). 353 | End DeltaIntf. 354 | 355 | Module MkJudge(Delta:DeltaIntf). 356 | 357 | (** The typing judgment *) 358 | 359 | Reserved Notation "K ; E | gc |= t ~: T" (at level 69). 360 | 361 | Inductive gc_kind : Set := GcAny | GcLet. 362 | Definition gc_info : Set := (bool * gc_kind)%type. 363 | Fixpoint gc_ok (gc:gc_info) := fst gc = true. 364 | Fixpoint gc_raise (gc:gc_info) : gc_info := 365 | match snd gc with 366 | | GcLet => (true, GcLet) 367 | | _ => gc 368 | end. 369 | Fixpoint gc_lower (gc:gc_info) : gc_info := 370 | match snd gc with 371 | | GcLet => (false, GcLet) 372 | | _ => gc 373 | end. 374 | 375 | Inductive typing(gc:gc_info) : kenv -> env -> trm -> typ -> Prop := 376 | | typing_var : forall K E x M Us, 377 | kenv_ok K -> 378 | env_ok E -> 379 | binds x M E -> 380 | proper_instance K (sch_kinds M) Us -> 381 | K ; E | gc |= (trm_fvar x) ~: (M ^^ Us) 382 | | typing_abs : forall L K E U T t1, 383 | type U -> 384 | (forall x, x \notin L -> 385 | K ; (E & x ~ Sch U nil) | gc_raise gc |= (t1 ^ x) ~: T) -> 386 | K ; E | gc |= (trm_abs t1) ~: (typ_arrow U T) 387 | | typing_let : forall M L1 L2 K E T2 t1 t2, 388 | (forall Xs, fresh L1 (sch_arity M) Xs -> 389 | (K & kinds_open_vars (sch_kinds M) Xs); E | gc_raise gc |= 390 | t1 ~: (M ^ Xs)) -> 391 | (forall x, x \notin L2 -> 392 | K ; (E & x ~ M) | gc_raise gc |= (t2 ^ x) ~: T2) -> 393 | K ; E | gc |= (trm_let t1 t2) ~: T2 394 | | typing_app : forall K E S T t1 t2, 395 | K ; E | gc_lower gc |= t1 ~: (typ_arrow S T) -> 396 | K ; E | gc_lower gc |= t2 ~: S -> 397 | K ; E | gc |= (trm_app t1 t2) ~: T 398 | | typing_cst : forall K E Us c, 399 | kenv_ok K -> 400 | env_ok E -> 401 | proper_instance K (sch_kinds (Delta.type c)) Us -> 402 | K ; E | gc |= (trm_cst c) ~: (Delta.type c ^^ Us) 403 | | typing_gc : forall Ks L K E t T, 404 | gc_ok gc -> 405 | (forall Xs, fresh L (length Ks) Xs -> 406 | K & kinds_open_vars Ks Xs; E | gc |= t ~: T) -> 407 | K ; E | gc |= t ~: T 408 | 409 | where "K ; E | gc |= t ~: T" := (typing gc K E t T). 410 | 411 | 412 | (* ********************************************************************** *) 413 | (** ** Description of the semantics *) 414 | 415 | (** Reduction rules *) 416 | 417 | Inductive red : trm -> trm -> Prop := 418 | | red_abs : forall t1 t2, 419 | term (trm_abs t1) -> 420 | value t2 -> 421 | red (trm_app (trm_abs t1) t2) (t1 ^^ t2) 422 | | red_let : forall t1 t2, 423 | term (trm_let t1 t2) -> 424 | value t1 -> 425 | red (trm_let t1 t2) (t2 ^^ t1) 426 | | red_delta : forall c tl vl, 427 | red (const_app c tl) (@Delta.reduce c tl vl) 428 | | red_let_1 : forall t1 t1' t2, 429 | term_body t2 -> 430 | red t1 t1' -> 431 | red (trm_let t1 t2) (trm_let t1' t2) 432 | | red_app_1 : forall t1 t1' t2, 433 | value t2 -> 434 | red t1 t1' -> 435 | red (trm_app t1 t2) (trm_app t1' t2) 436 | | red_app_2 : forall t1 t2 t2', 437 | term t1 -> 438 | red t2 t2' -> 439 | red (trm_app t1 t2) (trm_app t1 t2'). 440 | 441 | Notation "t --> t'" := (red t t') (at level 68). 442 | 443 | 444 | (* ********************************************************************** *) 445 | (** ** Description of the results *) 446 | 447 | (** Goal is to prove preservation and progress *) 448 | 449 | Definition preservation := forall K E t t' T, 450 | K ; E | (true,GcAny) |= t ~: T -> 451 | t --> t' -> 452 | K ; E | (true,GcAny) |= t' ~: T. 453 | 454 | Definition progress := forall K t T, 455 | K ; empty | (true,GcAny) |= t ~: T -> 456 | value t 457 | \/ exists t', t --> t'. 458 | 459 | End MkJudge. 460 | 461 | End MkDefs. 462 | -------------------------------------------------------------------------------- /ML_SP_Extraction.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Code extraction for structural polymorphism * 3 | * Jacques Garrigue, October 2007 - Jun 2009 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | 8 | Require Import List Arith Metatheory ML_SP_Domain. 9 | Require Omega. 10 | Import Infer2. 11 | Import MyEval2. 12 | Import Sound3. 13 | Import Infer.Unify. 14 | Import MyEval. 15 | Import Rename.Sound.Infra. 16 | Import Defs. 17 | Import Rename2.Sound2.JudgInfra. 18 | Import Judge. 19 | Import Infer2. 20 | 21 | Definition t := 22 | trm_app 23 | (trm_cst (Const.matches (NoDup_nodup (5 :: nil)))) 24 | (trm_abs (trm_bvar O)). 25 | 26 | (* This doesn't seem to work inside Coq (some things don't get evaluated) *) 27 | (* Eval compute in typinf' t. *) 28 | 29 | Definition decidable (A : Set) (P : A -> Prop) := 30 | forall x, {P x} + {~P x}. 31 | 32 | Definition ok_dec : decidable (@ok sch). 33 | intro E; induction E; env_fix. 34 | auto. 35 | destruct a; env_fix. 36 | destruct IHE. 37 | case_eq (get v E); intros. 38 | right*; intro. 39 | elim (binds_fresh H). 40 | inversions* H0. 41 | auto. 42 | right*. 43 | Defined. 44 | 45 | Inductive type_n (n:nat) : typ -> Prop := 46 | | typn_bvar : forall m, m < n -> type_n n (typ_bvar m) 47 | | typn_fvar : forall x, type_n n (typ_fvar x) 48 | | typn_arrow : forall T1 T2, 49 | type_n n T1 -> type_n n T2 -> type_n n (typ_arrow T1 T2). 50 | Hint Constructors type_n : core. 51 | 52 | Definition type_n_dec : forall n , decidable (type_n n). 53 | introv T; induction T; simpl; auto. 54 | destruct~ (le_lt_dec n n0). 55 | right; intro. inversions H. Omega.omega. 56 | destruct IHT1. 57 | destruct* IHT2. 58 | right; intro. inversions* H. 59 | right; intro. inversions* H. 60 | Defined. 61 | 62 | Lemma type_n_typ_body : forall T Xs, 63 | (type_n (length Xs) T <-> type (typ_open_vars T Xs)). 64 | Proof. 65 | unfold typ_open_vars. 66 | intros; split. 67 | induction 1; simpl*. 68 | apply* Delta.type_nth_typ_vars. 69 | induction T; simpl; intros; auto. 70 | destruct~ (le_lt_dec (length Xs) n). 71 | unfold typ_fvars in H. 72 | rewrite <- (length_map typ_fvar) in l. 73 | rewrite (nth_overflow _ _ l) in H. inversion H. 74 | inversions* H. 75 | Qed. 76 | 77 | Definition list_forall_dec : forall (A:Set) (P:A->Prop), 78 | decidable P -> decidable (list_forall P). 79 | introv HP l; induction l. 80 | left*. 81 | destruct* (HP a). 82 | right; intro. inversion* H. 83 | Defined. 84 | 85 | Definition scheme_dec : decidable scheme. 86 | intros [T Ks]. 87 | unfold scheme, typ_body; simpl. 88 | set (n := length Ks). clearbody n. 89 | destruct (type_n_dec n T). 90 | puts (list_forall_dec (type_n_dec n)). 91 | unfold All_kind_types. 92 | puts (list_forall_dec (fun k => H (kind_types k))). 93 | destruct (H0 Ks). 94 | left*; intuition; subst. apply* (proj1 (type_n_typ_body T Xs)). 95 | apply* list_forall_imp; intros. simpl in H1. 96 | apply* list_forall_imp; intros. 97 | apply* (proj1 (type_n_typ_body x0 Xs)). 98 | right*; intro. 99 | elim n0; clear -H1. 100 | destruct (var_freshes {} n). 101 | destruct* (H1 x); clear H1. 102 | apply* list_forall_imp; intros. simpl in H1. 103 | refine (list_forall_imp _ _ H1); intros. 104 | rewrite (fresh_length _ _ _ f). 105 | refine (proj2 (type_n_typ_body x1 x) H2); auto*. 106 | right*; intro. 107 | elim n0; clear -H. 108 | destruct (var_freshes {} n). 109 | destruct* (H x); clear H. 110 | rewrite (fresh_length _ _ _ f). 111 | refine (proj2 (type_n_typ_body T x) H0); auto*. 112 | Defined. 113 | 114 | Definition env_prop_dec : forall (A:Set) (P:A->Prop), 115 | decidable P -> decidable (env_prop P). 116 | introv HP E; induction E. 117 | left*; intro; intros. elim H. 118 | destruct a; env_fix. 119 | destruct* (HP a). 120 | Defined. 121 | 122 | Lemma env_weaker_refl : forall E, Rename2.env_weaker E E. 123 | Proof. 124 | introv x; intros. 125 | exists (@nil kind). 126 | rewrite* <- app_nil_end. 127 | Qed. 128 | 129 | Definition typinf1 : forall E t, 130 | {p:kenv*typ | let (K,T) := p in K; E |Gc|= t ~: T}+ 131 | ({env_fv E <> {}}+{forall K T, ~ K;E |Gc|= t ~: T}). 132 | intros. 133 | case_eq (S.is_empty (env_fv E)); intros. 134 | assert (Hempty: env_fv E = {}). 135 | puts (S.is_empty_2 H). 136 | apply eq_ext; split2*. intro Ha; elim (H0 _ Ha). 137 | clear H; destruct (ok_dec E). 138 | destruct (env_prop_dec scheme_dec E). 139 | case_eq (typinf' E t0); intros. 140 | left*; exists p. destruct p. 141 | apply* typinf_sound'. 142 | right*; right*; introv Typ. 143 | destruct* (Rename2.typing_remove_gc Typ {} (@env_weaker_refl E)) 144 | as [K' [HK' Typ']]. 145 | destruct* (typinf_principal' Hempty Typ') as [K0 [T' [TI]]]. 146 | rewrite H in TI; discriminate. 147 | right*. 148 | right*. 149 | right*; left*. 150 | intro. rewrite H0 in H. 151 | puts (S.is_empty_1 (S.empty_1)). 152 | rewrite H1 in H; discriminate. 153 | Defined. 154 | 155 | Definition eval1 fenv t h := eval fenv h nil nil t nil. 156 | 157 | (* Export and try to do this in ocaml *) 158 | Require Import Extraction. 159 | (* Extraction Language Haskell. *) 160 | Set Extraction AccessOpaque. 161 | Extraction "typinf" typinf1 eval1. 162 | -------------------------------------------------------------------------------- /ML_SP_SoundDomain.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Preservation and Progress for mini-ML with structural polymorphism * 3 | * Jacques Garrigue, May 2008 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import Lib_FinSet Lib_FinSetImpl Metatheory List ListSet Arith. 8 | Require Import ML_SP_Soundness. 9 | Require Import Omega. 10 | 11 | Section ListSet. 12 | Variable A : Type. 13 | Hypothesis eq_dec : forall x y:A, {x = y} + {x <> y}. 14 | 15 | Definition set_incl : forall (s1 s2 : list A), 16 | sumbool (incl s1 s2) (~incl s1 s2). 17 | intros. 18 | induction s1. left; intro; simpl*. 19 | case_eq (set_mem eq_dec a s2); intros. 20 | destruct IHs1. 21 | left*; intros x Hx. simpl in Hx; destruct Hx. 22 | subst; apply* set_mem_correct1. apply* i. 23 | right*. 24 | right*. intro; elim (set_mem_complete1 eq_dec _ _ H). 25 | apply* H0. 26 | Defined. 27 | End ListSet. 28 | 29 | Module Cstr. 30 | Definition attr := nat. 31 | Definition eq_dec := eq_nat_dec. 32 | Inductive ksort := Ksum | Kprod | Kbot. 33 | Record cstr_impl : Set := C { 34 | cstr_sort : ksort; 35 | cstr_low : list nat; 36 | cstr_high : option (list nat) }. 37 | Definition cstr := cstr_impl. 38 | Definition valid c := 39 | cstr_sort c <> Kbot /\ 40 | match cstr_high c with 41 | | None => True 42 | | Some h => incl (cstr_low c) h 43 | end. 44 | Definition le_sort s1 s2 := 45 | s1 = Kbot \/ s1 = s2. 46 | Definition entails c1 c2 := 47 | le_sort (cstr_sort c1) (cstr_sort c2) /\ 48 | incl (cstr_low c2) (cstr_low c1) /\ 49 | match cstr_high c2 with 50 | | None => True 51 | | Some h2 => 52 | match cstr_high c1 with 53 | | None => False 54 | | Some h1 => incl h1 h2 55 | end 56 | end. 57 | Lemma entails_refl : forall c, entails c c. 58 | Proof. 59 | intros. 60 | split. right*. 61 | split. apply incl_refl. 62 | destruct* (cstr_high c). 63 | Qed. 64 | 65 | Lemma entails_trans : forall c1 c2 c3, 66 | entails c1 c2 -> entails c2 c3 -> entails c1 c3. 67 | Proof. 68 | introv. intros [CS1 [CL1 CH1]] [CS2 [CL2 CH2]]. 69 | split. 70 | unfold le_sort. 71 | destruct* CS1. 72 | rewrite H. destruct* CS2. 73 | split. apply* incl_tran. 74 | clear CL1 CL2. 75 | destruct* (cstr_high c3). 76 | destruct* (cstr_high c2). 77 | destruct* (cstr_high c1). 78 | Qed. 79 | 80 | Definition unique c v := set_mem eq_nat_dec v (cstr_low c). 81 | 82 | Definition sort_lub s1 s2 := 83 | match s1, s2 with 84 | | Ksum, Ksum => Ksum 85 | | Kprod, Kprod => Kprod 86 | | _, _ => Kbot 87 | end. 88 | 89 | Definition lub c1 c2 := 90 | C (sort_lub (cstr_sort c1) (cstr_sort c2)) 91 | (set_union eq_nat_dec (cstr_low c1) (cstr_low c2)) 92 | (match cstr_high c1, cstr_high c2 with 93 | | None, h => h 94 | | h, None => h 95 | | Some s1, Some s2 => Some (set_inter eq_nat_dec s1 s2) 96 | end). 97 | 98 | Lemma ksort_dec : forall s, {s=Kbot} + {s<>Kbot}. 99 | intro; destruct* s; right*; intro; discriminate. 100 | Qed. 101 | 102 | Hint Resolve incl_tran : core. 103 | Lemma entails_lub : forall c1 c2 c, 104 | (entails c c1 /\ entails c c2) <-> entails c (lub c1 c2). 105 | Proof. 106 | unfold entails, lub; simpl; split; intros. 107 | intuition. 108 | destruct (ksort_dec (cstr_sort c)). 109 | rewrite e; left*. 110 | destruct* H. 111 | destruct* H0. 112 | rewrite <- H, <- H0. 113 | right*; destruct* (cstr_sort c). 114 | intros x Hx; destruct* (set_union_elim _ _ _ _ Hx). 115 | case_rewrite R1 (cstr_high c1); 116 | case_rewrite R2 (cstr_high c2); 117 | try case_rewrite R3 (cstr_high c); intuition. 118 | intros x Hx; apply set_inter_intro. apply* H4. apply* H5. 119 | destruct H as [HS [HL HH]]. 120 | split; split; try split; 121 | try (intros x Hx; apply HL; 122 | solve [ apply* set_union_intro2 | apply* set_union_intro1 ]); 123 | try solve [ 124 | case_rewrite R1 (cstr_high c1); 125 | case_rewrite R2 (cstr_high c2); 126 | try case_rewrite R3 (cstr_high c); intuition; 127 | intros x Hx; destruct* (set_inter_elim _ _ _ _ (HH _ Hx))]; 128 | clear -HS; destruct HS; try solve [left*]; 129 | rewrite H; 130 | unfold le_sort; destruct (cstr_sort c1); destruct* (cstr_sort c2). 131 | Qed. 132 | 133 | Definition valid_dec : forall c, sumbool (valid c) (~valid c). 134 | intros. 135 | unfold valid. 136 | destruct* (ksort_dec (cstr_sort c)). 137 | case_eq (cstr_high c); intros. 138 | destruct (set_incl eq_nat_dec (cstr_low c) l). left*. 139 | right*. 140 | auto. 141 | Defined. 142 | 143 | Lemma entails_unique : forall c1 c2 v, 144 | entails c1 c2 -> unique c2 v = true -> unique c1 v = true. 145 | Proof. 146 | unfold entails, unique. 147 | intros. 148 | destruct H as [_ [H _]]. 149 | apply set_mem_correct2. 150 | apply H. 151 | apply* set_mem_correct1. 152 | Qed. 153 | 154 | Lemma entails_valid : forall c1 c2, 155 | entails c1 c2 -> valid c1 -> valid c2. 156 | Proof. 157 | unfold entails, valid. 158 | intros. 159 | destruct H as [HS [HL HH]]; destruct H0. 160 | destruct* HS. 161 | rewrite <- H1. split2*. 162 | case_rewrite R1 (cstr_high c1); case_rewrite R2 (cstr_high c2); 163 | auto*. 164 | Qed. 165 | End Cstr. 166 | 167 | Section NoDup. 168 | Fixpoint nodup (l:list nat) := 169 | match l with 170 | | nil => nil 171 | | v :: l' => 172 | if In_dec eq_nat_dec v l' then nodup l' else v :: nodup l' 173 | end. 174 | Lemma nodup_elts : forall a l, In a l <-> In a (nodup l). 175 | Proof. 176 | induction l; split2*; simpl; intro; destruct IHl. 177 | destruct H. subst. 178 | destruct* (In_dec eq_nat_dec a l). 179 | destruct* (In_dec eq_nat_dec a0 l). 180 | destruct* (In_dec eq_nat_dec a0 l). 181 | simpl in H; destruct* H. 182 | Qed. 183 | Lemma NoDup_nodup : forall l, NoDup (nodup l). 184 | Proof. 185 | induction l; simpl. constructor. 186 | destruct* (In_dec eq_nat_dec a l). 187 | constructor; auto. 188 | intro Ha. elim n. apply (proj2 (nodup_elts _ _) Ha). 189 | Qed. 190 | End NoDup. 191 | 192 | Module Const. 193 | Inductive ops : Set := 194 | | tag : Cstr.attr -> ops 195 | | matches : forall (l:list Cstr.attr), NoDup l -> ops 196 | | record : forall (l:list Cstr.attr), NoDup l -> ops 197 | | sub : Cstr.attr -> ops 198 | | recf : ops. 199 | Definition const := ops. 200 | Definition arity op := 201 | match op with 202 | | tag _ => 1 203 | | @matches l _ => length l 204 | | @record l _ => length l 205 | | sub _ => 0 206 | | recf => 1 207 | end. 208 | End Const. 209 | 210 | 211 | Module Sound := MkSound(Cstr)(Const). 212 | Import Sound. 213 | Import Infra. 214 | Import Defs. 215 | 216 | Lemma list_snd_combine : forall (A B:Set) (l1:list A) (l2:list B), 217 | length l1 = length l2 -> 218 | list_snd (combine l1 l2) = l2. 219 | Proof. 220 | induction l1; destruct l2; intros; try discriminate. reflexivity. 221 | inversions H. 222 | simpl. rewrite* (IHl1 l2). 223 | Qed. 224 | 225 | Module Delta. 226 | Definition matches_arg n := typ_arrow (typ_bvar n) (typ_bvar 1). 227 | 228 | Lemma valid_tag : forall s t, 229 | s <> Cstr.Kbot -> Cstr.valid (Cstr.C s (t::nil) None). 230 | Proof. 231 | intros. split2*. compute. auto. 232 | Qed. 233 | 234 | Lemma ksum : Cstr.Ksum <> Cstr.Kbot. 235 | intro; discriminate. 236 | Qed. 237 | 238 | Lemma kprod : Cstr.Kprod <> Cstr.Kbot. 239 | intro; discriminate. 240 | Qed. 241 | 242 | Lemma coherent_tag : forall s t, 243 | coherent (Cstr.C s (t::nil) None) ((t, typ_bvar 0) :: nil). 244 | Proof. 245 | intros s t x; intros. 246 | destruct H0; destruct H1; try contradiction. 247 | inversions H0. inversions* H1. 248 | Qed. 249 | 250 | Lemma valid_matches : forall s l, 251 | s <> Cstr.Kbot -> Cstr.valid (Cstr.C s nil (Some l)). 252 | Proof. 253 | intros; split2*. intro; simpl*. 254 | Qed. 255 | 256 | Lemma coherent_matches : forall s n l, 257 | NoDup l -> 258 | coherent (Cstr.C s nil (Some l)) 259 | (combine l (map typ_bvar (seq n (length l)))). 260 | Proof. 261 | intros; intro; intros. 262 | clear H0; revert H1 H2; gen n. 263 | induction H; simpl; intros. elim H1. 264 | destruct H1. inversions H1. 265 | destruct H2. inversions* H2. 266 | elim (H (in_combine_l _ _ _ _ H2)). 267 | destruct H2. inversions H2. 268 | elim (H (in_combine_l _ _ _ _ H1)). 269 | apply* (IHNoDup (S n)). 270 | Qed. 271 | 272 | Definition type (op:Const.const) := 273 | match op with 274 | | Const.tag t => 275 | Sch (typ_arrow (typ_bvar 0) (typ_bvar 1)) 276 | (None :: Some (Kind (valid_tag t ksum) (@coherent_tag _ t)) :: nil) 277 | | @Const.matches l ND => 278 | Sch (fold_right typ_arrow (typ_arrow (typ_bvar 0) (typ_bvar 1)) 279 | (map matches_arg (seq 2 (length l)))) 280 | (Some (Kind (@valid_matches _ l ksum) (coherent_matches 2 ND)) :: 281 | map (fun _ => None) (seq 0 (S (length l)))) 282 | | @Const.record l ND => 283 | Sch (fold_right typ_arrow (typ_bvar 0) (map typ_bvar (seq 1 (length l)))) 284 | (Some (Kind (@valid_matches _ l kprod) (coherent_matches 1 ND)) :: 285 | map (fun _ => None) (seq 0 (length l))) 286 | | Const.sub f => 287 | Sch (typ_arrow (typ_bvar 1) (typ_bvar 0)) 288 | (None :: Some (Kind (valid_tag f kprod) (@coherent_tag _ f)) :: nil) 289 | | Const.recf => 290 | let T := typ_arrow (typ_bvar 0) (typ_bvar 1) in 291 | Sch (typ_arrow (typ_arrow T T) T) (None :: None :: nil) 292 | end. 293 | 294 | Definition trm_default := trm_abs trm_def. 295 | 296 | Fixpoint record_args (t : trm) tl {struct t} : list nat * list trm := 297 | match t with 298 | | trm_app t1 t2 => record_args t1 (t2 :: tl) 299 | | trm_cst (@Const.record l _) => (l, tl) 300 | | _ => (nil, nil) 301 | end. 302 | 303 | Lemma record_args_ok : forall l nd tl tl', 304 | record_args (const_app (@Const.record l nd) tl) tl' = (l, tl++tl'). 305 | Proof. 306 | induction tl using rev_ind; simpl; intros. auto. 307 | rewrite const_app_app, <- app_assoc; simpl. 308 | apply* IHtl. 309 | Qed. 310 | 311 | Definition is_record c := 312 | match c with 313 | | Const.record _ => true 314 | | _ => false 315 | end. 316 | 317 | Lemma record_args_other : forall c tl tl', 318 | is_record c = false -> record_args (const_app c tl) tl' = (nil, nil). 319 | Proof. 320 | induction tl using rev_ind; simpl; intros. destruct* c. discriminate. 321 | rewrite const_app_app. simpl. 322 | apply* IHtl. 323 | Qed. 324 | 325 | Definition reduce c tl (vl:list_for_n value (S(Const.arity c)) tl) := 326 | match c with 327 | | Const.tag _ => trm_default 328 | | @Const.matches l nd => 329 | match nth (length l) tl trm_def with 330 | | trm_app (trm_cst (Const.tag t)) t1 => 331 | match index eq_nat_dec 0 t l with 332 | | Some i => trm_app (nth i tl trm_def) t1 333 | | None => trm_default 334 | end 335 | | _ => trm_default 336 | end 337 | | Const.record _ => trm_default 338 | | Const.sub f => 339 | match tl with 340 | | nil => trm_default 341 | | t :: _ => 342 | let (l, fl) := record_args t nil in 343 | match index eq_nat_dec 0 f l with 344 | | Some i => nth i fl trm_default 345 | | None => trm_default 346 | end 347 | end 348 | | Const.recf => 349 | match tl with 350 | | f :: a :: _ => trm_app (trm_app f (trm_app (trm_cst Const.recf) f)) a 351 | | _ => trm_default 352 | end 353 | end. 354 | 355 | Lemma term_default : term trm_default. 356 | Proof. 357 | unfold trm_default, trm_def. 358 | apply (@term_abs {}). 359 | intros. unfold trm_open; simpl. auto. 360 | Qed. 361 | Hint Resolve term_default : core. 362 | 363 | 364 | Lemma value_term : forall e, 365 | value e -> term e. 366 | Proof. 367 | intros. destruct H. induction H; auto. 368 | Qed. 369 | Hint Resolve value_term : core. 370 | 371 | Lemma term : forall c tl vl, 372 | term (@reduce c tl vl). 373 | Proof. 374 | intros. 375 | case vl; clear vl; intros. 376 | destruct c; simpl in *; auto. 377 | (* matches *) 378 | assert (length l0 < S (length l0)) by auto. 379 | rewrite e in H. 380 | case_eq (nth (length l0) tl trm_def); intros; auto. 381 | destruct* t. 382 | destruct* c. 383 | case_eq (index eq_nat_dec 0 a l0); intros; auto. 384 | apply term_app. 385 | destruct (index_ok _ 0 _ _ H1). 386 | apply value_term. 387 | apply* (list_forall_out l). 388 | apply* nth_In. 389 | assert (term (nth (length l0) tl trm_def)). 390 | apply value_term. 391 | apply* (list_forall_out l). 392 | apply* nth_In. 393 | rewrite H0 in H2. 394 | inversion* H2. 395 | (* sub *) 396 | destruct* tl. 397 | case_eq (record_args t nil); intros. 398 | assert (list_forall value l1). 399 | inversions l. 400 | clear -H H3. 401 | assert (list_forall value nil) by auto. 402 | revert H H0. generalize (@nil trm). 403 | destruct H3. 404 | induction H; simpl; intros; try solve [inversions* H0]. 405 | destruct c; inversions* H. 406 | apply* (IHvalu1 _ H1). 407 | constructor; auto. 408 | exists* n2. 409 | clear -H0. 410 | destruct* (index eq_nat_dec 0 a l0). 411 | gen n; induction H0; simpl; intros; auto; destruct* n. 412 | (* recf *) 413 | inversions* l; clear l. 414 | inversions* H; clear H. 415 | Qed. 416 | 417 | Lemma closed : forall c, sch_fv (Delta.type c) = {}. 418 | Proof. 419 | intros. 420 | induction c; unfold sch_fv; simpl. 421 | (* tag *) 422 | unfold kind_fv; simpl. 423 | repeat rewrite union_empty_l. auto. 424 | (* matches *) 425 | set (x := 1). unfold x at 1. 426 | unfold kind_fv_list, kind_fv; simpl. 427 | generalize x; clear. 428 | induction l; simpl in *; intros; 429 | repeat rewrite union_empty_l. 430 | auto. 431 | use (IHl (S x)); clear IHl. rewrite union_empty_l in H. auto. 432 | (* record *) 433 | set (x := 0). unfold x at 1. 434 | unfold kind_fv_list, kind_fv; simpl. 435 | generalize x; clear. 436 | induction l; simpl in *; intros; 437 | repeat rewrite union_empty_l. 438 | auto. 439 | use (IHl (S x)). 440 | (* sub *) 441 | unfold kind_fv; simpl. 442 | repeat rewrite union_empty_l. auto. 443 | (* recf *) 444 | simpl. 445 | repeat rewrite union_empty_l. auto. 446 | Qed. 447 | 448 | Lemma type_nth_typ_vars : forall n Xs, 449 | n < length Xs -> Defs.type (nth n (typ_fvars Xs) typ_def). 450 | Proof. 451 | induction n; destruct Xs; simpl; intros; try (exfalso; omega). 452 | auto. 453 | apply IHn; omega. 454 | Qed. 455 | 456 | Hint Extern 1 (Defs.type (nth _ (typ_fvars _) typ_def)) => 457 | solve [apply type_nth_typ_vars; omega] : core. 458 | 459 | Lemma scheme : forall c, scheme (Delta.type c). 460 | Proof. 461 | intros. intro; intros. 462 | unfold typ_open_vars. 463 | destruct c; simpl in *. 464 | (* tag *) 465 | do 3 (destruct Xs; try discriminate). 466 | simpl. 467 | split2*. 468 | unfold All_kind_types. 469 | repeat (constructor; simpl*). 470 | (* matches *) 471 | rewrite length_map, length_seq in H. 472 | split. 473 | do 2 (destruct Xs; try discriminate). 474 | inversion H; clear H. 475 | replace Xs with (nil ++ Xs) by simpl*. 476 | set (Xs0 := nil(A:=var)). 477 | replace 2 with (S (S (length Xs0))) by simpl*. 478 | gen Xs; generalize Xs0; induction n; simpl; intros. auto. 479 | destruct Xs; try discriminate. 480 | constructor. 481 | unfold typ_fvars. 482 | rewrite map_app. 483 | rewrite app_nth2 by (rewrite length_map; omega). 484 | rewrite length_map. 485 | replace (length Xs1 - length Xs1) with 0 by omega. 486 | simpl*. 487 | replace (v1 :: Xs) with ((v1 :: nil) ++ Xs) by simpl*. 488 | rewrite app_assoc. 489 | simpl in IHn. 490 | replace (S (length Xs1)) with (length (Xs1 ++ v1 :: nil)) 491 | by (rewrite length_app; simpl; omega). 492 | apply* IHn. 493 | unfold All_kind_types. 494 | repeat (constructor; auto). 495 | induction (seq 1 (length l)); simpl; try constructor; simpl*. 496 | simpl. 497 | assert (length Xs >= 2 + length l) by omega. 498 | clear H; revert H0. 499 | generalize 2; induction n; simpl; intros. auto. 500 | constructor. apply* IHn. 501 | simpl*. 502 | (* record *) 503 | rewrite length_map, length_seq in H. 504 | assert (1 + length l = length Xs) by omega. clear H. 505 | split. 506 | assert (0 < 1) by omega. 507 | revert H0 H. 508 | generalize 1 Xs. clear. 509 | induction (length l); simpl; intros. auto. 510 | constructor. auto. 511 | apply* IHn. 512 | repeat (constructor; auto). 513 | clear; generalize 0; induction (length l); simpl*. 514 | unfold All_kind_types. simpl. 515 | rewrite list_snd_combine by (rewrite length_map, length_seq; auto). 516 | revert H0; generalize 1. 517 | unfold Cstr.attr in *. 518 | clear; induction (length l); simpl; intros. auto. 519 | constructor. apply* IHn. 520 | simpl*. 521 | (* sub *) 522 | split. simpl*. 523 | repeat (constructor; auto). 524 | apply type_nth_typ_vars; omega. 525 | unfold All_kind_types. simpl. 526 | repeat (constructor; simpl*). 527 | Qed. 528 | End Delta. 529 | 530 | Module Sound2 := Sound.Mk2(Delta). 531 | Import Sound2. 532 | Import JudgInfra. 533 | Import Judge. 534 | 535 | Module SndHyp. 536 | Lemma fold_arrow_eq : forall T1 TL1 T2 TL2 Us, 537 | typ_open (fold_right typ_arrow T1 TL1) Us = fold_right typ_arrow T2 TL2 -> 538 | length TL1 = length TL2 -> 539 | typ_open T1 Us = T2 /\ 540 | list_forall2 (fun U1 U2 => typ_open U1 Us = U2) TL1 TL2. 541 | Proof. 542 | induction TL1; intros; destruct* TL2; simpl in *; try discriminate; auto. 543 | inversions H. 544 | destruct* (IHTL1 T2 TL2 Us). 545 | Qed. 546 | 547 | Lemma gc_lower_false : forall gc, gc_lower (false, gc) = (false, gc). 548 | Proof. destruct gc; simpl*. Qed. 549 | 550 | Lemma fold_app_inv : forall K E gc t tl T, 551 | K ; E |(false,gc)|= fold_left trm_app tl t ~: T -> 552 | exists TL, 553 | K ; E |(false,gc)|= t ~: fold_right typ_arrow T TL /\ 554 | list_forall2 (typing (false,gc) K E) tl TL. 555 | Proof. 556 | induction tl using rev_ind; simpl; intros. 557 | exists* (nil(A:=typ)). 558 | rewrite fold_left_app in *; simpl in *. 559 | inversions H; try discriminate. rewrite gc_lower_false in *. 560 | destruct (IHtl (typ_arrow S T) H4). 561 | exists (x0 ++ S :: nil). 562 | rewrite fold_right_app; simpl. 563 | split2*. 564 | Qed. 565 | 566 | Lemma map_nth : forall (A B:Set) d1 d2 (f:A->B) k l, 567 | k < length l -> nth k (map f l) d1 = f (nth k l d2). 568 | Proof. 569 | induction k; intros; destruct l; simpl in *; auto; 570 | try elim (Nat.nlt_0_r _ H). 571 | apply IHk. 572 | apply* PeanoNat.lt_S_n. 573 | Qed. 574 | 575 | Hint Rewrite length_combine combine_nth : list. 576 | 577 | Definition delta_typed_cst c := 578 | forall tl vl K E gc T TL, 579 | K; E | (false, gc) |= trm_cst c ~: fold_right typ_arrow T TL -> 580 | list_forall2 (typing (false, gc) K E) tl TL -> 581 | K ; E |(false,gc)|= @Delta.reduce c tl vl ~: T. 582 | 583 | Lemma delta_typed_tag : forall v, delta_typed_cst (Const.tag v). 584 | Proof. 585 | intros; intro; introv TypC TypA. 586 | exfalso. 587 | destruct vl as [Hv _]. 588 | inversions TypC; try discriminate. clear TypC H4. 589 | destruct* tl; try destruct* tl; try destruct* tl; try discriminate. 590 | inversions TypA; clear TypA. 591 | inversions H6; clear H6. 592 | inversions H8; clear H8. 593 | unfold sch_open in H0. 594 | simpl in H0. 595 | inversions H0. clear H0. 596 | destruct H5 as [_ WK]. simpl in WK. 597 | inversions WK; clear WK. 598 | inversions H7; clear H7. 599 | inversions H5; clear H5. 600 | discriminate. 601 | Qed. 602 | 603 | Lemma typing_value_inv : forall gc K E t T n, 604 | K;E |(false,gc)|= t ~: T -> valu n t -> 605 | (exists t1, t = trm_abs t1) \/ 606 | (exists c, exists vl, t = const_app c vl /\ 607 | n <= Const.arity c /\ list_for_n value (Const.arity c - n) vl). 608 | Proof. 609 | intros. 610 | gen K E T. induction H0; intros. 611 | left*. 612 | right*; exists c; exists (@nil trm). 613 | split2*. 614 | rewrite Nat.sub_diag. split2*. split2*. 615 | clear IHvalu2. 616 | inversions H; clear H; try discriminate. 617 | rewrite gc_lower_false in *. 618 | destruct (IHvalu1 _ _ _ H4); clear IHvalu1. 619 | destruct H. subst. inversion H0_. 620 | destruct H as [c [vl [HE [HC HV]]]]. 621 | subst. 622 | right*; exists c; exists (vl ++ t2 :: nil). 623 | rewrite const_app_app. split2*. 624 | split. omega. 625 | destruct HV. split2*. 626 | assert (value t2) by exists* n2. 627 | apply* list_forall_app. 628 | Qed. 629 | 630 | Lemma value_fvar_is_const : forall K E gc t x k, 631 | K; E |(false,gc)|= t ~: typ_fvar x -> 632 | binds x (Some k) K -> 633 | value t -> 634 | match Cstr.cstr_sort (kind_cstr k) with 635 | | Cstr.Ksum => 636 | exists v, exists t2, t = trm_app (trm_cst (Const.tag v)) t2 637 | | Cstr.Kprod => 638 | exists l, exists nd, exists tl, 639 | t = const_app (@Const.record l nd) tl /\ length tl = length l 640 | | Cstr.Kbot => True 641 | end. 642 | Proof. 643 | introv Typ B HV. 644 | destruct HV as [i vi]. 645 | destruct (typing_value_inv Typ vi). 646 | destruct H. subst. inversions Typ. discriminate. 647 | destruct H as [c [vl [HE [HA Hvl]]]]. 648 | subst. 649 | destruct (fold_app_inv _ _ Typ) as [TL [Typ0 TypA]]; clear Typ. 650 | inversions Typ0; clear Typ0; try discriminate. 651 | unfold sch_open in H0. destruct H5. clear H. 652 | destruct c; simpl in *. 653 | (* tag *) 654 | destruct TL; try discriminate. 655 | inversions H0; clear H0. 656 | inversions H2; clear H2. 657 | inversions H7; clear H7. 658 | simpl in H5. 659 | inversions H2; clear H2. 660 | destruct TL; try discriminate. 661 | inversions H9; clear H9. 662 | puts (binds_func H0 B). inversions H; clear H H0. 663 | puts (proj1 (proj1 H7)). simpl in H. 664 | destruct H; rewrite* H. 665 | (* Ksum *) 666 | inversions TypA; clear TypA; try discriminate. 667 | inversions H9; clear H9; try discriminate. 668 | esplit; esplit; reflexivity. 669 | (* matches *) 670 | case_eq (cut (length TL) (map Delta.matches_arg (seq 2 (length l)))); 671 | introv R1. 672 | assert (length TL <= length (map Delta.matches_arg (seq 2 (length l)))). 673 | rewrite length_map, length_seq. 674 | rewrite <- (list_forall2_length TypA). rewrite <- (proj1 Hvl). 675 | omega. 676 | destruct (cut_ok _ H R1). clear H. rewrite H5 in H0. 677 | rewrite fold_right_app in H0. 678 | destruct* (fold_arrow_eq _ _ _ _ _ H0). 679 | destruct l1; discriminate. 680 | (* record *) 681 | case_eq (cut (length TL) (map typ_bvar (seq 1 (length l)))); introv R1. 682 | assert (length TL <= length (map typ_bvar (seq 1 (length l)))). 683 | rewrite length_map, length_seq. 684 | rewrite <- (list_forall2_length TypA). rewrite <- (proj1 Hvl). 685 | omega. 686 | destruct (cut_ok _ H R1). clear H. rewrite H5 in H0. 687 | rewrite fold_right_app in H0. 688 | destruct* (fold_arrow_eq _ _ _ _ _ H0). 689 | destruct l1; try discriminate. 690 | inversions H2; clear H2. simpl in H; subst. 691 | inversions H9; clear H9. 692 | puts (binds_func H8 B). inversions H; clear H H8. 693 | puts (proj1 (proj1 H10)). simpl in H. 694 | destruct H; rewrite~ H. 695 | (* Kprod *) 696 | rewrite app_nil_r in H5. 697 | exists l. exists n. exists vl. 698 | split2*. subst. length_hyps. rewrite length_seq in *. omega. 699 | (* sub *) 700 | destruct Hvl. 701 | inversions TypA; try discriminate. 702 | (* recf *) 703 | destruct Hvl. 704 | inversions TypA; clear TypA; try discriminate. 705 | inversions H0; clear H0. 706 | inversions H6; clear H6; try discriminate. 707 | destruct i; discriminate. 708 | Qed. 709 | 710 | Lemma fold_right_app_end : forall T1 T2 TL, 711 | fold_right typ_arrow (typ_arrow T1 T2) TL = 712 | fold_right typ_arrow T2 (TL ++ T1 :: nil). 713 | Proof. 714 | intros; rewrite* fold_right_app. 715 | Qed. 716 | 717 | Lemma in_matches_types : forall i l TL Us, 718 | i < length l -> 719 | In (nth i l 0, nth i Us typ_def) 720 | (map_snd (fun T => typ_open T (TL ++ Us)) 721 | (combine l (map typ_bvar (seq (length TL) (length l))))). 722 | Proof. 723 | intros. 724 | remember (0 + i) as m. 725 | pattern i at 2. 726 | replace i with (0+i) by reflexivity. rewrite <- Heqm. 727 | generalize 0 at 1. intro d. 728 | replace (length TL) with (0 + length TL) by auto. 729 | gen i; generalize 0; induction l; intros. 730 | simpl in H; exfalso; omega. 731 | destruct i; simpl. 732 | replace m with n; auto. rewrite app_nth2. 733 | replace (n + length TL - length TL) with n by omega. 734 | auto. omega. omega. 735 | right*. 736 | apply (IHl (S n) i). simpl in H. omega. 737 | omega. 738 | Qed. 739 | 740 | Lemma delta_typed_matches : forall l n, delta_typed_cst (@Const.matches l n). 741 | Proof. 742 | intros; intro; introv Typ0 TypA. 743 | inversions Typ0; try discriminate; clear Typ0. 744 | clear H1 H4. 745 | unfold sch_open in H0; simpl in H0. 746 | rewrite fold_right_app_end in H0. 747 | poses Hlen (proj1 vl). simpl in Hlen. 748 | destruct (fold_arrow_eq _ _ _ _ _ H0) as [HT HA]; clear H0. 749 | autorewrite with list. simpl*. 750 | forward~ (list_forall2_nth trm_def typ_def TypA (n:=length l)) as Typn. 751 | rewrite* <- Hlen. 752 | forward~ (list_forall2_nth typ_def typ_def HA (n:=length l)) as Eqn. 753 | autorewrite with list. simpl. omega. 754 | rewrite app_nth2 in Eqn by (rewrite length_map, length_seq; auto). 755 | autorewrite with list in Eqn. 756 | rewrite Nat.sub_diag in Eqn. 757 | simpl in Eqn. 758 | destruct H5 as [_ WUs]. simpl in WUs. 759 | inversions WUs; clear WUs. 760 | inversions H3; clear H3. 761 | simpl in Eqn. 762 | inversions H1; clear H1 H2. 763 | rewrite <- H6 in *; clear H6. 764 | puts (value_fvar_is_const Typn H0). 765 | assert (Cstr.cstr_sort (kind_cstr k') = Cstr.Ksum). 766 | destruct H4 as [[HE _] _]. simpl in HE. destruct* HE. 767 | clear -H1. destruct k'. 768 | exfalso. simpl in H1. 769 | destruct* kind_valid0. 770 | rewrite H1 in H; clear H1. 771 | destruct H as [v [t2 Hv]]. 772 | apply* (list_forall_out (proj2 vl)). apply* nth_In. 773 | simpl. 774 | rewrite Hv in *; clear Hv. 775 | inversions Typn; clear Typn; try discriminate. 776 | rewrite gc_lower_false in *. 777 | inversions H6; clear H6; try discriminate. 778 | destruct H11. 779 | destruct H as [H _]. simpl in H. 780 | do 3 (destruct Us; try discriminate). clear H. 781 | simpl in *. 782 | inversions H1; clear H1 H9 H10. 783 | inversions H12; clear H12 H7. 784 | inversions H3; clear H3 H9. 785 | puts (binds_func H6 H0). 786 | inversions H; clear H H0 H6 H5. 787 | destruct H4; destruct H7. destruct H; destruct H1. 788 | destruct k' as [[ks kcl kch] kv kr kh]; simpl in *. 789 | destruct* kch. 790 | assert (Hvt: In (v,t) kr) by auto*. clear H2 H1. 791 | unfold Cstr.valid in kv; simpl in kv. 792 | assert (Hvl: In v l) by auto*. 793 | case_eq (index eq_nat_dec 0 v l); introv R2; 794 | try elim (index_none_notin _ _ _ _ R2 Hvl). 795 | destruct (index_ok _ 0 _ _ R2). 796 | unfold Cstr.attr in *. 797 | forward~ (list_forall2_nth trm_def typ_def TypA (n:=n0)) as Typv. 798 | omega. 799 | forward~ (list_forall2_nth typ_def typ_def HA (n:=n0)) as Eqv. 800 | rewrite (list_forall2_length HA). 801 | rewrite <- (list_forall2_length TypA). 802 | omega. 803 | rewrite <- Eqv in Typv; clear HA TypA Eqv. 804 | rewrite app_nth1 in Typv by (rewrite length_map, length_seq; auto). 805 | rewrite (map_nth _ 0) in Typv by rewrite* length_seq. 806 | rewrite seq_nth in Typv by auto. 807 | simpl in Typv. 808 | rewrite (kh v (nth n0 lb0 typ_def) t) in Typv; auto. 809 | eapply typing_app; rewrite* gc_lower_false. 810 | unfold Cstr.unique. simpl. apply* set_mem_correct2. 811 | unfold set_In; auto*. 812 | rewrite <- H2; apply H0. 813 | apply* (@in_matches_types n0 l (typ_fvar x :: b0 :: nil) lb0). 814 | Qed. 815 | 816 | Lemma delta_typed_record : forall l n, delta_typed_cst (@Const.record l n). 817 | Proof. 818 | intros; intro; introv Typ0 TypA. 819 | exfalso. 820 | inversions Typ0; try discriminate. 821 | puts (proj1 vl). 822 | simpl in H. 823 | unfold sch_open in H0. simpl in H0. 824 | destruct TL using rev_ind. inversion TypA. subst; discriminate. 825 | clear IHTL. rewrite fold_right_app in H0. 826 | puts (fold_arrow_eq _ _ _ _ _ H0). 827 | rewrite length_map, length_seq in H2. 828 | puts (list_forall2_length TypA). rewrite length_app in H3; simpl in H3. 829 | destruct H2. omega. 830 | destruct H5. 831 | simpl in H7. inversions H7; clear H7. 832 | inversions H10; clear H10. discriminate. 833 | Qed. 834 | 835 | Lemma delta_typed_sub : forall a, delta_typed_cst (Const.sub a). 836 | Proof. 837 | intros; intro; introv Typ0 TypA. 838 | inversions Typ0; clear Typ0; try discriminate. 839 | puts (proj1 vl). simpl in H. 840 | do 2 (destruct tl; try discriminate). 841 | inversions TypA; clear TypA H H1 H4. 842 | inversions H8; clear H8. 843 | unfold sch_open in H0; simpl in H0. 844 | destruct H5. 845 | inversions H1; clear H1. 846 | inversions H7; clear H7 H4. 847 | inversions H8; clear H8; simpl in *. 848 | inversions H0; clear H0 H. 849 | inversions H3; clear H3. 850 | puts (value_fvar_is_const H6 H0). 851 | assert (Cstr.cstr_sort (kind_cstr k') = Cstr.Kprod). 852 | destruct H2 as [[HE _] _]. simpl in HE. destruct* HE. 853 | clear -H1. destruct k'. 854 | exfalso. simpl in H1. 855 | destruct* kind_valid0. 856 | rewrite H1 in H; clear H1. 857 | destruct* H as [l' [nd [tl [HE HL]]]]. 858 | destruct vl. inversions* H1. 859 | clear vl; subst. 860 | rewrite Delta.record_args_ok. rewrite app_nil_r. 861 | destruct (fold_app_inv _ _ H6) as [TL [Typ0 TypA]]; clear H6. 862 | inversions Typ0; clear Typ0; try discriminate. 863 | unfold sch_open in H1; simpl in H1; clear H3 H6. 864 | destruct (fold_arrow_eq _ _ _ _ _ H1) as [Hx HTL]. 865 | rewrite length_map, length_seq. auto. 866 | clear H1; destruct H7 as [_ WS]. 867 | simpl in WS. 868 | inversions WS; clear WS. 869 | simpl in Hx; subst. 870 | inversions H3; clear H5 H3. 871 | puts (binds_func H6 H0). 872 | inversions H; clear H H0 H6. 873 | puts (proj33 (proj1 H7)). 874 | puts (proj32 (proj1 H2)). 875 | destruct k' as [[ks kl km] kv kr kh]; simpl in *. 876 | destruct* km. 877 | puts (proj2 kv). simpl in H1. 878 | assert (In a l') by simpl*. clear H H1. 879 | case_eq (index eq_nat_dec 0 a l'); introv R1; 880 | [|elim (index_none_notin _ _ _ _ R1 H3)]. 881 | destruct (index_ok _ 0 _ _ R1). 882 | forward~ (list_forall2_nth Delta.trm_default typ_def TypA (n:=n)) as Typ. 883 | rewrite* HL. 884 | rewrite (kh a T (nth n TL typ_def)). auto. 885 | unfold Cstr.unique. simpl. apply* set_mem_correct2. apply* H0. 886 | apply* (proj2 H2). simpl*. 887 | forward~ (list_forall2_nth typ_def typ_def HTL (n:=n)) as Eqv. 888 | rewrite length_map, length_seq; auto. 889 | rewrite <- Eqv; clear Eqv. 890 | rewrite (map_nth _ 0) by rewrite* length_seq. 891 | rewrite* seq_nth. 892 | apply (proj2 H7). 893 | simpl. 894 | rewrite <- H1. 895 | apply* (@in_matches_types n l' (typ_fvar x::nil) lb). 896 | Qed. 897 | 898 | Lemma delta_typed : forall c tl vl K E gc T, 899 | K ; E |(false,gc)|= const_app c tl ~: T -> 900 | K ; E |(false,gc)|= @Delta.reduce c tl vl ~: T. 901 | Proof. 902 | intros. 903 | destruct (fold_app_inv _ _ H) as [TL [Typ0 TypA]]; clear H. 904 | destruct c. 905 | apply* delta_typed_tag. 906 | apply* delta_typed_matches. 907 | apply* delta_typed_record. 908 | apply* delta_typed_sub. 909 | (* recf *) 910 | puts (proj1 vl). 911 | inversions TypA; clear TypA; try discriminate. 912 | inversions H1; clear H1; try discriminate. 913 | inversions H3; clear H3; try discriminate. 914 | simpl in *. 915 | inversions Typ0; try discriminate. 916 | eapply typing_app; rewrite* gc_lower_false. 917 | eapply typing_app; rewrite* gc_lower_false. 918 | eapply typing_app; rewrite* gc_lower_false. 919 | Qed. 920 | End SndHyp. 921 | 922 | Module Sound3 := Sound2.Mk3(SndHyp). 923 | -------------------------------------------------------------------------------- /ML_SP_Soundness.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Preservation and Progress for mini-ML (CBV) - Proofs * 3 | * Arthur Chargueraud, March 2007, Coq v8.1 * 4 | * Extension to structural polymorphism * 5 | * Jacques Garrigue, October 2007 - June 2008 * 6 | ***************************************************************************) 7 | 8 | Set Implicit Arguments. 9 | Require Import Arith List Metatheory 10 | ML_SP_Definitions ML_SP_Infrastructure. 11 | Require Omega. 12 | 13 | Module MkSound(Cstr:CstrIntf)(Const:CstIntf). 14 | 15 | Module Infra := MkInfra(Cstr)(Const). 16 | Import Infra. 17 | Import Defs. 18 | 19 | Module Mk2(Delta:DeltaIntf). 20 | Module JudgInfra := MkJudgInfra(Delta). 21 | Import JudgInfra. 22 | Import Judge. 23 | 24 | Lemma kenv_ok_concat : forall K1 K2, 25 | kenv_ok K1 -> kenv_ok K2 -> disjoint (dom K1) (dom K2) -> kenv_ok (K1 & K2). 26 | Proof. auto. Qed. 27 | 28 | Lemma ok_kinds_open_vars : forall K Ks Xs, 29 | ok K -> fresh (dom K) (length Ks) Xs -> 30 | ok (K & kinds_open_vars Ks Xs). 31 | Proof. 32 | intros. 33 | unfold kinds_open_vars. 34 | apply* disjoint_ok. 35 | apply* ok_combine_fresh. 36 | Qed. 37 | 38 | Hint Resolve ok_kinds_open_vars : core. 39 | 40 | (* ********************************************************************** *) 41 | (** Typing is preserved by weakening *) 42 | 43 | Lemma typing_weaken : forall gc G E F K t T, 44 | K ; (E & G) |gc|= t ~: T -> 45 | env_ok (E & F & G) -> 46 | K ; (E & F & G) |gc|= t ~: T. 47 | Proof. 48 | introv Typ. gen_eq (E & G) as H. gen G. 49 | induction Typ; introv EQ Ok; subst. 50 | apply* typing_var. apply* binds_weaken. 51 | apply_fresh* (@typing_abs gc) as y. apply_ih_bind* H1. 52 | forward~ (H0 y) as Q. 53 | apply_fresh* (@typing_let gc M L1) as y. apply_ih_bind* H2. 54 | forward~ (H1 y) as Q. 55 | auto*. 56 | auto. 57 | apply_fresh* (@typing_gc gc Ks) as y. 58 | Qed. 59 | 60 | Lemma proper_instance_weaken : forall K K' K'' Ks Us, 61 | ok (K & K' & K'') -> 62 | proper_instance (K & K'') Ks Us -> 63 | proper_instance (K & K' & K'') Ks Us. 64 | Proof. 65 | intros. 66 | destruct* H0 as [TM FM]; split2*. 67 | Qed. 68 | 69 | Lemma typing_weaken_kinds : forall gc K K' K'' E t T, 70 | K & K''; E |gc|= t ~: T -> 71 | kenv_ok (K & K' & K'') -> 72 | K & K' & K''; E |gc|= t ~: T. 73 | Proof. 74 | introv Typ. gen_eq (K & K'') as H. gen K''. 75 | induction Typ; introv EQ Ok; subst. 76 | apply* typing_var. apply* proper_instance_weaken. 77 | apply_fresh* (@typing_abs gc) as y. 78 | apply_fresh* (@typing_let gc M (L1 \u dom(K&K'&K''))) as y. 79 | intros. clear H1 H2. 80 | rewrite concat_assoc. 81 | apply* H0; clear H0. rewrite* concat_assoc. 82 | forward~ (H Xs) as Typ. 83 | apply* typing_app. 84 | apply* typing_cst. apply* proper_instance_weaken. 85 | apply_fresh* (@typing_gc gc Ks) as y. 86 | intros. 87 | rewrite concat_assoc. 88 | apply* (H1 Xs); clear H1. 89 | rewrite* concat_assoc. 90 | forward~ (H0 Xs) as Typ; clear H0. 91 | Qed. 92 | 93 | Lemma typing_weaken_kinds' : forall gc K K' E t T, 94 | kenv_ok (K & K') -> 95 | K ; E |gc|= t ~: T -> K & K' ; E |gc|= t ~: T. 96 | Proof. 97 | intros. 98 | replace (K & K') with (K & K' & empty) by simpl*. 99 | apply* typing_weaken_kinds. 100 | Qed. 101 | 102 | Lemma proper_instance_subst : forall K K' K'' Ks Us S, 103 | env_prop type S -> 104 | proper_instance (K & K' & K'') Ks Us -> 105 | well_subst (K & K' & K'') (K & map (kind_subst S) K'') S -> 106 | proper_instance (K & map (kind_subst S) K'') (List.map (kind_subst S) Ks) 107 | (List.map (typ_subst S) Us). 108 | Proof. 109 | introv TS PI WS. 110 | destruct* PI. 111 | split. rewrite length_map. apply* typ_subst_type_list. 112 | rewrite* <- kinds_subst_open. 113 | Qed. 114 | 115 | Lemma well_subst_fresh : forall K K' K'' S Ys Ks, 116 | well_subst (K & K' & K'') (K & map (kind_subst S) K'') S -> 117 | fresh (dom S \u dom K \u dom K'') (length Ks) Ys -> 118 | well_subst (K & K' & K'' & kinds_open_vars Ks Ys) 119 | (K & map (kind_subst S) (K'' & kinds_open_vars Ks Ys)) S. 120 | Proof. 121 | introv WS Fr. 122 | assert (KxYs: disjoint (dom K \u dom K'') 123 | (dom (kinds_open_vars Ks Ys))) by auto. 124 | intro x; intros. 125 | rewrite map_concat. rewrite <- concat_assoc. 126 | destruct* (binds_concat_inv H) as [[N B]|B]; clear H. 127 | destruct k; try constructor. 128 | simpl. rewrite get_notin_dom by auto. 129 | puts (binds_map (kind_subst S) B). 130 | apply* wk_kind. 131 | Qed. 132 | 133 | Lemma All_kind_types_subst : forall k S, 134 | All_kind_types type k -> 135 | env_prop type S -> All_kind_types type (kind_subst S k). 136 | Proof. 137 | intros; unfold kind_subst; apply All_kind_types_map. 138 | apply* All_kind_types_imp. 139 | Qed. 140 | 141 | Hint Resolve All_kind_types_subst : core. 142 | 143 | Lemma kenv_ok_subst : forall K K' K'' S, 144 | env_prop type S -> 145 | kenv_ok (K & K' & K'') -> kenv_ok (K & map (kind_subst S) K''). 146 | Proof. 147 | introv HS H. 148 | kenv_ok_solve. auto. 149 | intro; intros. 150 | destruct (in_map_inv _ _ _ _ H1) as [b [Hb B]]. 151 | subst*. 152 | Qed. 153 | 154 | Lemma env_ok_subst : forall E E' S, 155 | env_prop type S -> 156 | env_ok (E & E') -> env_ok (E & map (sch_subst S) E'). 157 | Proof. 158 | introv HS H. 159 | env_ok_solve. auto. 160 | intro; intros. 161 | destruct (in_map_inv _ _ _ _ H0) as [b [Hb B]]. 162 | subst*. 163 | Qed. 164 | 165 | Hint Resolve kenv_ok_subst env_ok_subst : core. 166 | 167 | (* ********************************************************************** *) 168 | (** Type substitution preserves typing *) 169 | 170 | Lemma typing_typ_subst : forall gc F K'' S K K' E t T, 171 | disjoint (dom S) (env_fv E \u fv_in kind_fv K) -> 172 | env_prop type S -> 173 | well_subst (K & K' & K'') (K & map (kind_subst S) K'') S -> 174 | K & K' & K''; E & F |gc|= t ~: T -> 175 | K & map (kind_subst S) K''; E & (map (sch_subst S) F) |gc|= 176 | t ~: (typ_subst S T). 177 | Proof. 178 | introv. intros Dis TS WS Typ. 179 | gen_eq (K & K' & K'') as GK; gen_eq (E & F) as G; gen K''; gen F. 180 | induction Typ; introv WS EQ EQ'; subst; simpls typ_subst. 181 | (* Var *) 182 | rewrite~ sch_subst_open. apply* typing_var. 183 | binds_cases H1. 184 | apply* binds_concat_fresh. 185 | rewrite* sch_subst_fresh. 186 | use (fv_in_spec sch_fv _ _ _ (binds_in B)). 187 | auto*. 188 | destruct M as [T Ks]. simpl. 189 | apply* proper_instance_subst. 190 | (* Abs *) 191 | apply_fresh* (@typing_abs gc) as y. 192 | replace (Sch (typ_subst S U) nil) with (sch_subst S (Sch U nil)) by auto. 193 | apply_ih_map_bind* H1. 194 | (* Let *) 195 | apply_fresh* (@typing_let gc (sch_subst S M) 196 | (L1 \u dom S \u dom K \u dom K'')) as y. 197 | clear H H1 H2. clear L2 T2 t2 Dis. 198 | simpl. intros Ys Fr. 199 | rewrite* <- sch_subst_open_vars. 200 | rewrite* <- kinds_subst_open_vars. 201 | rewrite concat_assoc. rewrite <- map_concat. 202 | rewrite length_map in Fr. 203 | apply* H0; clear H0. 204 | apply* well_subst_fresh. 205 | rewrite* concat_assoc. 206 | apply_ih_map_bind* H2. 207 | (* App *) 208 | auto*. 209 | (* Cst *) 210 | rewrite* sch_subst_open. 211 | assert (disjoint (dom S) (sch_fv (Delta.type c))). 212 | intro x. rewrite* Delta.closed. 213 | rewrite sch_subst_fresh; try disjoint_solve. 214 | apply* typing_cst. 215 | rewrite* <- (sch_subst_fresh S H2). 216 | destruct (Delta.type c) as [T Ks]; simpl. 217 | apply* proper_instance_subst. 218 | (* GC *) 219 | apply* (@typing_gc gc (List.map (kind_subst S) Ks) 220 | (L \u dom S \u dom K \u dom K'')). 221 | rewrite length_map; intros. 222 | rewrite* <- kinds_subst_open_vars. 223 | rewrite concat_assoc. rewrite <- map_concat. 224 | apply* (H1 Xs); clear H1. 225 | apply* well_subst_fresh. 226 | rewrite* concat_assoc. 227 | Qed. 228 | 229 | Lemma typing_typ_substs : forall gc K' S K E t T, 230 | disjoint (dom S) (env_fv E \u fv_in kind_fv K \u dom K) -> 231 | env_prop type S -> 232 | well_subst (K & K') K S -> 233 | K & K'; E |gc|= t ~: T -> 234 | K ; E |gc|= t ~: (typ_subst S T). 235 | Proof. 236 | intros. 237 | generalize (@typing_typ_subst gc empty empty); intro TTS. 238 | simpl in TTS. 239 | apply* TTS. 240 | Qed. 241 | 242 | (* ********************************************************************** *) 243 | (** Typing schemes for expressions *) 244 | 245 | Definition has_scheme_vars gc L (K:kenv) E t M := forall Xs, 246 | fresh L (sch_arity M) Xs -> 247 | K & kinds_open_vars (sch_kinds M) Xs; E |gc|= t ~: (M ^ Xs). 248 | 249 | Definition has_scheme gc K E t M := forall Vs, 250 | types (sch_arity M) Vs -> 251 | list_forall2 (well_kinded K) (kinds_open (sch_kinds M) Vs) Vs -> 252 | K ; E |gc|= t ~: (M ^^ Vs). 253 | 254 | (* ********************************************************************** *) 255 | (** Type schemes of terms can be instanciated *) 256 | 257 | Lemma kind_subst_open_combine : forall Xs Vs Ks, 258 | fresh (kind_fv_list Ks) (length Xs) Xs -> 259 | types (length Xs) Vs -> 260 | forall k : kind, 261 | In k Ks -> 262 | kind_open k Vs = kind_subst (combine Xs Vs) (kind_open k (typ_fvars Xs)). 263 | Proof. 264 | introv Fr. intros. 265 | destruct H. 266 | rewrite* kind_subst_open. 267 | rewrite* kind_subst_fresh. 268 | rewrite* (fresh_subst {}). 269 | rewrite* <- H. 270 | rewrite* dom_combine. 271 | use (kind_fv_fresh _ _ _ _ H0 Fr). 272 | Qed. 273 | 274 | Lemma well_subst_open_vars : forall (K:kenv) Vs (Ks:list kind) Xs, 275 | fresh (fv_in kind_fv K) (length Ks) Xs -> 276 | fresh (kind_fv_list Ks) (length Xs) Xs -> 277 | types (length Xs) Vs -> 278 | list_forall2 (well_kinded K) (kinds_open Ks Vs) Vs -> 279 | well_subst (K & kinds_open_vars Ks Xs) K (combine Xs Vs). 280 | Proof. 281 | introv Fr Fr' TV WK. 282 | intro x; intros. 283 | destruct* (binds_concat_inv H) as [[N B]|B]; clear H. 284 | unfold kinds_open_vars in N. 285 | rewrite* kind_subst_fresh. 286 | simpl. 287 | rewrite* get_notin_dom. 288 | destruct* k. 289 | use (fv_in_spec kind_fv _ _ _ (binds_in B)). 290 | unfold kinds_open_vars, kinds_open in *. 291 | rewrite <- map_combine in B. 292 | destruct (binds_map_inv _ _ B) as [k0 [Hk0 Bk0]]. subst. 293 | puts (binds_map (kind_subst (combine Xs Vs)) B). 294 | simpl in H; do 2 rewrite map_combine in H. 295 | rewrite list_map_comp in H. 296 | refine (list_forall2_get (P:=well_kinded K) Xs _ H _). 297 | instantiate (1:=Vs). 298 | rewrite* <- (list_map_ext Ks _ _ (kind_subst_open_combine _ _ Fr' TV)). 299 | simpl; case_eq (get x (combine Xs Vs)); intros. auto. 300 | elim (get_contradicts _ _ _ _ Bk0 H0); auto. 301 | Qed. 302 | 303 | Lemma has_scheme_from_vars : forall gc L K E t M, 304 | has_scheme_vars gc L K E t M -> 305 | has_scheme gc K E t M. 306 | Proof. 307 | intros gc L K E t [T Ks] H Vs TV. unfold sch_open. simpls. 308 | fold kind in K. fold kenv in K. 309 | pick_freshes (length Ks) Xs. 310 | rewrite (fresh_length _ _ _ Fr) in TV. 311 | rewrite~ (@typ_subst_intro Xs Vs T). 312 | unfolds has_scheme_vars sch_open_vars. simpls. 313 | intro WK. 314 | apply* (@typing_typ_substs gc (kinds_open_vars Ks Xs)). 315 | apply list_forall_env_prop. destruct* TV. 316 | apply* well_subst_open_vars. 317 | Qed. 318 | 319 | (* ********************************************************************** *) 320 | (** Typing is preserved by term substitution *) 321 | 322 | Lemma typing_trm_subst : forall gc F M K E t T z u, 323 | K ; E & z ~ M & F |(gc,GcAny)|= t ~: T -> 324 | (exists L:vars, has_scheme_vars (gc,GcAny) L K E u M) -> 325 | term u -> 326 | K ; E & F |(gc,GcAny)|= (trm_subst z u t) ~: T. 327 | Proof. 328 | introv Typt. intros Typu Wu. 329 | gen_eq (E & z ~ M & F) as G. gen_eq (gc, GcAny) as gc0. gen F. 330 | induction Typt; introv EQ1 EQ2; subst; simpl trm_subst; 331 | destruct Typu as [Lu Typu]. 332 | case_var. 333 | binds_get H1. apply_empty* (@typing_weaken (gc,GcAny)). 334 | destruct H2; apply* (has_scheme_from_vars Typu). 335 | binds_cases H1; apply* typing_var. 336 | apply_fresh* (@typing_abs (gc,GcAny)) as y. 337 | rewrite* trm_subst_open_var. 338 | apply_ih_bind* H1. 339 | apply_fresh* (@typing_let (gc,GcAny) M0 L1) as y. 340 | intros; apply* H0. 341 | exists (Lu \u mkset Xs); intros Ys TypM. 342 | forward~ (Typu Ys) as Typu'; clear Typu. 343 | apply* typing_weaken_kinds. 344 | forward~ (H Xs). 345 | rewrite* trm_subst_open_var. 346 | apply_ih_bind* H2. 347 | assert (exists L : vars, has_scheme_vars (gc,GcAny) L K E u M). exists* Lu. 348 | auto*. 349 | auto*. 350 | apply_fresh* (@typing_gc (gc,GcAny) Ks) as y. 351 | intros Xs Fr. 352 | apply* H1; clear H1. 353 | exists (Lu \u dom K \u mkset Xs); intros Ys Fr'. 354 | forward~ (Typu Ys) as Typu'; clear Typu. 355 | apply* typing_weaken_kinds. 356 | forward~ (H0 Xs). 357 | Qed. 358 | 359 | (* ********************************************************************** *) 360 | (** Canonical derivations *) 361 | 362 | (* less than 100 lines! *) 363 | 364 | Lemma typing_gc_any : forall gc K E t T, 365 | K ; E |gc|= t ~: T -> K ; E |(true,GcAny)|= t ~: T. 366 | Proof. 367 | induction 1; auto*. 368 | apply* typing_gc. simpl; auto. 369 | Qed. 370 | 371 | Lemma typing_gc_raise : forall gc K E t T, 372 | K ; E |gc|= t ~: T -> K ; E |gc_raise gc|= t ~: T. 373 | Proof. 374 | induction 1; destruct gc; destruct g; simpl; auto*. 375 | apply* typing_gc. simpl; auto. 376 | Qed. 377 | 378 | Definition typing_gc_let K E t T := K; E |(true,GcLet)|= t ~: T. 379 | 380 | Lemma typing_gc_ind : forall (P: kenv -> env -> trm -> typ -> Prop), 381 | (forall K E t T, K; E |(false,GcLet)|= t ~: T -> P K E t T) -> 382 | (forall Ks L K E t T, 383 | (forall Xs : list var, 384 | fresh L (length Ks) Xs -> P (K & kinds_open_vars Ks Xs) E t T) -> 385 | P K E t T) -> 386 | forall K E t T, typing_gc_let K E t T -> P K E t T. 387 | Proof. 388 | intros. 389 | unfold typing_gc_let in H1. 390 | gen_eq (true,GcLet) as gc. 391 | induction H1; intros; subst; try solve [apply* H]. 392 | apply* H0. 393 | Qed. 394 | 395 | Lemma typing_canonize : forall gc K E t T, 396 | K ; E |gc|= t ~: T -> K ; E |(true,GcLet)|= t ~: T. 397 | Proof. 398 | induction 1; auto*. 399 | (* App *) 400 | clear H H0. 401 | gen IHtyping1. 402 | fold (typing_gc_let K E t2 S) in IHtyping2. 403 | apply (proj2 (A:=kenv_ok K)). 404 | induction IHtyping2 using typing_gc_ind. 405 | split2*; intros; subst. 406 | gen H. gen_eq (typ_arrow T0 T) as S. 407 | fold (typing_gc_let K E t1 S) in IHtyping1. 408 | apply (proj2 (A:=kenv_ok K)). 409 | induction IHtyping1 using typing_gc_ind. 410 | split2*; intros; subst. 411 | apply* typing_app. 412 | split. 413 | destruct (var_freshes L (length Ks)) as [Xs HXs]. 414 | destruct* (H Xs HXs). 415 | intros; subst. 416 | apply* (@typing_gc (true,GcLet) Ks L). 417 | simpl; auto. 418 | intros. 419 | destruct (H Xs H0); clear H. 420 | apply* H3; clear H3. 421 | apply* typing_weaken_kinds'. 422 | split. 423 | destruct (var_freshes L (length Ks)) as [Xs HXs]. 424 | destruct* (H Xs HXs). 425 | 426 | intros. 427 | apply* (@typing_gc (true,GcLet) Ks L). 428 | simpl; auto. 429 | intros. 430 | destruct (H Xs H0); clear H. 431 | apply* H2; clear H2. 432 | apply* typing_weaken_kinds'. 433 | (* GC *) 434 | apply* typing_gc. 435 | simpl; auto. 436 | Qed. 437 | 438 | (* End of canonical derivations *) 439 | 440 | (* ********************************************************************** *) 441 | (** Extra hypotheses for main results *) 442 | 443 | Module Type SndHypIntf. 444 | Parameter delta_typed : forall c tl vl K E gc T, 445 | K ; E |(false,gc)|= const_app c tl ~: T -> 446 | K ; E |(false,gc)|= @Delta.reduce c tl vl ~: T. 447 | End SndHypIntf. 448 | 449 | Module Mk3(SH:SndHypIntf). 450 | Import SH. 451 | 452 | (* ********************************************************************** *) 453 | (** Preservation: typing is preserved by reduction *) 454 | 455 | Lemma typ_open_vars_nil : forall T, 456 | type T -> typ_open_vars T nil = T. 457 | Proof. 458 | induction T; unfold typ_open_vars; simpl; intros; auto*. 459 | inversion H. 460 | unfold typ_open_vars in *; simpls. 461 | rewrite IHT1. rewrite* IHT2. inversion* H. inversion* H. 462 | Qed. 463 | 464 | Lemma typing_abs_inv : forall gc K E t1 t2 T1 T2, 465 | K ; E |(gc,GcAny)|= trm_abs t1 ~: typ_arrow T1 T2 -> 466 | K ; E |(gc,GcAny)|= t2 ~: T1 -> 467 | K ; E |(gc,GcAny)|= t1 ^^ t2 ~: T2. 468 | Proof. 469 | introv Typ1 Typ2. 470 | gen_eq (gc,GcAny) as gcs. 471 | gen_eq (trm_abs t1) as t. 472 | gen_eq (typ_arrow T1 T2) as T. 473 | induction Typ1; intros; subst; try discriminate. 474 | inversions H2; inversions H3; clear H2 H3. 475 | pick_fresh x. 476 | rewrite* (@trm_subst_intro x). 477 | apply_empty* (@typing_trm_subst gc). 478 | exists {}. intro. unfold kinds_open_vars, sch_open_vars; simpl. 479 | destruct Xs; simpl*. rewrite* typ_open_vars_nil. 480 | apply* (@typing_gc (gc,GcAny) Ks L). 481 | intros. 482 | puts (H0 Xs H2); clear H0. 483 | apply* H1. 484 | apply* typing_weaken_kinds'. 485 | Qed. 486 | 487 | Lemma preservation_result : preservation. 488 | Proof. 489 | introv Typ. gen_eq (true, GcAny) as gc. gen t'. 490 | induction Typ; introv EQ Red; subst; inversions Red; 491 | try solve [apply* typing_gc]; 492 | try (destruct (const_app_inv c tl) as [eq | [T1' [T2' eq]]]; 493 | rewrite eq in *; discriminate). 494 | (* Let *) 495 | pick_fresh x. rewrite* (@trm_subst_intro x). 496 | simpl in H1. 497 | apply_empty* (@typing_trm_subst true). 498 | apply* H1. 499 | (* Let *) 500 | apply* (@typing_let (true,GcAny) M L1). 501 | (* Beta *) 502 | apply* typing_abs_inv. 503 | (* Delta *) 504 | assert (K;E |(true,GcAny)|= trm_app t1 t2 ~: T) by auto*. 505 | use (typing_canonize H). 506 | fold (typing_gc_let K E (trm_app t1 t2) T) in H1. 507 | rewrite <- H0 in *. 508 | clear -H1. 509 | gen_eq (const_app c tl) as t1. 510 | induction H1 using typing_gc_ind; intros; subst. 511 | apply* typing_gc_any. 512 | apply* delta_typed. 513 | apply* typing_gc. simpl*. 514 | (* App1 *) 515 | auto*. 516 | (* App2 *) 517 | auto*. 518 | (* Delta/cst *) 519 | apply* (@typing_gc_any (false,GcAny)). 520 | apply* delta_typed. 521 | rewrite* H3. 522 | Qed. 523 | 524 | (* ********************************************************************** *) 525 | (** Progress: typed terms are values or can reduce *) 526 | 527 | Lemma value_app_const : forall t1 t2 n, 528 | valu n (trm_app t1 t2) -> 529 | exists c:Const.const, exists vl:list trm, 530 | length vl + n = Const.arity c /\ trm_app t1 t2 = const_app c vl /\ 531 | list_forall value vl. 532 | Proof. 533 | induction t1; intros; inversions H; try (inversion H3; fail). 534 | clear IHt1_2. 535 | destruct (IHt1_1 _ _ H3) as [c [vl [Hlen [Heq Hv]]]]. 536 | exists c. exists (vl ++ t2 :: nil). 537 | split. rewrite length_app. rewrite <- Hlen. simpl. ring. 538 | split. rewrite Heq. unfold const_app. 539 | rewrite fold_left_app. simpl. auto. 540 | apply* list_forall_concat. 541 | constructor; auto. exists* n2. 542 | exists c. exists (t2 :: nil). 543 | inversions H3. rewrite H1. 544 | unfold const_app. simpl; auto. 545 | split3*. constructor; auto. exists* n2. 546 | Qed. 547 | 548 | Lemma progress_delta : forall K t0 t3 t2 T, 549 | K; empty |(false,GcLet)|= trm_app (trm_app t0 t3) t2 ~: T -> 550 | valu 0 (trm_app t0 t3) -> 551 | value t2 -> 552 | exists t' : trm, trm_app (trm_app t0 t3) t2 --> t'. 553 | Proof. 554 | intros. 555 | destruct (value_app_const H0) as [c [vl [Hlen [Heq Hv]]]]. 556 | unfold const_app in *. 557 | rewrite Heq in *. 558 | change (exists t', fold_left trm_app (t2::nil) (const_app c vl) --> t'). 559 | unfold const_app; rewrite <- fold_left_app. 560 | assert (list_for_n value (S(Const.arity c)) (vl ++ t2 :: nil)). 561 | split2*. apply* list_forall_app. 562 | exists (Delta.reduce H2). 563 | apply red_delta. 564 | Qed. 565 | 566 | Lemma progress_result : progress. 567 | Proof. 568 | introv Typ. gen_eq (empty:env) as E. gen_eq (true,GcAny) as gc. 569 | poses Typ' Typ. 570 | induction Typ; intros; subst; 571 | try (pick_freshes (length Ks) Xs; apply* (H0 Xs)). 572 | inversions H1. 573 | left*. exists* 0. 574 | right*. pick_freshes (sch_arity M) Ys. 575 | destructi~ (@H0 Ys) as [[n Val1] | [t1' Red1]]. 576 | assert (value t1). exists* n. 577 | exists* (t2 ^^ t1). 578 | exists* (trm_let t1' t2). 579 | destruct~ IHTyp2 as [Val2 | [t2' Red2]]. 580 | destruct~ IHTyp1 as [Val1 | [t1' Red1]]. 581 | use (typing_canonize Typ'). 582 | remember (empty(A:=sch)) as E. 583 | remember (trm_app t1 t2) as t. 584 | clear Typ1 Typ2 Typ'. 585 | fold (typing_gc_let K E t T) in H. 586 | apply (proj2 (A:=kenv_ok K)). 587 | induction H using typing_gc_ind. 588 | split2*; intros; subst. 589 | destruct Val1 as [n Val1]; inversions Val1. 590 | right*; exists* (t0 ^^ t2). 591 | case_eq (Const.arity c); intros. 592 | right*. rewrite H0 in Val1. 593 | assert (list_for_n value 1 (t2 :: nil)) by split2*. 594 | rewrite <- H0 in H1. 595 | exists (Delta.reduce H1). 596 | apply (red_delta H1). 597 | left*. exists n. rewrite H0 in Val1. destruct* Val2. 598 | destruct n. 599 | right*; apply* progress_delta. 600 | left*. destruct Val2. exists* n. 601 | destruct (var_freshes L (length Ks)) as [Xs HXs]. 602 | destruct* (H Xs); clear H. 603 | right*; exists* (trm_app t1' t2). 604 | right*; exists* (trm_app t1 t2'). 605 | left*; exists* (Const.arity c). 606 | destruct (var_freshes L (length Ks)) as [Xs HXs]. 607 | apply* (H1 Xs). 608 | Qed. 609 | 610 | Lemma value_irreducible : forall t t', 611 | value t -> ~(t --> t'). 612 | Proof. 613 | induction t; introv HV; destruct HV as [k HV']; inversions HV'; 614 | intro R; inversions R. 615 | destruct (const_app_inv c tl) as [eq | [t1' [t2' eq]]]; 616 | rewrite eq in *; discriminate. 617 | inversions H2. 618 | destruct (value_app_const HV'). 619 | destruct H as [vl' [Hl [He Hv]]]. 620 | rewrite He in H0; clear He. 621 | destruct (const_app_eq _ _ _ _ H0). subst. 622 | clear -vl Hl; destruct vl. 623 | Omega.omega. 624 | elim (IHt1 t1'). exists* (S k). auto. 625 | elim (IHt2 t2'). exists* n2. auto. 626 | clear -vl H0. 627 | destruct vl. 628 | destruct (const_app_inv c0 tl) as [eq | [t1' [t2' eq]]]; 629 | rewrite eq in *; discriminate. 630 | Qed. 631 | 632 | End Mk3. 633 | 634 | End Mk2. 635 | 636 | End MkSound. 637 | -------------------------------------------------------------------------------- /Metatheory.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Bundles of Modules for Developments in Programming Language Metatheory * 3 | * Brian Aydemir & Arthur Charguéraud, July 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | Require Export Lib_Tactic Lib_FinSet. 7 | Require Export Metatheory_Var. 8 | Require Export Metatheory_Fresh. 9 | Require Export Metatheory_Env. 10 | Require Export Metatheory_SP. 11 | Export Env. 12 | Open Scope set_scope. 13 | Open Scope env_scope. 14 | -------------------------------------------------------------------------------- /Metatheory_Env.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Generic Environments for Programming Language Metatheory * 3 | * Brian Aydemir & Arthur Charguéraud, July 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import Lib_Tactic List Decidable Metatheory_Var Metatheory_Fresh. 8 | 9 | (* ********************************************************************** *) 10 | (** * Definitions of Environments *) 11 | 12 | Declare Scope env_scope. 13 | 14 | Module Env. 15 | 16 | (* ********************************************************************** *) 17 | (** ** Structure and Basic Operations on Environments *) 18 | 19 | Section Definitions. 20 | 21 | Variable A : Set. 22 | Implicit Types x : var. 23 | Implicit Types a : A. 24 | 25 | (** Environments is represented as an associative list. *) 26 | 27 | Definition env := list (var * A). 28 | 29 | (** Empty environment. *) 30 | 31 | Definition empty : env := nil. 32 | 33 | (** An environment containing only one binding mapping x to a. *) 34 | 35 | Definition single x a := (x,a) :: nil. 36 | 37 | (** Concatenation of environments E and F. *) 38 | 39 | Definition concat (E F : env) := F ++ E. 40 | 41 | (** Comutes the domain of the environment, i.e. the set of vars mapped. *) 42 | 43 | Fixpoint dom (E : env) : vars := 44 | match E with 45 | | nil => {} 46 | | (x, _) :: E' => {{x}} \u (dom E') 47 | end. 48 | 49 | (** [map] applies a function to all the values mapped by the environment. *) 50 | 51 | Fixpoint map (f : A -> A) (E : env) {struct E} : env := 52 | match E with 53 | | nil => nil 54 | | (x, V) :: E' => (x, f V) :: map f E' 55 | end. 56 | 57 | (** [get] locates a binding in an environment. *) 58 | 59 | Fixpoint get (x : var) (E : env) {struct E} : option A := 60 | match E with 61 | | nil => None 62 | | (y,a) :: E' => if eq_var_dec x y then Some a else get x E' 63 | end. 64 | 65 | End Definitions. 66 | 67 | 68 | (* ********************************************************************** *) 69 | (** ** Notations *) 70 | 71 | Arguments empty {A}. 72 | 73 | (** [x ~ a] is the notation for a singleton environment mapping x to a. *) 74 | 75 | Notation "x ~ a" := (single x a) 76 | (at level 27, left associativity) : env_scope. 77 | 78 | (** [E & F] is the notation for concatenation of E and F. *) 79 | 80 | Notation "E & F" := (concat E F) 81 | (at level 28, left associativity) : env_scope. 82 | 83 | (** [x # E] to be read x fresh from E captures the fact that 84 | x is unbound in E . *) 85 | 86 | Notation "x '#' E" := (x \notin (dom E)) (at level 67) : env_scope. 87 | 88 | Bind Scope env_scope with env. 89 | Local Open Scope env_scope. 90 | 91 | 92 | (* ********************************************************************** *) 93 | (** ** Relations on environemnts *) 94 | 95 | Section Relations. 96 | 97 | Variable A : Set. 98 | 99 | (** An environment is well-formed iff it contains no duplicate 100 | bindings. *) 101 | 102 | Inductive ok : env A -> Prop := 103 | | ok_empty : 104 | ok empty 105 | | ok_push : forall E x a, 106 | ok E -> x # E -> ok (E & x ~ a). 107 | 108 | (** An environment contains a binding from x to a iff the most recent 109 | binding on x is mapped to a *) 110 | 111 | Definition binds x a (E : env A) := 112 | get x E = Some a. 113 | 114 | (** Proposition capturing the idea that a proposition P holds on all 115 | values containted in the environment. 116 | Use In rather than binds for compatibility with fv_in. *) 117 | 118 | Definition env_prop (P : A -> Prop) (E : env A) := 119 | forall x a, In (x,a) E -> P a. 120 | 121 | (** Inclusion of an environment in another one. *) 122 | 123 | Definition extends (E E' : env A) := 124 | forall x a, binds x a E -> binds x a E'. 125 | 126 | End Relations. 127 | 128 | 129 | (* ********************************************************************** *) 130 | (** * Properties of Environemnts *) 131 | 132 | Hint Constructors ok : core. 133 | 134 | Local Open Scope env_scope. 135 | 136 | 137 | (* ********************************************************************** *) 138 | (** ** Properties of Operations *) 139 | 140 | Section OpProperties. 141 | Variable A : Set. 142 | Implicit Types E F : env A. 143 | Implicit Types a b : A. 144 | 145 | (** Concatenation with an empty environment is identity. *) 146 | 147 | Lemma concat_empty : forall E, 148 | E & empty = E. 149 | Proof. 150 | auto. 151 | Qed. 152 | 153 | (** Concatenation is associative. *) 154 | 155 | Lemma concat_assoc : forall E F G, 156 | (E & F) & G = E & (F & G). 157 | Proof. 158 | induction G; simpl; congruence. 159 | Qed. 160 | 161 | (** Map commutes with substitution. *) 162 | 163 | Lemma map_concat : forall f E F, 164 | map f (E & F) = (map f E) & (map f F). 165 | Proof. 166 | induction F as [|(x,a)]; simpl; congruence. 167 | Qed. 168 | 169 | (** Interaction of associativity and map. *) 170 | 171 | Lemma concat_assoc_map_push : forall f E F y V, 172 | (E & map f F) & y ~ (f V) = E & (map f (F & y ~ V)). 173 | Proof. 174 | auto. 175 | Qed. 176 | 177 | (** Inversion results. *) 178 | 179 | Lemma eq_empty_inv : forall x a E F, 180 | empty = E & x ~ a & F -> False. 181 | Proof. 182 | induction F as [|(y,b)]; simpl; discriminate. 183 | Qed. 184 | 185 | Lemma eq_push_inv : forall E x a E' x' a', 186 | E & x ~ a = E' & x' ~ a' -> E = E' /\ x = x' /\ a = a'. 187 | Proof. 188 | simpl. intros. inversions* H. 189 | Qed. 190 | 191 | (** Domain of an empty environment. *) 192 | 193 | Lemma dom_empty : 194 | dom (A:=A) empty = {}. 195 | Proof. 196 | auto. 197 | Qed. 198 | 199 | (** Domain of a singleton environment. *) 200 | 201 | Lemma dom_single : forall x a, 202 | dom (x ~ a) = {{x}}. 203 | Proof. 204 | simpl. intros. apply union_empty_r. 205 | Qed. 206 | 207 | (** Domain of a push is the head variable union the domain of the tail. *) 208 | 209 | Lemma dom_push : forall x a E, 210 | dom (E & x ~ a) = {{x}} \u dom E. 211 | Proof. 212 | auto. 213 | Qed. 214 | 215 | (** Domain of a cons is the head variable union the domain of the tail. *) 216 | 217 | Lemma dom_cons : forall x a E, 218 | dom ((x,a) :: E) = {{x}} \u dom E. 219 | Proof. 220 | auto. 221 | Qed. 222 | 223 | (** Domain of a concatenation is the union of the domains. *) 224 | 225 | Lemma dom_concat : forall E F, 226 | dom (E & F) = dom E \u dom F. 227 | Proof. 228 | induction F as [|(x,a)]; simpl. 229 | symmetry. apply union_empty_r. 230 | rewrite IHF. apply union_comm_assoc. 231 | Qed. 232 | 233 | (** Map preserves the domain. *) 234 | 235 | Lemma dom_map : forall f E, 236 | dom (map f E) = dom E. 237 | Proof. 238 | induction E as [|(x,a)]; simpl; congruence. 239 | Qed. 240 | 241 | End OpProperties. 242 | 243 | Arguments eq_empty_inv [A x a E F]. 244 | Arguments eq_push_inv [A E x a E' x' a']. 245 | 246 | (** Simplification tactics *) 247 | 248 | Hint Rewrite 249 | concat_empty 250 | map_concat 251 | dom_empty dom_single dom_push dom_cons dom_concat dom_map : rew_env. 252 | 253 | Hint Rewrite <- concat_assoc : rew_env. 254 | 255 | Tactic Notation "simpl_env" := 256 | autorewrite with rew_env in *. 257 | 258 | Hint Extern 1 (_ # _) => simpl_env; notin_solve : core. 259 | 260 | (** The [env_fix] tactic is used to convert environments 261 | from [(x,a)::E] to [E & x ~ a]. *) 262 | 263 | (* Duplication in first two cases is to work around a Coq bug 264 | of the change tactic *) 265 | 266 | Ltac env_fix_core := 267 | let go := try env_fix_core in 268 | match goal with 269 | | H: context [(?x,?a)::?E] |- _ => 270 | ( (progress (change ((x,a)::E) with (E & x ~ a) in H)) 271 | || (progress (unsimpl (E & x ~ a) in H)) ); go 272 | | |- context [(?x,?a)::?E] => 273 | ( (progress (change ((x,a)::E) with (E & x ~ a))) 274 | || (progress (unsimpl (E & x ~ a))) ); go 275 | | H: context [@nil ((var * ?A)%type)] |- _ => 276 | progress (change (@nil ((var * A)%type)) with (@empty A) in H); go 277 | | |- context [@nil ((var * ?A)%type)] => 278 | progress (change (@nil ((var * A)%type)) with (@empty A)); go 279 | end. 280 | 281 | Ltac env_fix := try env_fix_core. 282 | 283 | 284 | (* ********************************************************************** *) 285 | (** ** Properties of Well-formedness and Freshness *) 286 | 287 | Section OkProperties. 288 | Variable A : Set. 289 | Implicit Types E F : env A. 290 | Implicit Types a b : A. 291 | Hint Constructors ok : core. 292 | 293 | (** Inversion for ok on concat *) 294 | 295 | Lemma ok_concat_inv : forall E F, 296 | ok (E & F) -> ok E /\ ok F. 297 | Proof. 298 | induction F as [|(x,a)]; simpl; env_fix; intros Ok. 299 | auto*. 300 | inversions Ok. split. auto*. apply* ok_push. 301 | Qed. 302 | 303 | (** Removing bindings preserves ok *) 304 | 305 | Lemma ok_remove : forall F E G, 306 | ok (E & F & G) -> ok (E & G). 307 | Proof. 308 | induction G as [|(y,a)]; simpl; env_fix; intros Ok. 309 | induction F as [|(y,a)]; simpl. 310 | auto. 311 | inversions* Ok. 312 | inversions* Ok. 313 | Qed. 314 | 315 | (** Mapping values preserves ok *) 316 | 317 | Lemma ok_map : forall E F (f : A -> A), 318 | ok (E & F) -> ok (E & map f F). 319 | intros. induction F as [|(y,a)]; simpl; env_fix. 320 | auto. 321 | rewrite <- concat_assoc in H. inversions* H. 322 | Qed. 323 | 324 | (** A binding in the middle of an environment has a var fresh 325 | from all the bindings before it *) 326 | 327 | Lemma fresh_mid : forall E F x a, 328 | ok (E & x ~ a & F) -> x # E. 329 | Proof. 330 | induction F; simpl; introv Ok; inversions Ok; env_fix. 331 | auto. 332 | inversions H1; env_fix. 333 | contradictions. fold (@empty A) in H0. apply* eq_empty_inv. 334 | simpls*. 335 | Qed. 336 | 337 | End OkProperties. 338 | 339 | (** Automation *) 340 | 341 | Hint Resolve fresh_mid ok_map : core. 342 | 343 | (* Hint Extern 1 (ok (?E & ?G)) => 344 | match goal with H: context [E & ?F & G] |- _ => 345 | apply (@ok_remove _ F) end. *) 346 | 347 | 348 | (* ********************************************************************** *) 349 | (** ** Properties of the binds relation *) 350 | 351 | Section BindsProperties. 352 | Variable A : Set. 353 | Implicit Types E F : env A. 354 | Implicit Types a b : A. 355 | 356 | Hint Extern 1 (_ \notin _) => notin_solve : core. 357 | 358 | (** Binds at head *) 359 | 360 | Lemma binds_head : forall x a E, 361 | binds x a (E & x ~ a). 362 | Proof. 363 | intros. unfold binds. simpl. case_var*. 364 | Qed. 365 | 366 | (** Binds in tail *) 367 | 368 | Lemma binds_tail : forall x y a b E, 369 | x <> y -> binds x a E -> binds x a (E & y ~ b). 370 | Proof. 371 | intros. unfold binds. simpl. case_var*. 372 | Qed. 373 | 374 | (** Binds is functional *) 375 | 376 | Lemma binds_func : forall x a b E, 377 | binds x a E -> binds x b E -> a = b. 378 | Proof. 379 | unfold binds. introv B1 B2. 380 | poses H (trans_eq (sym_eq B1) B2). 381 | inversion* H. 382 | Qed. 383 | 384 | (** No binding in an empty environment *) 385 | 386 | Lemma binds_empty : forall x a, 387 | binds x a empty -> False. 388 | Proof. 389 | introv H. inversions H. 390 | Qed. 391 | 392 | (** No binding on x if x is fresh from the environment *) 393 | 394 | Lemma binds_fresh : forall x a E, 395 | binds x a E -> x # E -> False. 396 | Proof. 397 | unfold binds. induction E as [|(y,b)]; simpl; intros Has Fresh. 398 | contradictions. 399 | case_var*. notin_contradiction. 400 | Qed. 401 | 402 | (** Binding on x preserved by concatenation if x is fresh 403 | for the extension *) 404 | 405 | Lemma binds_concat_fresh : forall x a E F, 406 | binds x a E -> x # F -> binds x a (E & F). 407 | Proof. 408 | unfold binds. induction F as [|(y,b)]; simpl; intros. 409 | auto. case_var*. notin_contradiction. 410 | Qed. 411 | 412 | (** Binding on x preserved by concatenation if the result 413 | of the concatenation is a well-formed environment *) 414 | 415 | Lemma binds_concat_ok : forall x a E F, 416 | binds x a E -> ok (E & F) -> binds x a (E & F). 417 | Proof. 418 | unfolds binds. induction F as [|(y,b)]; simpl; intros Map Ok. 419 | auto. 420 | inversions Ok. case_var. 421 | contradictions. apply* (@binds_fresh y a E). 422 | auto. 423 | Qed. 424 | 425 | (** Bindings preserved by pre-catenation *) 426 | 427 | Lemma binds_prepend : forall x a E F, 428 | binds x a F -> binds x a (E & F). 429 | Proof. 430 | unfold binds. induction F as [|(y,b)]; simpl; intros Map. 431 | contradictions. case_var*. 432 | Qed. 433 | 434 | (** Case analysis on the belonging of a binding to a concatenation *) 435 | 436 | Lemma binds_concat_inv : forall x a E F, 437 | binds x a (E & F) -> (x # F /\ binds x a E) \/ (binds x a F). 438 | Proof. 439 | unfold binds. induction F as [|(y,b)]; simpl; intros Map. 440 | left*. 441 | case_var. 442 | right*. 443 | destruct (IHF Map) as [[Fr Map2] | Map2]. left*. right*. 444 | Qed. 445 | 446 | (** Interaction of binds and map *) 447 | 448 | Lemma binds_map : forall x a f E, 449 | binds x a E -> binds x (f a) (map f E). 450 | Proof. 451 | unfold binds. induction E as [|(y,b)]; simpl; intros Map. 452 | contradictions. case_var*. inversions* Map. 453 | Qed. 454 | 455 | (** Retreive the binding on x from an environment [E & (x~a) & F] *) 456 | 457 | Lemma binds_mid : forall x a E F, 458 | ok (E & x ~ a & F) -> binds x a (E & x ~ a & F). 459 | Proof. 460 | unfold binds. induction F as [|(y,b)]; simpl; intros Ok. 461 | case_var*. 462 | inversions Ok. case_var. 463 | simpl_env. notin_contradiction. 464 | auto. 465 | Qed. 466 | 467 | (** The binding on x in the environment [E & (x~a) & F] can only 468 | be bound to a if that environment is well-formed *) 469 | 470 | Lemma binds_mid_eq : forall z a b E F, 471 | binds z a (E & z ~ b & F) -> ok (E & z ~ b & F) -> a = b. 472 | Proof. 473 | unfold binds. induction F as [|(y,c)]; simpl; intros Map Ok. 474 | case_var*. inversions* Map. 475 | inversions Ok. case_var. 476 | inversions Ok. simpl_env. notin_contradiction. 477 | auto. 478 | Qed. 479 | 480 | (** Inversion on a binding in an atomic environment *) 481 | 482 | Lemma binds_single_inv : forall x y a b, 483 | binds x a (y ~ b) -> x = y /\ a = b. 484 | Proof. 485 | unfold binds. simpl. intros. case_var. 486 | inversions* H. 487 | contradictions. 488 | Qed. 489 | 490 | (** Interaction between binds and the insertion of bindings. 491 | In theory we don't need this lemma since it would suffice 492 | to use the binds_cases tactics, but since weakening is a 493 | very common operation we provide a lemma for it. *) 494 | 495 | Lemma binds_weaken : forall x a E F G, 496 | binds x a (E & G) -> ok (E & F & G) -> 497 | binds x a (E & F & G). 498 | Proof. 499 | unfold binds. induction G as [|(y,b)]; simpl; intros Map Ok. 500 | apply* binds_concat_ok. 501 | inversions Ok. case_var*. 502 | Qed. 503 | 504 | (* Whether bindings are or not in the context is decidable *) 505 | 506 | Lemma binds_dec : forall E x a, 507 | (forall a1 a2 : A, {a1 = a2} + {a1 <> a2}) -> 508 | decidable (binds x a E). 509 | Proof. 510 | introv Dec. remember (get x E) as B. symmetry in HeqB. 511 | unfold binds. destruct B as [a'|]. 512 | destruct (Dec a a'). subst. 513 | left*. 514 | right*. intros H. congruence. 515 | right*. intros H. congruence. 516 | Qed. 517 | 518 | End BindsProperties. 519 | 520 | Arguments binds_concat_inv [A x a E F]. 521 | 522 | Hint Resolve binds_head binds_tail : core. 523 | 524 | 525 | (* ********************************************************************** *) 526 | (** ** Properties of environment inclusion *) 527 | 528 | Section ExtendsProperties. 529 | Variable A : Set. 530 | Implicit Types E F : env A. 531 | 532 | Lemma extends_self : forall E, 533 | extends E E. 534 | Proof. 535 | introz. auto. 536 | Qed. 537 | 538 | Lemma extends_push : forall E x a, 539 | x # E -> extends E (E & x ~ a). 540 | Proof. 541 | introz. apply* binds_tail. 542 | destruct (x0 == x). 543 | subst. contradictions. apply* binds_fresh. 544 | auto. 545 | Qed. 546 | 547 | Lemma binds_inj : forall E x a b, 548 | binds x a E -> binds x b E -> a = b. 549 | Proof. 550 | unfold binds. intros. congruence. 551 | Qed. 552 | 553 | Lemma extends_binds : forall E x a, 554 | binds x a E -> extends E (E & x ~ a). 555 | Proof. 556 | introz. unfolds binds. simpl. case_var. congruence. auto. 557 | Qed. 558 | 559 | End ExtendsProperties. 560 | 561 | Hint Resolve extends_self extends_push extends_binds : core. 562 | 563 | 564 | (* ********************************************************************** *) 565 | (** ** Iteration of pushing bindings in the environment *) 566 | 567 | Section IterPush. 568 | Variable A : Set. 569 | Implicit Types E F : env A. 570 | 571 | Definition iter_push (xs : list var) (vs : list A) := 572 | rev (combine xs vs). 573 | 574 | Lemma iter_push_nil : forall xs, 575 | iter_push xs nil = nil. 576 | Proof. 577 | induction xs; simpl*. 578 | Qed. 579 | 580 | Lemma iter_push_cons : forall x (v : A) xs vs, 581 | iter_push (x::xs) (v::vs) = (x ~ v) & (iter_push xs vs). 582 | Proof. 583 | auto*. 584 | Qed. 585 | 586 | Lemma ok_concat_iter_push : forall xs E (Us : list A), 587 | ok E -> 588 | fresh (dom E) (length xs) xs -> 589 | ok (E & iter_push xs Us). 590 | Proof. 591 | induction xs; simpl; intros. auto. 592 | destruct H0. destruct Us; simpls. auto. 593 | rewrite iter_push_cons. rewrite* <- concat_assoc. 594 | Qed. 595 | 596 | End IterPush. 597 | 598 | Hint Resolve ok_concat_iter_push : core. 599 | 600 | Section Fv_in. 601 | 602 | (* ********************************************************************** *) 603 | (** ** Gathering free variables of the values contained in an 604 | environment *) 605 | 606 | Variables (A : Set) (fv : A -> vars). 607 | 608 | (** Computing free variables contained in an environment. *) 609 | 610 | Fixpoint fv_in (E : env A) : vars := 611 | match E with 612 | | nil => {} 613 | | (x,a)::E' => fv a \u fv_in E' 614 | end. 615 | 616 | (** Specification of the above function in terms of [bind]. *) 617 | 618 | Lemma fv_in_spec : forall E, env_prop (fun a => fv a << fv_in E) E. 619 | Proof. 620 | induction E; intro; simpl; intros. elim H. 621 | destruct H; intro; intros. 622 | subst. now apply S.union_2. 623 | destruct a. 624 | apply S.union_3. 625 | puts (IHE _ _ H). 626 | auto with sets. 627 | Qed. 628 | 629 | End Fv_in. 630 | 631 | (* ********************************************************************** *) 632 | (** ** Tactics *) 633 | 634 | Opaque binds. 635 | 636 | (** [binds_get H as EQ] produces from an hypothesis [H] of 637 | the form [binds x a (E & x ~ b & F)] the equality [EQ: a = b]. *) 638 | 639 | Tactic Notation "binds_get" constr(H) "as" ident(EQ) := 640 | match type of H with binds ?z ?a (?E & ?x ~ ?b & ?F) => 641 | let K := fresh "K" in assert (K : ok (E & x ~ b & F)); 642 | [ auto | poses EQ (@binds_mid_eq _ z a b E F H K); clear K ] end. 643 | 644 | (** [binds_get H] expects an hypothesis [H] of the form 645 | [binds x a (E & x ~ b & F)] and substitute [a] for [b] in the goal. *) 646 | 647 | Tactic Notation "binds_get" constr(H) := 648 | let EQ := fresh in binds_get H as EQ; 649 | try match type of EQ with 650 | | ?f _ = ?f _ => inversions EQ; clear EQ 651 | | ?x = ?y => subst x end. 652 | 653 | (** [binds_single H] derives from an hypothesis [H] of the form 654 | [binds x a (y ~ b)] the equalities [x = y] and [a = b], then 655 | it substitutes [x] for [y] in the goal or deduce a contradiction 656 | if [x <> y] can be found in the context. *) 657 | 658 | Ltac binds_single H := 659 | match type of H with binds ?x ?a (?y ~ ?b) => 660 | destruct (binds_single_inv H); 661 | try discriminate; try subst y; 662 | try match goal with N: ?x <> ?x |- _ => 663 | contradictions; apply N; reflexivity end end. 664 | 665 | (** [binds_case H as B1 B2] derives from an hypothesis [H] of the form 666 | [binds x a (E & F)] two subcases: [B1: binds x a E] (with a freshness 667 | condition [x # F]) and [B2: binds x a F]. *) 668 | 669 | Tactic Notation "binds_case" constr(H) "as" ident(B1) ident(B2) := 670 | let Fr := fresh "Fr" in 671 | destruct (binds_concat_inv H) as [[Fr B1] | B2]. 672 | 673 | (** [binds_case H] makes a case analysis on an hypothesis [H] of the form 674 | [binds x a E] where E can be constructed using concatenation and 675 | singletons. It calls [binds_single] when reaching a singleton. *) 676 | 677 | Ltac binds_cases H := 678 | let go B := clear H; 679 | first [ binds_single B | binds_cases B | idtac ] in 680 | let B1 := fresh "B" in let B2 := fresh "B" in 681 | binds_case H as B1 B2; simpl_env; [ go B1 | go B2 ]. 682 | 683 | 684 | (** Automation *) 685 | 686 | Hint Resolve 687 | binds_concat_fresh binds_concat_ok 688 | binds_prepend binds_map : core. 689 | 690 | 691 | (* ********************************************************************** *) 692 | (** ** Other tactics (syntactical sugar for tactics) *) 693 | 694 | (** Tactic to apply an induction hypothesis modulo a rewrite of 695 | the associativity of the environment necessary to handle the 696 | inductive rules dealing with binders. [apply_ih_bind] is in 697 | fact just a syntactical sugar for [do_rew concat_assoc (eapply H)] 698 | which in turns stands for 699 | [rewrite concat_assoc; eapply H; rewrite <- concat_assoc]. *) 700 | 701 | Tactic Notation "apply_ih_bind" constr(H) := 702 | do_rew concat_assoc (eapply H). 703 | 704 | Tactic Notation "apply_ih_bind" "*" constr(H) := 705 | do_rew* concat_assoc (eapply H). 706 | 707 | (** Similar as the above, except in the case where there 708 | is also a map function, so we need to use [concat_assoc_map_push] 709 | for rewriting. *) 710 | 711 | Tactic Notation "apply_ih_map_bind" constr(H) := 712 | do_rew concat_assoc_map_push (eapply H). 713 | 714 | Tactic Notation "apply_ih_map_bind" "*" constr(H) := 715 | do_rew* concat_assoc_map_push (eapply H). 716 | 717 | (** [apply_empty] applies a lemma of the form "forall E:env, P E" 718 | in the particular case where E is empty. The tricky step is 719 | the simplification of "P empty" before the "apply" tactic is 720 | called, and this is necessary for Coq to recognize that the 721 | lemma indeed applies. *) 722 | 723 | Tactic Notation "apply_empty" constr(lemma) := 724 | let H := fresh in poses H (@lemma empty); 725 | simpl in H; eapply H; env_fix; clear H. 726 | 727 | Tactic Notation "apply_empty" "*" constr(lemma) := 728 | apply_empty lemma; auto*. 729 | 730 | 731 | End Env. 732 | -------------------------------------------------------------------------------- /Metatheory_Fresh.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Tactics to Automate Reasoning about Freshness * 3 | * Brian Aydemir & Arthur Charguéraud, July 2007, Coq v8.1 * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import Lib_Tactic List Metatheory_Var. 8 | 9 | 10 | Ltac fresh_simpl_to_notin_in_context := 11 | repeat (match goal with H: fresh _ _ _ |- _ => 12 | progress (simpl in H; destruct H) end). 13 | 14 | 15 | (* ********************************************************************** *) 16 | (** ** Tactics for notin *) 17 | 18 | (** For efficiency, we avoid rewrites by splitting equivalence. *) 19 | 20 | Lemma notin_singleton_r : forall x y, 21 | x \notin {{y}} -> x <> y. 22 | intros. rewrite* <- notin_singleton. 23 | Qed. 24 | 25 | Lemma notin_singleton_l : forall x y, 26 | x <> y -> x \notin {{y}}. 27 | intros. rewrite* notin_singleton. 28 | Qed. 29 | 30 | Lemma notin_singleton_swap : forall x y, 31 | x \notin {{y}} -> y \notin {{x}}. 32 | intros. apply notin_singleton_l. 33 | apply sym_not_eq. apply* notin_singleton_r. 34 | Qed. 35 | 36 | Lemma notin_union_r : forall x E F, 37 | x \notin (E \u F) -> (x \notin E) /\ (x \notin F). 38 | intros. rewrite* <- notin_union. 39 | Qed. 40 | 41 | Lemma notin_union_l : forall x E F, 42 | x \notin E -> x \notin F -> x \notin (E \u F). 43 | intros. rewrite* notin_union. 44 | Qed. 45 | 46 | (** Tactics to deal with notin. It interacts with union 47 | singleton and empty, but not inclusion. *) 48 | 49 | Ltac notin_solve_from x E H := 50 | match type of H with x \notin ?L => 51 | match L with context[E] => 52 | match L with 53 | | E => exact H 54 | | ?F \u ?G => 55 | let P := constr:(@notin_union_r x F G H) in 56 | let PF := eval simpl in (proj1 P) in 57 | let PG := eval simpl in (proj2 P) in 58 | solve [ notin_solve_from x E PF 59 | | notin_solve_from x E PG ] 60 | end 61 | end 62 | end. 63 | 64 | Ltac notin_solve_from_context x E := 65 | match goal with H: x \notin _ |- _ => 66 | notin_solve_from x E H end. 67 | 68 | Ltac notin_solve_one := 69 | match goal with 70 | | |- _ \notin {} => 71 | apply notin_empty 72 | | |- ?x \notin {{?y}} => (* by x <> y or y <> x *) 73 | apply notin_singleton_l; solve 74 | [ assumption | apply sym_not_eq; assumption ] 75 | | |- ?x \notin {{?y}} => (* by y \notin {{x}} *) 76 | apply notin_singleton_swap; notin_solve_from_context y ({{x}}) 77 | | |- ?x \notin ?E => (* default search *) 78 | notin_solve_from_context x E 79 | end. 80 | 81 | Ltac notin_simpl := 82 | try match goal with |- _ \notin (_ \u _) => 83 | apply notin_union_l; notin_simpl end. 84 | 85 | Ltac notin_simpl_hyps := (* forward-chaining *) 86 | try match goal with 87 | | H: _ \notin {} |- _ => 88 | clear H; notin_simpl_hyps 89 | | H: ?x \notin {{?y}} |- _ => 90 | puts (@notin_singleton_r x y H); 91 | clear H; notin_simpl_hyps 92 | | H: ?x \notin (?E \u ?F) |- _ => 93 | let H1 := fresh "Fr" in let H2 := fresh "Fr" in 94 | destruct (@notin_union_r x E F H) as [H1 H2]; 95 | clear H; notin_simpl_hyps 96 | end. 97 | 98 | Ltac notin_simpls := 99 | notin_simpl_hyps; notin_simpl. 100 | 101 | Ltac notin_solve := 102 | fresh_simpl_to_notin_in_context; 103 | notin_simpl; notin_solve_one. 104 | 105 | Ltac notin_contradiction := 106 | match goal with H: ?x \notin ?E |- _ => 107 | match E with context[x] => 108 | let K := fresh in assert (K : x \notin {{x}}); 109 | [ notin_solve_one 110 | | contradictions; apply (notin_same K) ] 111 | end 112 | end. 113 | 114 | Ltac notin_neq_solve := 115 | apply notin_singleton_r; notin_solve. 116 | 117 | Ltac fold_notin := 118 | repeat match goal with 119 | | H: context [?x \in ?E -> False] |- _ => 120 | fold (not (x \in E)) in H 121 | | |- context [?x \in ?E -> False] => 122 | fold (not (x \in E)) end. 123 | 124 | 125 | (* ********************************************************************** *) 126 | (** Demo for notin *) 127 | 128 | Lemma test_notin_solve_1 : forall x E F G, 129 | x \notin E \u F -> x \notin G -> x \notin (E \u G). 130 | intros. notin_solve. 131 | Qed. 132 | 133 | Lemma test_notin_solve_2 : forall x y E F G, 134 | x \notin E \u {{y}} \u F -> x \notin G -> 135 | x \notin {{y}} /\ y \notin {{x}}. 136 | intros. split. notin_solve. notin_solve. 137 | Qed. 138 | 139 | Lemma test_notin_solve_3 : forall x y, 140 | x <> y -> x \notin {{y}} /\ y \notin {{x}}. 141 | intros. split. notin_solve. notin_solve. 142 | Qed. 143 | 144 | Lemma test_notin_contradiction : forall x y E F G, 145 | x \notin (E \u {{x}} \u F) -> y \notin G. 146 | intros. notin_contradiction. 147 | Qed. 148 | 149 | Lemma test_neq_solve : forall x y E F, 150 | x \notin (E \u {{y}} \u F) -> y \notin E -> 151 | y <> x /\ x <> y. 152 | intros. split. notin_neq_solve. notin_neq_solve. 153 | Qed. 154 | 155 | 156 | 157 | (***************************************************************************) 158 | (** Automation: hints to solve "notin" subgoals automatically. *) 159 | 160 | Hint Extern 1 (_ \notin _) => notin_solve : core. 161 | Hint Extern 1 (_ \in _ -> False) => fold_notin : core. 162 | Hint Extern 1 (_ <> _ :> var) => notin_neq_solve : core. 163 | Hint Extern 1 (_ <> _ :> S.elt) => notin_neq_solve : core. 164 | 165 | 166 | (* ********************************************************************** *) 167 | (** ** Tactics for fresh *) 168 | 169 | Hint Extern 1 (fresh _ _ _) => simpl : core. 170 | 171 | Lemma fresh_union_r : forall xs L1 L2 n, 172 | fresh (L1 \u L2) n xs -> fresh L1 n xs /\ fresh L2 n xs. 173 | Proof. 174 | induction xs; simpl; intros; destruct n; 175 | try solve [ contradictions* ]. auto. 176 | destruct H. splits~. 177 | forward* (@IHxs (L1 \u {{a}}) L2 n). 178 | rewrite union_comm. 179 | rewrite union_comm_assoc. 180 | rewrite* union_assoc. 181 | forward* (@IHxs L1 (L2 \u {{a}}) n). 182 | rewrite* union_assoc. 183 | Qed. 184 | 185 | Lemma fresh_union_l : forall xs L1 L2 n, 186 | fresh L1 n xs -> fresh L2 n xs -> fresh (L1 \u L2) n xs. 187 | Proof. 188 | induction xs; simpl; intros; destruct n; 189 | try solve [ contradictions* ]. auto. 190 | destruct H. destruct H0. split2~. 191 | forward~ (@IHxs (L1 \u {{a}}) (L2 \u {{a}}) n). 192 | intros K. 193 | rewrite <- (union_same {{a}}). 194 | rewrite union_assoc. 195 | rewrite <- (@union_assoc L1). 196 | rewrite (@union_comm L2). 197 | rewrite (@union_assoc L1). 198 | rewrite* <- union_assoc. 199 | Qed. 200 | 201 | Lemma fresh_empty : forall L n xs, 202 | fresh L n xs -> fresh {} n xs. 203 | Proof. 204 | intros. rewrite <- (@union_empty_r L) in H. 205 | destruct* (fresh_union_r _ _ _ _ H). 206 | Qed. 207 | 208 | Lemma fresh_length : forall xs L n, 209 | fresh L n xs -> n = length xs. 210 | Proof. 211 | induction xs; simpl; intros L n Fr; destruct n; 212 | try solve [contradictions*]. auto. 213 | rewrite* <- (@IHxs (L \u {{a}}) n). 214 | Qed. 215 | 216 | Lemma fresh_resize : forall n xs L, 217 | fresh L n xs -> fresh L (length xs) xs. 218 | Proof. 219 | introv Fr. rewrite* <- (fresh_length _ _ _ Fr). 220 | Qed. 221 | 222 | Ltac fresh_solve_from xs n E H := 223 | match type of H with fresh ?L _ _ => 224 | match L with context[E] => 225 | match L with 226 | | E => exact H 227 | | ?F \u ?G => 228 | let P := constr:(@fresh_union_r xs F G n H) in 229 | let PF := eval simpl in (proj1 P) in 230 | let PG := eval simpl in (proj2 P) in 231 | solve [ fresh_solve_from xs n E PF 232 | | fresh_solve_from xs n E PG ] 233 | end 234 | end 235 | end. 236 | 237 | Ltac fresh_solve_from_context xs n E := 238 | match goal with H: fresh _ n xs |- _ => 239 | fresh_solve_from xs n E H end. 240 | 241 | Ltac fresh_solve_one := 242 | assumption || 243 | match goal with 244 | | |- fresh {} ?n ?xs => 245 | match goal with H: fresh ?L n xs |- _ => 246 | apply (@fresh_empty L n xs H) end 247 | (* | H: fresh _ ?n ?xs |- fresh ?E ?n ?xs => 248 | fresh_solve_from xs n E H *) 249 | | |- fresh _ (length ?xs) ?xs => 250 | match goal with H: fresh _ ?n xs |- _ => 251 | progress (apply (@fresh_resize n)); fresh_solve_one end 252 | end. 253 | 254 | Ltac fresh_simpl := 255 | try match goal with |- fresh (_ \u _) _ _ => 256 | apply fresh_union_l; fresh_simpl end. 257 | 258 | Ltac fresh_split := 259 | match goal with 260 | | H: fresh (?L1 \u ?L2) ?n ?xs |- fresh _ _ ?xs => 261 | destruct (fresh_union_r xs L1 L2 n H); clear H; fresh_split 262 | | _ => try fresh_simpl 263 | end. 264 | 265 | Ltac fresh_simpl_to_notin_in_goal := 266 | simpl; splits. 267 | 268 | Ltac fresh_simpl_to_notin_solve := 269 | fresh_simpl_to_notin_in_context; 270 | fresh_simpl_to_notin_in_goal; 271 | notin_solve. 272 | 273 | Ltac fresh_solve := 274 | (fresh_split; fresh_solve_one) || (fresh_simpl; fresh_simpl_to_notin_solve). 275 | 276 | 277 | (* ********************************************************************** *) 278 | (** Demo for notin *) 279 | 280 | Lemma test_fresh_solve_1 : forall xs L1 L2 n, 281 | fresh (L1 \u L2) n xs -> fresh L1 n xs. 282 | Proof. 283 | intros. fresh_solve. 284 | Qed. 285 | 286 | Lemma test_fresh_solve_2 : forall xs L1 L2 n, 287 | fresh (L1 \u L2) n xs -> fresh L2 n xs. 288 | Proof. 289 | intros. fresh_solve. 290 | Qed. 291 | 292 | Lemma test_fresh_solve_3 : forall xs L1 L2 n, 293 | fresh (L1 \u L2) n xs -> fresh {} n xs. 294 | Proof. 295 | intros. fresh_solve. 296 | Qed. 297 | 298 | Lemma test_fresh_solve_4 : forall xs L1 L2 n, 299 | fresh (L1 \u L2) n xs -> fresh L1 (length xs) xs. 300 | Proof. 301 | intros. fresh_solve. 302 | Qed. 303 | 304 | (***************************************************************************) 305 | (** Automation: hints to solve "fresh" subgoals automatically. *) 306 | 307 | Hint Extern 1 (fresh _ _ _) => fresh_solve : core. 308 | 309 | 310 | -------------------------------------------------------------------------------- /Metatheory_Var.v: -------------------------------------------------------------------------------- 1 | (*************************************************************************** 2 | * Generic Variables for Programming Language Metatheory * 3 | * Brian Aydemir & Arthur Charguéraud, July 2007, Coq v8.1 é * 4 | ***************************************************************************) 5 | 6 | Set Implicit Arguments. 7 | Require Import List Omega OrderedType OrderedTypeEx. 8 | Require Import Lib_Tactic Lib_ListFacts Lib_FinSet Lib_FinSetImpl. 9 | Require Export Lib_ListFactsMore. 10 | 11 | (* ********************************************************************** *) 12 | (** * Abstract Definition of Variables *) 13 | 14 | Module Type VARIABLES. 15 | 16 | (** We leave the type of variables abstract. *) 17 | 18 | Parameter var : Set. 19 | 20 | (** This type is inhabited. *) 21 | 22 | Parameter var_default : var. 23 | 24 | (** Variables are ordered. *) 25 | 26 | Declare Module Var_as_OT : UsualOrderedType with Definition t := var. 27 | 28 | (** We can form sets of variables. *) 29 | 30 | Declare Module Import VarSet : FinSet with Module E := Var_as_OT. 31 | Local Open Scope set_scope. 32 | 33 | Definition vars := VarSet.S.t. 34 | 35 | (** Finally, we have a means of generating fresh variables. *) 36 | 37 | Parameter var_generate : vars -> var. 38 | Parameter var_generate_spec : forall E, (var_generate E) \notin E. 39 | Parameter var_fresh : forall (L : vars), { x : var | x \notin L }. 40 | 41 | (** Variables can be enumerated *) 42 | 43 | Parameter var_of_Z : Z -> var. 44 | Parameter Z_of_var : var -> Z. 45 | 46 | End VARIABLES. 47 | 48 | 49 | (* ********************************************************************** *) 50 | (** * Concrete Implementation of Variables *) 51 | 52 | Module Variables : VARIABLES. 53 | 54 | Local Open Scope Z_scope. 55 | 56 | Definition var := Z. 57 | 58 | Definition var_default : var := 0. 59 | 60 | Definition var_of_Z x : var := x. 61 | Definition Z_of_var x : Z := x. 62 | 63 | Module Var_as_OT : UsualOrderedType with Definition t := var := Z_as_OT. 64 | 65 | Module Import VarSet : FinSet with Module E := Var_as_OT := 66 | Lib_FinSetImpl.Make Var_as_OT. 67 | 68 | Local Open Scope set_scope. 69 | 70 | Definition vars := VarSet.S.t. 71 | 72 | Lemma max_lt_l : 73 | forall (x y z : Z), x <= y -> x <= Z.max y z. 74 | Proof. 75 | intros. 76 | apply (Z.le_trans _ _ _ H). 77 | apply Z.le_max_l. 78 | Qed. 79 | 80 | Lemma finite_nat_list_max : forall (l : list Z), 81 | { n : Z | forall x, In x l -> x <= n }. 82 | Proof. 83 | induction l as [ | l ls IHl ]. 84 | exists 0; intros x H; inversion H. 85 | inversion IHl as [x H]. 86 | exists (Z.max x l); intros y J; simpl in J; inversion J. 87 | subst; apply Z.le_max_r. 88 | assert (y <= x); auto using max_lt_l. 89 | Qed. 90 | 91 | Lemma finite_nat_list_max' : forall (l : list Z), 92 | { n : Z | ~ In n l }. 93 | Proof. 94 | intros l. 95 | case (finite_nat_list_max l); intros x H. 96 | exists (x+1). 97 | intros J. 98 | assert (K := H _ J); omega. 99 | Qed. 100 | 101 | Definition var_generate (L : vars) : var := 102 | proj1_sig (finite_nat_list_max' (S.elements L)). 103 | 104 | Lemma var_generate_spec : forall E, (var_generate E) \notin E. 105 | Proof. 106 | unfold var_generate. intros E. 107 | destruct (finite_nat_list_max' (S.elements E)) as [n pf]. 108 | simpl. intros a. 109 | assert (In n (S.elements E)). rewrite <- InA_iff_In. 110 | auto using S.elements_1. 111 | intuition. 112 | Qed. 113 | 114 | Lemma var_fresh : forall (L : vars), { x : var | x \notin L }. 115 | Proof. 116 | intros L. exists (var_generate L). apply var_generate_spec. 117 | Qed. 118 | 119 | End Variables. 120 | 121 | 122 | (* ********************************************************************** *) 123 | (** * Properties of variables *) 124 | 125 | Export Variables. 126 | Export Variables.VarSet. 127 | Module Export VarSetFacts := FinSetFacts VarSet. 128 | 129 | Open Scope set_scope. 130 | 131 | (** Equality on variables is decidable. *) 132 | 133 | Module Import Var_as_OT_Facts := OrderedTypeFacts Variables.Var_as_OT. 134 | 135 | Lemma eq_var_dec : forall x y : var, {x = y} + {x <> y}. 136 | Proof. 137 | exact Var_as_OT_Facts.eq_dec. 138 | Qed. 139 | 140 | (* ********************************************************************** *) 141 | (** ** Dealing with list of variables *) 142 | 143 | (** Freshness of n variables from a set L and from one another. *) 144 | 145 | Fixpoint fresh (L : vars) (n : nat) (xs : list var) {struct xs} : Prop := 146 | match xs, n with 147 | | nil, O => True 148 | | x::xs', S n' => x \notin L /\ fresh (L \u {{x}}) n' xs' 149 | | _,_ => False 150 | end. 151 | 152 | Hint Extern 1 (fresh _ _ _) => simpl : core. 153 | 154 | (** Triviality : If a list xs contains n fresh variables, then 155 | the length of xs is n. *) 156 | 157 | Lemma fresh_length : forall xs L n, 158 | fresh L n xs -> n = length xs. 159 | Proof. 160 | induction xs; simpl; intros; destruct n; 161 | try solve [ contradictions* | f_equal* ]. 162 | Qed. 163 | 164 | (* It is possible to build a list of n fresh variables. *) 165 | 166 | Lemma var_freshes : forall L n, 167 | { xs : list var | fresh L n xs }. 168 | Proof. 169 | intros. gen L. induction n; intros L. 170 | exists* (nil : list var). 171 | destruct (var_fresh L) as [x Fr]. 172 | destruct (IHn (L \u {{x}})) as [xs Frs]. 173 | exists* (x::xs). 174 | Qed. 175 | 176 | 177 | (* ********************************************************************** *) 178 | (** ** Tactics: Case Analysis on Variables *) 179 | 180 | (** We define notations for the equality of variables (our free variables) 181 | and for the equality of naturals (our bound variables represented using 182 | de Bruijn indices). *) 183 | 184 | Notation "x == y" := (eq_var_dec x y) (at level 67). 185 | Notation "i === j" := (Peano_dec.eq_nat_dec i j) (at level 70). 186 | 187 | (** Tactic for comparing two bound or free variables. *) 188 | 189 | Ltac case_nat := 190 | let destr x y := destruct (x === y); [try subst x | idtac] in 191 | match goal with 192 | | H: context [?x === ?y] |- _ => destr x y 193 | | |- context [?x === ?y] => destr x y 194 | end. 195 | 196 | Tactic Notation "case_nat" "*" := case_nat; auto*. 197 | 198 | Tactic Notation "case_var" := 199 | let destr x y := destruct (x == y); [try subst x | idtac] in 200 | match goal with 201 | | H: context [?x == ?y] |- _ => destr x y 202 | | |- context [?x == ?y] => destr x y 203 | end. 204 | 205 | Tactic Notation "case_var" "*" := case_var; auto*. 206 | 207 | 208 | (* ********************************************************************** *) 209 | (** ** Tactics: Picking Names Fresh from the Context *) 210 | 211 | (** [gather_vars_for_type T F] return the union of all the finite sets 212 | of variables [F x] where [x] is a variable from the context such that 213 | [F x] type checks. In other words [x] has to be of the type of the 214 | argument of [F]. The resulting union of sets does not contain any 215 | duplicated item. This tactic is an extreme piece of hacking necessary 216 | because the tactic language does not support a "fold" operation on 217 | the context. *) 218 | 219 | Ltac gather_vars_with F := 220 | let rec gather V := 221 | match goal with 222 | | H: ?S |- _ => 223 | let FH := constr:(F H) in 224 | match V with 225 | | {} => gather FH 226 | | context [FH] => fail 1 227 | | _ => gather (FH \u V) 228 | end 229 | | _ => V 230 | end in 231 | let L := gather {} in eval simpl in L. 232 | 233 | (** [beautify_fset V] assumes that [V] is built as a union of finite 234 | sets and return the same set cleaned up: empty sets are removed and 235 | items are laid out in a nicely parenthesized way *) 236 | 237 | Ltac beautify_fset V := 238 | let rec go Acc E := 239 | match E with 240 | | ?E1 \u ?E2 => let Acc1 := go Acc E1 in 241 | go Acc1 E2 242 | | {} => Acc 243 | | ?E1 => match Acc with 244 | | {} => E1 245 | | _ => constr:(Acc \u E1) 246 | end 247 | end 248 | in go {} V. 249 | 250 | (** [pick_fresh_gen L Y] expects [L] to be a finite set of variables 251 | and adds to the context a variable with name [Y] and a proof that 252 | [Y] is fresh for [L]. *) 253 | 254 | Ltac pick_fresh_gen L Y := 255 | let Fr := fresh "Fr" in 256 | let L := beautify_fset L in 257 | (destruct (var_fresh L) as [Y Fr]). 258 | 259 | (** [pick_fresh_gens L n Y] expects [L] to be a finite set of variables 260 | and adds to the context a list of variables with name [Y] and a proof 261 | that [Y] is of length [n] and contains variable fresh for [L] and 262 | distinct from one another. *) 263 | 264 | Ltac pick_freshes_gen L n Y := 265 | let Fr := fresh "Fr" in 266 | let L := beautify_fset L in 267 | (destruct (var_freshes L n) as [Y Fr]). 268 | 269 | (** Demo of pick_fresh_gen *) 270 | 271 | Ltac test_pick_fresh_filter Y := 272 | let A := gather_vars_with (fun x : vars => x) in 273 | let B := gather_vars_with (fun x : var => {{ x }}) in 274 | let C := gather_vars_with (fun x : var => {}) in 275 | pick_fresh_gen (A \u B \u C) Y. 276 | 277 | Lemma test_pick_fresh : forall (x y z : var) (L1 L2 L3: vars), True. 278 | Proof. 279 | intros. test_pick_fresh_filter k. auto. 280 | Qed. 281 | 282 | (** The above invokation of [pick_fresh] generates a 283 | variable [k] and the hypothesis 284 | [k \notin L1 \u L2 \u L3 \u {{x}} \u {{y}} \u {{z}}] *) 285 | 286 | 287 | (* ********************************************************************** *) 288 | (** ** Tactics: Applying Lemmas With Quantification Over Cofinite Sets *) 289 | 290 | (** [apply_fresh_base] tactic is a helper to build tactics that apply an 291 | inductive constructor whose first argument should be instanciated 292 | by the set of names already used in the context. Those names should 293 | be returned by the [gather] tactic given in argument. For each premise 294 | of the inductive rule starting with an universal quantification of names 295 | outside the set of names instanciated, a subgoal with be generated by 296 | the application of the rule, and in those subgoal we introduce the name 297 | quantified as well as its proof of freshness. *) 298 | 299 | Ltac apply_fresh_base_simple lemma gather := 300 | let L0 := gather in let L := beautify_fset L0 in 301 | first [apply (@lemma L) | eapply (@lemma L)]. 302 | 303 | Ltac apply_fresh_base lemma gather var_name := 304 | apply_fresh_base_simple lemma gather; 305 | try match goal with |- forall _, _ \notin _ -> _ => 306 | let Fr := fresh "Fr" in intros var_name Fr end. 307 | 308 | 309 | (** [inst_notin H y as H'] expects [H] to be of the form 310 | [forall x, x \notin L, P x] and creates an hypothesis [H'] 311 | of type [P y]. It tries to prove the subgoal [y \notin L] 312 | by [auto]. This tactic is very useful to apply induction 313 | hypotheses given in the cases with binders. *) 314 | 315 | Tactic Notation "inst_notin" constr(lemma) constr(var) 316 | "as" ident(hyp_name) := 317 | let go L := let Fr := fresh in assert (Fr : var \notin L); 318 | [ auto | poses hyp_name (@lemma var Fr); clear Fr ] in 319 | match type of lemma with 320 | | forall _, _ \notin ?L -> _ => go L 321 | | forall _, (_ \in ?L -> False) -> _ => go L 322 | end. 323 | 324 | Tactic Notation "inst_notin" "*" constr(lemma) constr(var) 325 | "as" ident(hyp_name) := 326 | inst_notin lemma var as hyp_name; auto*. 327 | -------------------------------------------------------------------------------- /Omega.v: -------------------------------------------------------------------------------- 1 | Require Export ZArith Lia. 2 | Ltac omega := lia. 3 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | ## A Certified Interpreter for ML with Structural Polymorphism 2 | ### Jacques Garrigue, updated April 2025 3 | 4 | The files in this archive contain a proof of type soundness for structural 5 | polymorphism based on local constraints [1], together with soundness 6 | and principality of the type inference algorithm, and a certified interpreter. 7 | See the 8 | [support page](http://www.math.nagoya-u.ac.jp/~garrigue/papers/aplas2010.html) 9 | for detailed information. 10 | 11 | I started from a proof of type soundness for core ML by Arthur 12 | Chargueraud, which accompanied "Engineering formal metatheory" [2]. 13 | The library files (Lib_* and Metatheory*) are almost untouched. 14 | For compatibility we also copied FSetList.v from Coq 8.2. 15 | 16 | The new files are as follows. 17 | * `Metatheory_SP`: new generic lemmas used in developments 18 | * `Cardinal`: lemmas about finite set cardinals 19 | * `ML_SP_Definitions`: basic definitions 20 | * `ML_SP_Infrastructure`: structural lemmas on kinds and types 21 | * `ML_SP_Soundness`: lemmas on derivations and proof of type soundness 22 | * `ML_SP_Eval`: proof of type soundness for a stack-based evaluator 23 | * `ML_SP_Rename`: renaming lemmas, and Gc elimination 24 | * `ML_SP_Unify`: soundness and completeness of unification 25 | * `ML_SP_Inference`: soundness and principality of type inference 26 | * `ML_SP_Domain`: instantiation of all the above proofs to polymorphic variants 27 | * `ML_SP_Unify/Inference_wf`: termination counters are in Prop 28 | 29 | Of the above, Definitions, Infrastructure and Soundness are base on 30 | the core ML proof, but were heavily modified. 31 | Eval, Unify, Inference and Domain are completely new. 32 | 33 | All the above development were checked with coq 8.2010. 34 | (Porting to 8.20.1 is the only change wrt. the version of september 2010) 35 | You can compile them with "sh build.sh" 36 | 37 | You can also play with the type checker. It does not work inside Coq 38 | at this point, but you should compile typinf.mli and typinf.ml 39 | (obtained by running ML_SP_Extraction.v), and look at use_typinf.ml for how 40 | to use them. (It contains a number of petty printers and conversion 41 | functions that make things easier). 42 | 43 | Here are all the steps: (the first 3 steps are optional) 44 | ``` 45 | $ sh mkmakefile.sh 46 | $ make 47 | $ make html 48 | $ ocamlc -c typinf.mli 49 | $ ocamlc -c typinf.ml 50 | $ ocaml 51 | Objective Caml version ... 52 | # #use "use_typinf.ml";; 53 | ``` 54 | 55 | [1] Jacques Garrigue: A Certified Interpreter of ML with Structural 56 | Polymorphism and Recursive Types. 57 | Mathematical Structures in Computer Science, November 2014, pages 1-25. 58 | Earlier version presented at APLAS'10, Shanghai, China, November 2010. 59 | [link](http://www.math.nagoya-u.ac.jp/~garrigue/papers/aplas2010.html) 60 | 61 | [2] Brian Aydemir, Arthur Chargueraud, Benjamin C. Pierce, Randy Pollack 62 | and Stephanie Weirich: Engineering Formal Metatheory. POPL'08. 63 | [link](http://www.chargueraud.org/arthur/research/2007/binders/) 64 | 65 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . StructPoly 2 | -------------------------------------------------------------------------------- /build.sh: -------------------------------------------------------------------------------- 1 | set -e 2 | wait="$1" 3 | makeit() { 4 | for i in "$@"; do 5 | if test "$wait" = "$i"; then wait=""; fi 6 | if test -z "$wait"; then echo coqc "$i" ; coqc "$i" ; fi 7 | done; } 8 | makeit Lib_Tactic Lib_ListFacts Lib_ListFactsMore 9 | makeit Lib_MyFSetInterface Lib_FinSet Lib_FinSetImpl 10 | makeit Metatheory_Var Metatheory_Fresh Metatheory_Env 11 | makeit Metatheory_SP Metatheory 12 | makeit ML_SP_Definitions ML_SP_Infrastructure 13 | makeit ML_SP_Soundness ML_SP_Rename ML_SP_Eval 14 | makeit Cardinal ML_SP_Unify_wf ML_SP_Inference_wf 15 | makeit ML_SP_Domain 16 | -------------------------------------------------------------------------------- /mkmakefile.sh: -------------------------------------------------------------------------------- 1 | coq_makefile -R . StructPoly FSetList.v Omega.v Lib*.v Metatheory*.v ML_SP*.v Cardinal.v -o Makefile 2 | 3 | -------------------------------------------------------------------------------- /times.txt: -------------------------------------------------------------------------------- 1 | coq-8.2: 2 | 628s on marie (Core solo 1.2GHz) 3 | 390s on hugo (Core duo 2GHz) 4 | 310s on tagen27 (Core duo 2.53GHz) 5 | coq-8.2pl1: 6 | 346s on tagen27 7 | -------------------------------------------------------------------------------- /use_typinf.ml: -------------------------------------------------------------------------------- 1 | #load"typinf.cmo";; 2 | open Format;; 3 | open Typinf;; 4 | open Cstr;; 5 | open Infer2;; 6 | open Infer.Unify.MyEval.Rename.Sound.Infra.Defs;; 7 | open Infer.Unify.MyEval;; 8 | open Const;; 9 | open Variables.VarSet;; 10 | 11 | let rec list_of_coqlist = function 12 | Nil -> [] 13 | | Cons (a, l) -> a::list_of_coqlist l;; 14 | 15 | (** Wrap typinf for easier use *) 16 | let typinf2 env trm = 17 | match typinf1 env trm with 18 | | Inl (Pair (kenv, typ)) -> 19 | List.map (fun (Pair(a,b)) -> a,b) (list_of_coqlist kenv), typ 20 | | Inr Left -> failwith "Environment is not closed" 21 | | Inr Right -> failwith "Type Error";; 22 | 23 | (** Conversion functions *) 24 | let rec int_of_nat = function O -> 0 | S x -> succ (int_of_nat x);; 25 | let rec nat_of_int n = if n <= 0 then O else S (nat_of_int (n-1));; 26 | 27 | let rec coqlist_of_list = function 28 | [] -> Nil 29 | | a :: l -> Cons (a, coqlist_of_list l);; 30 | 31 | let rec int_of_pos = function 32 | XH -> 1 33 | | XO p -> int_of_pos p * 2 34 | | XI p -> int_of_pos p * 2 + 1 35 | let rec pos_of_int n = 36 | if n <= 1 then XH else 37 | let p = pos_of_int (n/2) in 38 | if n mod 2 = 0 then XO p else XI p 39 | let int_of_z = function 40 | Z0 -> 0 41 | | Zpos p -> int_of_pos p 42 | | Zneg p -> - int_of_pos p 43 | let z_of_int z = 44 | if z < 0 then Zneg (pos_of_int (-z)) else 45 | if z > 0 then Zpos (pos_of_int z) else Z0 46 | ;; 47 | 48 | (** Abbreviations and pretty printers *) 49 | 50 | (* nat and var *) 51 | let rec omega = S omega;; (* In ocaml we can define infinity! *) 52 | let print_nat ppf n = 53 | match n with 54 | | S m when m == n -> fprintf ppf "omega" 55 | | _ -> fprintf ppf "%i" (int_of_nat n);; 56 | let print_z ppf z = fprintf ppf "%i" (int_of_z z);; 57 | let var n = Variables.var_of_Z (z_of_int n) 58 | let print_var ppf v = print_z ppf (Variables.coq_Z_of_var v);; 59 | #install_printer print_nat;; 60 | #install_printer print_var;; 61 | 62 | (* Types *) 63 | let tv n = Coq_typ_fvar (var n) 64 | let bv n = Coq_typ_bvar (nat_of_int n) 65 | let (@>) t1 t2 = Coq_typ_arrow (t1, t2) 66 | let any = None 67 | let rec type_rec lv ppf = function 68 | | Coq_typ_bvar n -> fprintf ppf "bv %a" print_nat n 69 | | Coq_typ_fvar x -> fprintf ppf "tv %a" print_var x 70 | | Coq_typ_arrow (t1, t2) as t -> 71 | if lv > 0 then fprintf ppf "(%a)" (type_rec 0) t else 72 | fprintf ppf "@[%a @@>@ %a@]" (type_rec 1) t1 (type_rec 0) t2 73 | let print_type = type_rec 0 74 | let rec print_list pp ppf = function 75 | Nil -> () 76 | | Cons (a, Nil) -> pp ppf a 77 | | Cons (a, l) -> fprintf ppf "%a;@ %a" pp a (print_list pp) l;; 78 | let string_of_sort = function 79 | | Ksum -> "Ksum" 80 | | Kprod -> "Kprod" 81 | | Kbot -> "Kbot" 82 | let print_set ppf s = 83 | Format.fprintf ppf "{@[%a@]}" (print_list print_nat) s 84 | let print_set_option ppf = function 85 | | None -> fprintf ppf "any" 86 | | Some s -> print_set ppf s 87 | let print_type_list = 88 | print_list (fun ppf (Pair(n,t)) -> 89 | fprintf ppf "@[%a =>@ %a@]" print_nat n print_type t) 90 | let print_ckind ppf k = 91 | let c = k.kind_cstr in 92 | fprintf ppf "@[<1><%s,@ %a,@ %a,@ {%a}>@]" 93 | (string_of_sort c.cstr_sort) 94 | print_set c.cstr_low 95 | print_set_option c.cstr_high 96 | print_type_list k.kind_rel 97 | let print_kind ppf = function 98 | | None -> fprintf ppf "any" 99 | | Some k -> print_ckind ppf k;; 100 | #install_printer print_type;; 101 | #install_printer print_kind;; 102 | 103 | (* Terms *) 104 | let app a b = Coq_trm_app (a,b) 105 | let apps a l = List.fold_left app a l 106 | let abs a = Coq_trm_abs a 107 | let bvar n = Coq_trm_bvar (nat_of_int n) 108 | let matches l = apps 109 | (Coq_trm_cst (Const.Coq_matches (coqlist_of_list (List.map nat_of_int l)))) 110 | let tag n = 111 | app (Coq_trm_cst (Const.Coq_tag (nat_of_int n))) 112 | let record l = apps 113 | (Coq_trm_cst (Const.Coq_record (coqlist_of_list (List.map nat_of_int l)))) 114 | let sub n = 115 | app (Coq_trm_cst (Const.Coq_sub (nat_of_int n))) 116 | let recf = app (Coq_trm_cst Const.Coq_recf);; 117 | 118 | 119 | (* First example: (Coq_tag A) is a function constant, which takes any value 120 | and returns a polymorphic variant A with this value as argument *) 121 | (* This example is equivalent to the ocaml term [fun x -> `A0 x] *) 122 | typinf2 Nil (abs (tag 0 (bvar 0)));; 123 | 124 | (* Second example: (Coq_matches [A1; ..; An]) is a n+1-ary function constant 125 | which takes n functions and a polymorphic variant as argument, and 126 | (fi v) if the argument was (Ai v). *) 127 | (* This example is equivalent to [function `A20 x -> x | `A21 x -> x] *) 128 | let trm = matches [20; 21] [abs (bvar 0); abs (bvar 0)];; 129 | typinf2 Nil trm;; 130 | 131 | (* Another example, producing a recursive type *) 132 | (* OCaml equivalent: [fun x -> match x with `A20 y -> y | `A21 y -> y] *) 133 | let trm2 = 134 | abs (matches [20; 21] [abs (bvar 0); abs (bvar 0); bvar 0]) ;; 135 | typinf2 Nil trm2;; 136 | 137 | let trm3 = app trm2 (tag 20 (tag 21 (abs (bvar 0)))) ;; 138 | typinf2 Nil trm3;; 139 | 140 | let r1 = eval1 Nil trm3 (nat_of_int 10);; 141 | let r2 = eval1 Nil trm3 (nat_of_int 20);; 142 | let r3 = eval1 Nil trm3 omega;; 143 | 144 | (* More advanced example: the reverse function *) 145 | (* First define cons and nil. Since there is no unit we use a record *) 146 | let cons x y = tag 1 (record [0;1] [x;y]) 147 | let nil = tag 0 (record [] []);; 148 | typinf2 Nil nil;; 149 | typinf2 Nil (cons nil nil);; 150 | 151 | (* The rev_append function. DeBruijn indices can be tricky... *) 152 | let rev_append = recf 153 | (abs 154 | (abs 155 | (abs 156 | (matches [0;1] 157 | [abs (bvar 1); 158 | abs (apps (bvar 3) 159 | [sub 1 (bvar 0); cons (sub 0 (bvar 0)) (bvar 1)]); 160 | bvar 1])))) ;; 161 | typinf2 Nil rev_append;; 162 | 163 | (* A list of tagged units *) 164 | let mylist = 165 | app (abs (cons (tag 3 (bvar 0)) 166 | (cons (tag 4 (bvar 0)) (cons (tag 5 (bvar 0)) nil)))) 167 | (record [] []);; 168 | typinf2 Nil mylist;; 169 | 170 | (* Infer types and run! *) 171 | let rlist = apps rev_append [mylist; nil];; 172 | let t = typinf2 Nil rlist;; 173 | let r = eval1 Nil rlist (nat_of_int 134);; 174 | --------------------------------------------------------------------------------