]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/coq-contribs/LAMBDA-TYPES/contexts_defs.v
ocaml 3.09 transition
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / contexts_defs.v
index ec1f9e75866c5766667023de5bdc51884c5606cb..a9a689235c8f71f8e373e4fba9b1be1a7190ed63 100644 (file)
@@ -2,11 +2,17 @@
 
 Require Export terms_defs.
 
-      Inductive Set C := CSort : nat -> C
-                      |  CTail : C -> K -> T -> C.
+      Inductive Set C := CSort: nat -> C
+                      |  CTail: C -> K -> T -> C.
 
       Hint f3CKT : ltlc := Resolve (f_equal3 C K T).
 
+      Tactic Definition CGenBase :=
+         Match Context With
+            | [ H: (CSort ?) = (CSort ?) |- ? ]         -> Inversion H; Clear H
+            | [ H: (CTail ? ? ?) = (CTail ? ? ?) |- ? ] -> Inversion H; Clear H
+           | _                                         -> TGenBase.
+
       Definition r: K -> nat -> nat := [k;i] Cases k of
          | (Bind _) => i
          | (Flat _) => (S i)
@@ -58,11 +64,11 @@ Require Export terms_defs.
 
    Section app_props. (******************************************************)
 
-      Theorem app_csort : (t:?; i,n:?) (app (CSort n) i t) = t.
+      Theorem app_csort: (t:?; i,n:?) (app (CSort n) i t) = t.
       XElim i; Intros; Simpl; XAuto.
       Qed.
 
-      Theorem app_O : (c:?; t:?) (app c (0) t) = t.
+      Theorem app_O: (c:?; t:?) (app c (0) t) = t.
       XElim c; XAuto.
       Qed.