]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/coq-contribs/LAMBDA-TYPES/pc3_gen.v
contribution about \lambda-\delta
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / pc3_gen.v
diff --git a/helm/coq-contribs/LAMBDA-TYPES/pc3_gen.v b/helm/coq-contribs/LAMBDA-TYPES/pc3_gen.v
new file mode 100644 (file)
index 0000000..4e5daba
--- /dev/null
@@ -0,0 +1,80 @@
+(*#* #stop file *)
+
+Require lift_gen.
+Require pr3_props.
+Require pr3_gen.
+Require pc3_defs.
+Require pc3_props.
+
+   Section pc3_gen. (********************************************************)
+
+      Theorem pc3_gen_abst : (c:?; u1,u2,t1,t2:?)
+                             (pc3 c (TTail (Bind Abst) u1 t1)
+                                    (TTail (Bind Abst) u2 t2)
+                             ) ->
+                             (pc3 c u1 u2) /\
+                             (b:?; u:?) (pc3 (CTail c (Bind b) u) t1 t2).
+      Intros.
+      Pc3Confluence; Pr3Gen; Pr3Gen; Rewrite H0 in H; Clear H0 x.
+      Inversion H; Rewrite H5 in H1; Rewrite H6 in H2.
+      Split; XEAuto.
+      Qed.
+
+      Theorem pc3_gen_lift: (c:?; t1,t2:?; h,d:?)
+                            (pc3 c (lift h d t1) (lift h d t2)) ->
+                            (e:?) (drop h d c e) ->
+                            (pc3 e t1 t2).
+      Intros.
+      Pc3Confluence; Pr3Gen; Pr3Gen; Rewrite H1 in H; Clear H1 x.
+      LiftGen; Rewrite H in H2; XEAuto.
+      Qed.
+
+      Theorem pc3_gen_not_abst : (b:?) ~b=Abst -> (c:?; t1,t2,u1,u2:?)
+                                 (pc3 c (TTail (Bind b) u1 t1)
+                                        (TTail (Bind Abst) u2 t2)
+                                 ) ->
+                                 (pc3 (CTail c (Bind b) u1) t1
+                                      (lift (1) (0) (TTail (Bind Abst) u2 t2))
+                                 )
+             .
+      Intros b; XElim b; Intros;
+      Try EqFalse; Pc3Confluence; Pr3Gen; Pr3Gen;
+      Try (Rewrite H1 in H0; Inversion H0);
+      Rewrite H1 in H4; Pr3Context;
+      EApply pc3_pr3_t; XEAuto.
+      Qed.
+
+      Theorem pc3_gen_lift_abst : (c:?; t,t2,u2:?; h,d:?)
+                                  (pc3 c (lift h d t)
+                                         (TTail (Bind Abst) u2 t2)
+                                  ) ->
+                                  (e:?) (drop h d c e) ->
+                                  (EX u1 t1 | (pr3 e t (TTail (Bind Abst) u1 t1)) &
+                                              (pr3 c u2 (lift h d u1)) &
+                                              (b:B; u:T)(pr3 (CTail c (Bind b) u) t2 (lift h (S d) t1))
+                                  ).
+      Intros.
+      Pc3Confluence; Pr3Gen; Pr3Gen; Rewrite H1 in H; Clear H1 x.
+      LiftGenBase; Rewrite H in H4; Rewrite H1 in H2; Rewrite H5 in H3; XEAuto.
+      Qed.
+
+   End pc3_gen.
+
+      Tactic Definition Pc3Gen :=
+         Match Context With
+            | [ _: (pc3 ?1 (lift ?2 ?3 ?4) (lift ?2 ?3 ?5));
+                _: (drop ?2 ?3 ?1 ?6) |- ? ] ->
+               LApply (pc3_gen_lift ?1 ?4 ?5 ?2 ?3); [ Intros H_x | XAuto ];
+               LApply (H_x ?6); [ Clear H_x; Intros | XAuto ]
+            | [ H: (pc3 ?1 (TTail (Bind Abst) ?2 ?3) (TTail (Bind Abst) ?4 ?5)) |- ? ] ->
+               LApply (pc3_gen_abst ?1 ?2 ?4 ?3 ?5);[ Clear H; Intros H | XAuto ];
+               XElim H; Intros
+            | [ H: (pc3 ?1 (TTail (Bind ?2) ?3 ?4) (TTail (Bind Abst) ?5 ?6));
+                _: ~ ?2 = Abst |- ? ] ->
+               LApply (pc3_gen_not_abst ?2); [ Intros H_x | XAuto ];
+               LApply (H_x ?1 ?4 ?6 ?3 ?5); [ Clear H H_x; Intros | XAuto ]
+            | [ _: (pc3 ?1 (lift ?2 ?3 ?4) (TTail (Bind Abst) ?5 ?6));
+                _: (drop ?2 ?3 ?1 ?7) |- ? ] ->
+               LApply (pc3_gen_lift_abst ?1 ?4 ?6 ?5 ?2 ?3); [ Intros H_x | XAuto ];
+               LApply (H_x ?7); [ Clear H_x; Intros H_x | XAuto ];
+               XElim H_x; Intros.