1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "logic/equality.ma".
18 include "list/list.ma".
20 inductive list_x : Type ≝
22 | cons_x : ∀T:Type.∀x:T.list_x → list_x.
24 let rec mk_prod (l:list_x) ≝
27 | cons_x T x tl ⇒ ∀y0:T.∀p0:x = y0. mk_prod tl ].
29 let rec appl (l:list_x) : mk_prod l → Type ≝
30 match l return λl:list_x.mk_prod l →Type
32 [ nil_x ⇒ λT:mk_prod nil_x.T
33 | cons_x Ty hd tl ⇒ λT:mk_prod (cons_x Ty hd tl).appl tl (T hd (refl_eq Ty hd)) ].
35 inductive list_xyp : list_x → Type ≝
36 | nil_xyp : ∀l.list_xyp l
37 | cons_xyp : ∀l.∀T:mk_prod l.∀x:appl l T.list_xyp (cons_x ? x l) → list_xyp l.
39 (* let rec tau (l:list_x) (w:list_xyp l) on w: Type ≝
42 | cons_xyp l' T' x' w' ⇒
43 ∀y:appl l' T'.∀p:x' = y.
44 tau (cons_x ? y l') w' ].
49 ∀T1:∀x0:T0. a0=x0 → Type.
50 ∀a1:T1 a0 (refl_eq ? a0).
51 tau ? (cons_xyp nil_x T0 a0 (cons_xyp (cons_x T0 a0 nil_x) T1 a1 (nil_xyp ?))) Type).
53 inductive index_list (T: nat → Type): ∀m,n:nat.Type ≝
54 | il_s : ∀n.T n → index_list T n n
55 | il_c : ∀m,n. T m → index_list T (S m) n → index_list T m n.
57 lemma down_il : ∀T:nat → Type.∀m,n.∀l: index_list T m n.∀f:(∀p. T (S p) → T p).
58 ∀m',n'.S m' = m → S n' = n → index_list T m' n'.
60 [destruct;apply il_s;apply f;assumption
62 [apply f;rewrite > H;assumption
64 [rewrite > H;reflexivity
68 definition R1 ≝ eq_rect'.
73 ∀T1:∀x0:T0. a0=x0 → Type.
74 ∀a1:T1 a0 (refl_eq ? a0).
75 ∀T2:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0. R1 ??? a1 ? p0 = x1 → Type.
76 ∀a2:T2 a0 (refl_eq ? a0) a1 (refl_eq ? a1).
80 ∀e1:R1 ??? a1 ? e0 = b1.
82 intros (T0 a0 T1 a1 T2 a2);
83 apply (eq_rect' ????? e1);
84 apply (R1 ?? ? ?? e0);
91 ∀T1:∀x0:T0. a0=x0 → Type.
92 ∀a1:T1 a0 (refl_eq ? a0).
93 ∀T2:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0. R1 ??? a1 ? p0 = x1 → Type.
94 ∀a2:T2 a0 (refl_eq ? a0) a1 (refl_eq ? a1).
95 ∀T3:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0.∀p1:R1 ??? a1 ? p0 = x1.
96 ∀x2:T2 x0 p0 x1 p1.R2 T0 a0 T1 a1 T2 a2 ? p0 ? p1 = x2→ Type.
97 ∀a3:T3 a0 (refl_eq ? a0) a1 (refl_eq ? a1) a2 (refl_eq ? a2).
101 ∀e1:R1 ??? a1 ? e0 = b1.
103 ∀e2:R2 T0 a0 T1 a1 T2 a2 ? e0 ? e1 = b2.
104 T3 b0 e0 b1 e1 b2 e2.
105 intros (T0 a0 T1 a1 T2 a2 T3 a3);
106 apply (eq_rect' ????? e2);
107 apply (R2 ?? ? ???? e0 ? e1);
111 (*let rec iter_type n (l1 : lista dei nomi fin qui creati) (l2: lista dei tipi dipendenti da applicare) ≝
114 | S p ⇒ match l2 with
115 [ nil ⇒ Type (* dummy *)
116 | cons T tl ⇒ ∀t:app_type_list l1 T.iter_type p (l1@[t]) tl ]].
118 lemma app_type_list : ∀l:list Type.∀T:iter_type l.Type.
122 |simplify in i;apply F;apply i;apply a]
125 definition type_list_cons : ∀l:list Type.iter_type l → list Type ≝
126 λl,T.app_type_list l T :: l.
128 let rec build_dep_type n T acc ≝
132 Type → list Type → Type.
135 inductive II : nat → Type ≝
138 | k3 : ∀n,m. II n → II m → II O.
140 inductive JJ : nat → Type ≝
143 | h3 : ∀n,m. JJ n → JJ m → JJ O.
147 lemma test: h3 ?? h1 h2 = h3 ?? h2 h1 → True.
149 letin f ≝ (λx.S (S x));
150 cut (∃a,b.S a = f b);
151 [decompose;cut (S (S a) = S (f a1))
158 definition IId : ∀n,m.II m → II n → Type ≝
159 λa,b,x,y.match x with
160 [ k1 n ⇒ match y with
161 [ k1 n' ⇒ ∀P:Type.(n = n' → P) → P
163 | k3 m n p' q' ⇒ ∀P:Type.P ]
164 | k2 n ⇒ match y with
166 | k2 n' ⇒ ∀P:Type.(n = n' → P) → P
167 | k3 m' n' p' q' ⇒ ∀P:Type.P ]
168 | k3 m n p q ⇒ match y with
171 | k3 m' n' p' q' ⇒ ∀P:Type.(∀e1:m = m'.∀e2:n = n'. (eq_rect ??? p ? e1) = p'
172 → (eq_rect ??? q ? e2) = q' → P) → P ]].
174 lemma IInc : ∀n,x,y.x = y → IId n n x y.
175 intros;rewrite > H;elim y;simplify;intros;
178 |apply f;reflexivity]
183 lemma IIconflict: ∀c,d.
184 k3 c d (k1 c) (k2 d) = k3 d c (k2 d) (k1 c) → False.
185 intros;apply (IInc ??? H);clear H;intro;
186 apply (eq_rect' ?? (λx.λp.∀e2:x=c.eq_rect ℕ c II (k1 c) x p=k2 x→eq_rect nat x II (k2 x) c e2 = k1 c → False) ?? e1);
187 simplify;intros;apply (IInc ??? H);
189 inductive I1 : nat → Type ≝
192 inductive I2 : ∀n:nat.I1 n → Type ≝
195 inductive I3 : Type ≝
196 | kI3 : ∀x1:nat.∀x2:I1 x1.∀x3:I2 x1 x2.I3.
198 (* 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.
199 λ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.
200 kI3 y1 y2 y3 =kI3 u1 u2 u3)
201 t1 (refl_eq ℕ t1) t2 (refl_eq ((λy1:ℕ.λp1:t1=y1.I1 y1) t1 (refl_eq ℕ t1)) t2)
202 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)
207 definition I3d : ∀x,y:I3.x = y → Type ≝
208 λx,y.match x return (λx:I3.x = y → Type) with
209 [ kI3 x1 x2 x3 ⇒ match y return (λy:I3.kI3 x1 x2 x3 = y → Type) with
211 λe:kI3 x1 x2 x3 = kI3 y1 y2 y3.
212 ∀P:Prop.(∀e1: x1 = y1.
213 ∀e2: R1 ?? (λz1,p1.I1 z1) ?? e1 = y2.
214 ∀e3: R2 ???? (λz1,p1,z2,p2.I2 z1 z2) x3 ? e1 ? e2 = y3.
217 eq ? (kI3 z1 z2 z3) (kI3 y1 y2 y3)) e y1 e1 y2 e2 y3 e3
218 = refl_eq ? (kI3 y1 y2 y3)
221 definition I3d : ∀x,y:I3.x=y → Type.
222 intros 2;cases x;cases y;intro;
223 apply (∀P:Prop.(∀e1: x1 = x3.
224 ∀e2: R1 ?? (λy1,p1.I1 y1) ?? e1 = x4.
225 ∀e3: R2 ???? (λy1,p1,y2,p2.I2 y1 y2) i ? e1 ? e2 = i1.
228 eq ? (kI3 y1 y2 y3) (kI3 x3 x4 i1)) H x3 e1 x4 e2 i1 e3
229 = refl_eq ? (kI3 x3 x4 i1)
233 (* definition I3d : ∀x,y:nat. x = y → Type ≝
236 return (λx.x = y → Type)
238 [ O ⇒ match y return (λy.O = y → Type) with
239 [ O ⇒ λe:O = O.∀P.P → P
240 | S q ⇒ λe: O = S q. ∀P.P]
241 | S p ⇒ match y return (λy.S p = y → Type) with
242 [ O ⇒ λe:S p = O.∀P.P
243 | S q ⇒ λe: S p = S q. ∀P.(p = q → P) → P]].
246 ∀x,y:I3. x = y → Type
250 [ kI3 t1 t2 (t3:I2 t1 t2) ⇒ match y with
251 [ kI3 u1 u2 u3 ⇒ λe:kI3 t1 t2 t3 = kI3 u1 u2 u3.∀P:Type.
253 ∀e2: R1 nat t1 (λy1:nat.λp1:y1 = u1.I1 y1) t2 ? e1 = u2.
254 ∀e3: R2 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3 ? e1 ? e2 = u3.
255 (* ∀e: kI3 t1 t2 t3 = kI3 u1 u2 u3.*)
256 R3 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3
257 (λ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)
263 (∀x,y.match x with [ kI3 t1 t2 t3 ⇒
264 match y with [ kI3 u1 u2 u3 ⇒ kI3 t1 t2 t3 = kI3 u1 u2 u3]]) → Type
267 (∀x,y.match x with [ kI3 t1 t2 t3 ⇒
268 match y with [ kI3 u1 u2 u3 ⇒ kI3 t1 t2 t3 = kI3 u1 u2 u3]]).
270 [ kI3 t1 t2 (t3:I2 t1 t2) ⇒ match y with
271 [ kI3 u1 u2 u3 ⇒ ∀P:Type.
273 ∀e2: R1 ?? (λy1,p1.I1 y1) ?? e1 = u2.
274 ∀e3: R2 ???? (λy1,p1,y2,p2.I2 y1 y2) t3 ? e1 ? e2 = u3.
275 (* ∀e: kI3 t1 t2 t3 = kI3 u1 u2 u3.*)
276 R3 nat t1 (λy1,p1.I1 y1) t2 (λy1,p1,y2,p2.I2 y1 y2) t3
277 (λ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)
281 lemma I3nc : ∀a,b.∀e:a = b. I3d a b e.
282 intros;apply (R1 ????? e);elim a;whd;intros;apply H;reflexivity;
285 (*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.
286 intros;lapply (eq_rect' A x P p y (sym_eq A y x p1));
287 generalize in match Hletin;
288 cut (∀p1:y = x.sym_eq ??? (sym_eq ??? p1) = p1)
289 [rewrite > Hcut;intro;assumption
290 |intro;apply (eq_rect' ????? p2);simplify;reflexivity]
296 ∀T1:∀x0:T0. x0=a0 → Type.
297 ∀a1:T1 a0 (refl_eq ? a0).
298 ∀T2:∀x0:T0. ∀p0:x0=a0. ∀x1:T1 x0 p0. x1 = R1_r ??? a1 ? p0 → Type.
302 ∀e1:b1 = R1_r ??? a1 ? e0.
303 ∀so:T2 a0 (refl_eq ? a0) a1 (refl_eq ? a1).T2 b0 e0 b1 e1.
304 intros 8;intros 2 (e1 H);
305 apply (R1_r ????? e1);
306 apply (R1_r ?? ? ?? e0);
310 definition I3prova : ∀a,b,c,d,e,f.∀Heq:kI3 a b c = kI3 d e f.
311 ∀P:? → ? → ? → ? → Prop.
313 P a b c (refl_eq ??).
314 intros;apply (I3nc ?? Heq);
316 generalize in match H as H;generalize in match Heq as Heq;
317 generalize in match f as f;generalize in match e as e;
319 apply (R1 ????? e1);intros 5;simplify in e2;
320 generalize in match H as H;generalize in match Heq as Heq;
321 generalize in match f as f;
323 apply (R1 ????? e2);intros 4;simplify in e3;
324 generalize in match H as H;generalize in match Heq as Heq;
326 apply (R1 ????? e3);intros;simplify in H1;
332 definition KKd : ∀m,n,p,q.KK m n → KK p q → Type ≝
333 λa,b,c,d,x,y.match x with
334 [ c1 n ⇒ match y with
335 [ c1 n' ⇒ ∀P:Type.(n = n' → P) → P
337 | c3 m' n' p' q' h' i' ⇒ ∀P:Type.P ]
338 | c2 n ⇒ match y with
340 | c2 n' ⇒ ∀P:Type.(n = n' → P) → P
341 | c3 m' n' p' q' h' i' ⇒ ∀P:Type.P ]
342 | c3 m n p q h i ⇒ match y with
345 | c3 m' n' p' q' h' i' ⇒
346 ∀P:Type.(∀e1:m = m'.∀e2:n = n'. ∀e3:p = p'. ∀e4:q = q'.
347 (eq_rect ?? (λm.KK m n') (eq_rect ?? (λn.KK m n) h n' e2) m' e1) = h'
348 → (eq_rect ?? (λp.KK p q') (eq_rect ?? (λq.KK p q) i ? e4) ? e3) = i' → P) → P ]].
350 lemma KKnc : ∀a,b,e,f.e = f → KKd a b a b e f.
351 intros;rewrite > H;elim f;simplify;intros;apply f1;reflexivity;