]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/tests/destruct_bb.ma
Implementation of ndestruct tactic (including destruction of constructor forms
[helm.git] / helm / software / matita / tests / destruct_bb.ma
index 494d5805a286f72977ac220c98b2eeb28cb721e6..382620e34f23fa458a31029354c79039fca61833 100644 (file)
@@ -79,7 +79,7 @@ definition R2 :
   ∀b1: T1 b0 e0.
   ∀e1:R1 ??? a1 ? e0 = b1.
   T2 b0 e0 b1 e1.
-intros 9;intro e1;
+intros (T0 a0 T1 a1 T2 a2);
 apply (eq_rect' ????? e1);
 apply (R1 ?? ? ?? e0);
 simplify;assumption;
@@ -94,14 +94,15 @@ definition R3 :
   ∀a2:T2 a0 (refl_eq ? a0) a1 (refl_eq ? a1).
   ∀T3:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0.∀p1:R1 ??? a1 ? p0 = x1.
       ∀x2:T2 x0 p0 x1 p1.R2 T0 a0 T1 a1 T2 a2 ? p0 ? p1 = x2→ Type.
+  ∀a3:T3 a0 (refl_eq ? a0) a1 (refl_eq ? a1) a2 (refl_eq ? a2).
   ∀b0:T0.
   ∀e0:a0 = b0.
   ∀b1: T1 b0 e0.
   ∀e1:R1 ??? a1 ? e0 = b1.
   ∀b2: T2 b0 e0 b1 e1.
   ∀e2:R2 T0 a0 T1 a1 T2 a2 ? e0 ? e1 = b2.
-  ∀a3:T3 a0 (refl_eq ? a0) a1 (refl_eq ? a1) a2 (refl_eq ? a2).T3 b0 e0 b1 e1 b2 e2.      
-intros 12;intros 2 (e2 H);
+  T3 b0 e0 b1 e1 b2 e2.       
+intros (T0 a0 T1 a1 T2 a2 T3 a3);
 apply (eq_rect' ????? e2);
 apply (R2 ?? ? ???? e0 ? e1);
 simplify;assumption;
@@ -194,16 +195,91 @@ inductive I2 : ∀n:nat.I1 n → Type ≝
 inductive I3 : Type ≝
 | kI3 : ∀x1:nat.∀x2:I1 x1.∀x3:I2 x1 x2.I3.
 
-definition I3d: I3 → I3 → Type ≝ 
-λx,y.match x with
+(* lemma idfof : (∀t1,t2,t3,u1,u2,u3.((λy1:ℕ.λp1:t1=y1.λy2:I1 y1.λp2:R1 ℕ t1 (λy1:ℕ.λp1:t1=y1.I1 y1) t2 y1 p1 =y2.
+       λy3:I2 y1 y2.λp3:R2 ℕ t1 (λy1:ℕ.λp1:t1=y1.I1 y1) t2 (λy1:ℕ.λp1:t1=y1.λy2:I1 y1.λp2:R1 ℕ t1 (λy1:ℕ.λp1:t1 =y1.I1 y1) t2 y1 p1 =y2.I2 y1 y2) t3 y1 p1 y2 p2 =y3.
+                    kI3 y1 y2 y3 =kI3 u1 u2 u3)
+t1 (refl_eq ℕ t1) t2 (refl_eq ((λy1:ℕ.λp1:t1=y1.I1 y1) t1 (refl_eq ℕ t1)) t2)
+   t3 (refl_eq ((λy1:ℕ.λp1:t1=y1.λy2:I1 y1.λp2:R1 ℕ t1 (λy1:ℕ.λp1:t1=y1.I1 y1) t2 y1 p1 =y2.I2 y1 y2) t1 (refl_eq ℕ t1) t2 (refl_eq ((λy1:ℕ.λp1:t1=y1.I1 y1) t1 (refl_eq ℕ t1)) t2)) t3)
+   )
+    → True).
+simplify; *)
+
+definition I3d : ∀x,y:I3.x = y → Type ≝
+λx,y.match x return (λx:I3.x = y → Type) with
+[ kI3 x1 x2 x3 ⇒ match y return (λy:I3.kI3 x1 x2 x3 = y → Type) with
+  [ kI3 y1 y2 y3 ⇒ 
+    λe:kI3 x1 x2 x3 = kI3 y1 y2 y3.
+       ∀P:Prop.(∀e1: x1 = y1.
+                ∀e2: R1 ?? (λz1,p1.I1 z1) ?? e1 = y2.
+                ∀e3: R2 ???? (λz1,p1,z2,p2.I2 z1 z2) x3 ? e1 ? e2 = y3. 
+                R3 ?????? 
+                (λz1,p1,z2,p2,z3,p3.
+                  eq ? (kI3 z1 z2 z3) (kI3 y1 y2 y3)) e y1 e1 y2 e2 y3 e3
+                = refl_eq ? (kI3 y1 y2 y3)
+                → P) → P]].
+
+definition I3d : ∀x,y:I3.x=y → Type.
+intros 2;cases x;cases y;intro;
+apply (∀P:Prop.(∀e1: x1 = x3.
+                ∀e2: R1 ?? (λy1,p1.I1 y1) ?? e1 = x4.
+                ∀e3: R2 ???? (λy1,p1,y2,p2.I2 y1 y2) i ? e1 ? e2 = i1. 
+                R3 ?????? 
+                (λy1,p1,y2,p2,y3,p3.
+                  eq ? (kI3 y1 y2 y3) (kI3 x3 x4 i1)) H x3 e1 x4 e2 i1 e3
+                = refl_eq ? (kI3 x3 x4 i1)
+                → P) → P);
+qed.
+
+(* definition I3d : ∀x,y:nat. x = y → Type ≝
+λx,y.
+match x 
+ return (λx.x = y → Type)
+ with
+[ O ⇒ match y return (λy.O = y → Type) with
+  [ O ⇒ λe:O = O.∀P.P → P
+  | S q ⇒ λe: O = S q. ∀P.P]
+| S p ⇒ match y return (λy.S p = y → Type) with
+  [ O ⇒ λe:S p = O.∀P.P
+  | S q ⇒ λe: S p = S q. ∀P.(p = q → P) → P]]. 
+
+definition I3d: 
+  ∀x,y:I3. x = y → Type 
+  ≝ 
+λx,y. 
+match x with
+[ kI3 t1 t2 (t3:I2 t1 t2) ⇒ match y with
+  [ kI3 u1 u2 u3 ⇒ λe:kI3 t1 t2 t3 = kI3 u1 u2 u3.∀P:Type.
+    (∀e1: t1 = u1.
+     ∀e2: R1 nat t1 (λy1:nat.λp1:y1 = u1.I1 y1) t2 ? e1 = u2.
+     ∀e3: R2 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3 ? e1 ? e2 = u3. 
+     (* ∀e:  kI3 t1 t2 t3 = kI3 u1 u2 u3.*)
+     R3 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3 
+        (λy1,p1,y2,p2,y3,p3.eq I3 (kI3 y1 y2 y3) (kI3 u1 u2 u3)) e u1 e1 u2 e2 u3 e3 = refl_eq I3 (kI3 u1 u2 u3)
+     → P)
+    → P]].
+
+definition I3d: 
+  ∀x,y:I3.
+    (∀x,y.match x with [ kI3 t1 t2 t3 ⇒
+      match y with [ kI3 u1 u2 u3 ⇒ kI3 t1 t2 t3 = kI3 u1 u2 u3]]) → Type 
+  ≝ 
+λx,y.λe:
+   (∀x,y.match x with [ kI3 t1 t2 t3 ⇒
+      match y with [ kI3 u1 u2 u3 ⇒ kI3 t1 t2 t3 = kI3 u1 u2 u3]]). 
+match x with
 [ kI3 t1 t2 (t3:I2 t1 t2) ⇒ match y with
   [ kI3 u1 u2 u3 ⇒ ∀P:Type.
     (∀e1: t1 = u1.
      ∀e2: R1 ?? (λy1,p1.I1 y1) ?? e1 = u2.
-     ∀e3: R2 ???? (λy1,p1,y2,p2.I2 y1 y2) ? e1 ? e2 t3 = u3.P) → P]].
-
-lemma I3nc : ∀a,b.a = b → I3d a b.
-intros;rewrite > H;elim b;simplify;intros;apply f;reflexivity;
+     ∀e3: R2 ???? (λy1,p1,y2,p2.I2 y1 y2) t3 ? e1 ? e2 = u3. 
+     (* ∀e:  kI3 t1 t2 t3 = kI3 u1 u2 u3.*)
+     R3 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3 
+        (λy1,p1,y2,p2,y3,p3.eq I3 (kI3 y1 y2 y3) (kI3 u1 u2 u3)) (e (kI3 t1 t2 t3) (kI3 u1 u2 u3)) u1 e1 u2 e2 u3 e3 = refl_eq ? (kI3 u1 u2 u3)
+     → P)
+    → P]].*)
+     
+lemma I3nc : ∀a,b.∀e:a = b. I3d a b e.
+intros;apply (R1 ????? e);elim a;whd;intros;apply H;reflexivity;
 qed.
 
 (*lemma R1_r : ΠA:Type.Πx:A.ΠP:Πy:A.y=x→Type.P x (refl_eq A x)→Πy:A.Πp:y=x.P y p.
@@ -231,17 +307,25 @@ apply (R1_r ?? ? ?? e0);
 simplify;assumption;
 qed.*)
 
-definition I3prova : ∀a,b,c,d,e,f.kI3 a b c = kI3 d e f → ∃P.P d e f.
-intros;apply (I3nc ?? H);clear H;
+definition I3prova : ∀a,b,c,d,e,f.∀Heq:kI3 a b c = kI3 d e f.
+  ∀P:? → ? → ? → ? → Prop.
+  P d e f Heq → 
+  P a b c (refl_eq ??).
+intros;apply (I3nc ?? Heq);
 simplify;intro;
+generalize in match H as H;generalize in match Heq as Heq;
 generalize in match f as f;generalize in match e as e;
-generalize in match c as c;generalize in match b as b;
-clear f e c b;
+clear H Heq f e;
 apply (R1 ????? e1);intros 5;simplify in e2;
-generalize in match f as f;generalize in match c as c;
-clear f c;
-apply (R1 ????? e2);intros;simplify in H;
-elim daemon;
+generalize in match H as H;generalize in match Heq as Heq;
+generalize in match f as f;
+clear H Heq f;
+apply (R1 ????? e2);intros 4;simplify in e3;
+generalize in match H as H;generalize in match Heq as Heq;
+clear H Heq;
+apply (R1 ????? e3);intros;simplify in H1;
+apply (R1 ????? H1);
+assumption;
 qed.