]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/coq-contribs/LAMBDA-TYPES/pr0_subst0.v
contribution about \lambda-\delta
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / pr0_subst0.v
diff --git a/helm/coq-contribs/LAMBDA-TYPES/pr0_subst0.v b/helm/coq-contribs/LAMBDA-TYPES/pr0_subst0.v
new file mode 100644 (file)
index 0000000..ace613d
--- /dev/null
@@ -0,0 +1,105 @@
+Require subst0_gen.
+Require subst0_lift.
+Require subst0_subst0.
+Require subst0_confluence.
+Require pr0_defs.
+Require pr0_lift.
+
+(*#* #stop file *)
+
+   Section pr0_subst0. (*****************************************************)
+
+      Tactic Definition IH :=
+         Match Context With
+            | [ H1: (u1:?) (pr0 u1 ?1) -> ?; H2: (pr0 ?2 ?1) |- ? ] ->
+               LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ];
+               XElim H1; Intros
+            | [ H1: (u1:?) (pr0 ?1 u1) -> ?; H2: (pr0 ?1 ?2) |- ? ] ->
+               LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ];
+               XElim H1; Intros
+            | [ H1: (v1,w1:?; i:?) (subst0 i v1 ?1 w1) -> (v2:T) (pr0 v1 v2) -> ?;
+                H2: (subst0 ?2 ?3 ?1 ?4); H3: (pr0 ?3 ?5) |- ? ] ->
+               LApply (H1 ?3 ?4 ?2); [ Clear H1; Intros H1 | XAuto ];
+               LApply (H1 ?5); [ Clear H1; Intros H1 | XAuto ];
+               XElim H1; [ Intros | Intros H1; XElim H1; Intros ].
+
+      Theorem pr0_subst0_back: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) ->
+                               (u1:?) (pr0 u1 u2) ->
+                               (EX t | (subst0 i u1 t1 t) & (pr0 t t2)).
+      Intros until 1; XElim H; Intros;
+      Repeat IH; XEAuto.
+      Qed.
+
+      Theorem pr0_subst0_fwd: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) ->
+                              (u1:?) (pr0 u2 u1) ->
+                              (EX t | (subst0 i u1 t1 t) & (pr0 t2 t)).
+      Intros until 1; XElim H; Intros;
+      Repeat IH; XEAuto.
+      Qed.
+
+      Hints Resolve pr0_subst0_fwd : ltlc.
+
+(*#* #start file *)
+
+(*#* #caption "confluence with substitution" *)
+(*#* #cap #cap t1,t2,v1,v2 #alpha w1 in U1, w2 in U2 *)
+
+      Theorem pr0_subst0: (t1,t2:?) (pr0 t1 t2) ->
+                          (v1,w1:?; i:?) (subst0 i v1 t1 w1) ->
+                          (v2:?) (pr0 v1 v2) ->
+                          (pr0 w1 t2) \/
+                          (EX w2 | (pr0 w1 w2) & (subst0 i v2 t2 w2)).
+
+(*#* #stop file *)
+
+      Intros until 1; XElim H; Clear t1 t2; Intros.
+(* case 1: pr0_refl *)
+      XEAuto.
+(* case 2: pr0_cong *)
+      Subst0Gen; Rewrite H3; Repeat IH; XEAuto.
+(* case 3: pr0_beta *)
+      Repeat Subst0Gen; Rewrite H3; Try Rewrite H5; Try Rewrite H6;
+      Repeat IH; XEAuto.
+(* case 4: pr0_gamma *)
+      Repeat Subst0Gen; Rewrite H6; Try Rewrite H8; Try Rewrite H9;
+      Repeat IH; XDEAuto 7.
+(* case 5: pr0_delta *)
+      Subst0Gen; Rewrite H4; Repeat IH;
+      [ XEAuto | Idtac | XEAuto | Idtac | XEAuto | Idtac | Idtac | Idtac ].
+      Subst0Subst0; Arith9'In H9 i; XEAuto.
+      Subst0Confluence; XEAuto.
+      Subst0Subst0; Arith9'In H10 i; XEAuto.
+      Subst0Confluence; XEAuto.
+      Subst0Subst0; Arith9'In H11 i; Subst0Confluence; XDEAuto 6.
+(* case 6: pr0_zeta *)
+      Repeat Subst0Gen; Rewrite H2; Try Rewrite H4; Try Rewrite H5;
+      Try (Simpl in H5; Rewrite <- minus_n_O in H5);
+      Try (Simpl in H6; Rewrite <- minus_n_O in H6);
+      Try IH; XEAuto.
+(* case 7: pr0_epsilon *)
+      Subst0Gen; Rewrite H1; Try IH; XEAuto.
+      Qed.
+
+   End pr0_subst0.
+
+      Tactic Definition Pr0Subst0 :=
+         Match Context With
+            | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5);
+                H3: (pr0 ?4 ?6) |- ? ] ->
+               LApply (pr0_subst0 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
+               LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
+               LApply (H1 ?6); [ Clear H1; Intros H1 | XAuto ];
+               XElim H1; [ Intros | Intros H1; XElim H1; Intros ]
+            | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5) |- ? ] ->
+               LApply (pr0_subst0 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
+               LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
+               LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ];
+               XElim H1; [ Intros | Intros H1; XElim H1; Intros ]
+            | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (pr0 ?4 ?1) |- ? ] ->
+               LApply (pr0_subst0_back ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ];
+               LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ];
+               XElim H1; Intros
+            | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (pr0 ?1 ?4) |- ? ] ->
+               LApply (pr0_subst0_fwd ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ];
+               LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ];
+               XElim H1; Intros.