]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/subject.ma
be34cb07feea60d7700d1a7b11d10dfdb8740dfc
[helm.git] / matita / matita / lib / lambda / subject.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department of the University of Bologna, Italy.                     
5     ||I||                                                                 
6     ||T||  
7     ||A||  This file is distributed under the terms of the 
8     \   /  GNU General Public License Version 2        
9      \ /      
10       V_______________________________________________________________ *)
11
12 include "lambda/reduction.ma".
13 include "lambda/inversion.ma". 
14
15 (*
16 inductive T : Type[0] ≝
17   | Sort: nat → T
18   | Rel: nat → T 
19   | App: T → T → T 
20   | Lambda: T → T → T (* type, body *)
21   | Prod: T → T → T (* type, body *)
22   | D: T →T
23
24
25 inductive red : T →T → Prop ≝
26   | rbeta: ∀P,M,N. red (App (Lambda P M) N) (M[0 ≝ N])
27   | rappl: ∀M,M1,N. red M M1 → red (App M N) (App M1 N)
28   | rappr: ∀M,N,N1. red N N1 → red (App M N) (App M N1)
29   | rlaml: ∀M,M1,N. red M M1 → red (Lambda M N) (Lambda M1 N)
30   | rlamr: ∀M,N,N1. red N N1 → red(Lambda M N) (Lambda M N1)
31   | rprodl: ∀M,M1,N. red M M1 → red (Prod M N) (Prod M1 N)
32   | rprodr: ∀M,N,N1. red N N1 → red (Prod M N) (Prod M N1)
33   | d: ∀M,M1. red M M1 → red (D M) (D M1). *)
34   
35 lemma lift_D: ∀M,N. lift M 0 1 = D N → 
36   ∃P. N = lift P 0 1 ∧ M = D P.   
37 #M (cases M) normalize
38   [#i #N #H destruct
39   |#i #N #H destruct
40   |#A #B #N #H destruct
41   |#A #B #N #H destruct
42   |#A #B #N #H destruct
43   |#A #N #H destruct @(ex_intro … A) /2/
44   ]
45 qed. 
46
47 theorem type_of_type: ∀P,G,A,B. G ⊢_{P} A : B → (∀i. B ≠ Sort i) →
48   ∃i. G ⊢_{P} B : Sort i.
49 #Pts #G #A #B #t (elim t)
50   [#i #j #Aij #j @False_ind /2/ 
51   |#G1 #A #i #t1 #_ #P @(ex_intro … i) @(weak … t1 t1)
52   |#G1 #A #B #C #i #t1 #t2 #H1 #H2 #H3 (cases (H1 ?) )
53     [#i #Bi @(ex_intro … i) @(weak … Bi t2)
54     |#i @(not_to_not … (H3 i)) //
55     ]
56   |#G1 #A #B #i #j #k #h #t1 #t2 #_ #_ #H3 @False_ind /2/
57   |#G1 #A #B #C #D #t1 #t2 #H1 #H2 #H3 cases (H1 ?);
58     [#i #t3 cases (prod_inv … t3 … (refl …))
59      #s1 * #s2 * #s3 * * #Ci #H4 #H5 @(ex_intro … s2)
60      @(tj_subst_0 … t2 … H5)
61     |#i % #H destruct
62     ]  
63   |#G1 #A #B #C #i #t1 #t2 #H1 #H2 #H3 /2/
64   |#G1 #A #B #C #i #ch #t1 #t2 #H1 #H2 #H3 /2/
65   |#G1 #A #B #i #t1 #t2 #Hind1 #Hind2 #H /2/
66   ]
67 qed.
68
69 lemma prod_sort : ∀Pts,G,M,P,Q. G ⊢_{Pts} M :Prod P Q →
70  ∃i. P::G ⊢_{Pts} Q : Sort i.
71 #Pts #G #M #P #Q #t cases(type_of_type …t ?);
72   [#s #t2 cases(prod_inv … t2 …(refl …)) #s1 * #s2 * #s3 * * 
73    #_ #_ #H @(ex_intro … s2) //
74   | #i % #H destruct
75   ]
76 qed.
77     
78 axiom red_lift: ∀M,N. red (lift M 0 1) N → 
79   ∃P. N = lift P 0 1 ∧ red M P. 
80
81 theorem tj_d : ∀P,G,M,N. G ⊢_{P} D M : N → G ⊢_{P} M : N.
82 #Pts #G (cut (∀M,N. G ⊢_{Pts} M : N → ∀P. M = D P → G ⊢_{Pts} P : N)) [2: /2/] 
83 #M #N #t (elim t)
84   [#i #j #Aij #P #H destruct 
85   |#G1 #A #i #t1 #_ #P #H destruct
86   |#G1 #A #B #C #i #t1 #t2 #H1 #H2 #P #H3 
87    cases (lift_D … H3) #P * #eqP #eqA >eqP @(weak … i) /2/
88   |#G1 #A #B #i #j #k #h #t1 #t2 #_ #_ #P  #H destruct
89   |#G1 #A #B #C #D #t1 #t2 #_ #_ #P  #H destruct
90   |#G1 #A #B #C #D #t1 #t2 #_ #_ #P  #H destruct
91   |#G1 #A #B #C #i #ch #t1 #t2 #H #_ #P #H
92    @(conv… ch …t2) /2/
93   |#G1 #A #B #i #t1 #t2 #Hind1 #Hind2 #P #H destruct //
94   ]
95 qed.
96
97 definition red0 ≝ λM,N. M = N ∨ red M N.
98
99 axiom conv_lift: ∀P,i,M,N. Co P M N →
100   Co P (lift M 0 i) (lift N 0 i).
101 axiom red_to_conv : ∀P,M,N. red M N → Co P M N.
102 axiom red0_to_conv : ∀P,M,N. red0 M N → Co P M N.
103 axiom conv_prod: ∀P,A,B,M,N. Co P A B → Co P M N → 
104   Co P (Prod A M) (Prod B N). 
105 axiom conv_subst_1: ∀Pts,M,P,Q. red P Q → Co Pts (M[0≝Q]) (M[0≝P]).
106
107 inductive redG : list T → list T → Prop ≝
108  | rnil : redG (nil T) (nil T)
109  | rstep : ∀A,B,G1,G2. red0 A B → redG G1 G2 → 
110      redG (A::G1) (B::G2).
111
112 lemma redG_inv: ∀A,G,G1. redG (A::G) G1 → 
113   ∃B. ∃G2. red0 A B ∧ redG G G2 ∧ G1 = B::G2.
114 #A #G #G1 #rg (inversion rg) 
115   [#H destruct
116   |#A1 #B1 #G2 #G3 #r1 #r2 #_ #H destruct
117    #H1 @(ex_intro … B1) @(ex_intro … G3) % // % //
118   ]
119 qed.
120
121 lemma redG_nil: ∀G. redG (nil T) G → G = nil T.
122 #G #rg (inversion rg) // 
123 #A #B #G1 #G2 #r1 #r2 #_ #H destruct
124 qed.
125
126 axiom conv_prod_split: ∀P,A,A1,B,B1. 
127  Co P(Prod A B) (Prod A1 B1) → Co P A A1 ∧ Co P B B1.
128
129 axiom red0_prod : ∀M,N,P. red0 (Prod M N) P → 
130   (∃Q. P = Prod Q N ∧ red0 M Q) ∨
131   (∃Q. P = Prod M Q ∧ red0 N Q).
132
133 theorem subject_reduction: ∀P,G,M,N. G ⊢_{P} M:N → 
134  ∀M1. red0 M M1 → ∀G1. redG G G1 → G1 ⊢_{P} M1:N.
135 #Pts #G #M #N #d (elim d)
136   [#i #j #Aij #M1 * 
137     [#eqM1 <eqM1 #G1 #H >(redG_nil …H) /2/
138     |#H @False_ind /2/
139     ]
140   |#G1 #A #i #t1 #Hind #M1 * 
141     [#eqM1 <eqM1 #G2 #H cases (redG_inv … H)
142      #P * #G3 * * #r1 #r2 #eqG2 >eqG2
143      @(conv ??? (lift P O 1) ? i);
144       [@conv_lift @sym_conv @red0_to_conv //
145       |@(start … i) @Hind //
146       |@(weak … (Sort i) ? i); [@Hind /2/ | @Hind //]
147       ]
148     |#H @False_ind /2/
149     ]  
150   |#G1 #A #B #C #i #t1 #t2 #H1 #H2 #M1 
151    #H cases H;
152     [#eqM1 <eqM1 #G2 #rg (cases (redG_inv … rg)) 
153      #Q * #G3 * * #r2 #rg1 #eqG2 >eqG2 @(weak … i);
154       [@H1 /2/ | @H2 //] 
155     |#H (elim (red_lift … H)) #P * #eqM1 >eqM1 #redAP
156       #G2 #rg (cases (redG_inv … rg)) #Q * #G3 * * #r2 
157       #rg1 #eqG2 >eqG2 @(weak … i); 
158       [@H1 /2/ | @H2 //]
159     ]
160   |#G #A #B #i #j #k #Rjk #t1 #t2 #Hind1 #Hind2 #M1 #redP
161    (cases (red0_prod … redP))
162     [* #M2 * #eqM1 #redA >eqM1 #G1 #rg @(prod … Rjk);
163       [@Hind1 // | @Hind2 /2/]
164     |* #M2 * #eqM1 #redA >eqM1 #G1 #rg @(prod … Rjk); 
165       [@Hind1 /2/ | @Hind2 /3/] 
166     ]
167   |#G #A #B #C #P #t1 #t2 #Hind1 #Hind2 #M1 #red0a
168    (cases red0a)
169     [#eqM1 <eqM1 #G1 #rg @(app … B);
170       [@Hind1 /2/ | @Hind2 /2/ ]
171     |#reda (cases (red_app …reda))
172       [*
173         [* 
174           #M2 * #N1 * #eqA #eqM1 >eqM1 #G1 #rg
175            cut (G1 ⊢_{Pts} A: Prod B C); [@Hind1 /2/] #H1
176            (cases (abs_inv … H1 … eqA)) #i * #N2 * * 
177            #cProd #t3 #t4 
178            (cut (Co Pts B M2 ∧ Co Pts C N2) ) [/2/] * #convB #convC
179            (cases (prod_inv … t3 … (refl …) )) #i * #j * #k * *
180            #cik #t5 #t6 (cut (G1 ⊢_{Pts} P:B)) 
181             [@Hind2 /2/
182             |#Hcut cut (G1 ⊢_{Pts} N1[0:=P] : N2 [0:=P]);
183               [@(tj_subst_0 … M2) // @(conv … convB Hcut t5)
184               |#Hcut1 cases (prod_sort … H1) #s #Csort
185                @(conv … s  … Hcut1); 
186                 [@conv_subst /2/ | @(tj_subst_0 … Csort) //]
187               ]
188             ]
189         |* #M2 * #eqM1 >eqM1 #H #G1 #rg @(app … B);
190           [@Hind1 /2/ | @Hind2 /2/] 
191         ]      
192       |* #M2 * #eqM1 >eqM1 #H #G1 #rg 
193        cut (G1 ⊢_{Pts} A:Prod B C); [@Hind1 /2/] #t3
194        cases (prod_sort …t3) #i #Csort @(conv ??? C[O≝ M2] … i);
195         [@conv_subst_1 //
196         |@(app … B)  // @Hind2 /2/
197         |@(tj_subst_0 … Csort) @Hind2 /2/
198         ]
199       ]
200     ]
201   |#G #A #B #C #i #t1 #t2 #Hind1 #Hind2 #M2 #red0l #G1 #rg
202    cut (A::G1⊢_{Pts} C:B); [@Hind1 /3/] #t3
203    cut (G1 ⊢_{Pts} Prod A B : Sort i); [@Hind2 /2/] #t4
204    cases red0l;
205     [#eqM2 <eqM2 @(abs … t3 t4)
206     |#redl (cases (red_lambda … redl))  
207       [* #M3 * #eqM2 #redA >eqM2 
208          @(conv ??? (Prod M3 B) … t4);
209           [@conv_prod /2/
210           |@(abs … i); [@Hind1 /3/ |@Hind2 /3/]
211           ]
212         |* #M3 * #eqM3 #redC >eqM3
213          @(abs … t4) @Hind1 /3/
214       ]
215     ]
216   |#G #A #B #C #i #convBC #t1 #t2 #Hind1 #Hind2 #M1 #redA
217    #G1 #rg @(conv … i … convBC); [@Hind1 // |@Hind2 /2/]
218   |#G #A #B #i #t1 #t2 #Hind1 #Hind2 #M1 #red0d #G1 #rg
219    cases red0d; 
220     [#eqM1 <eqM1 @(dummy … i); [@Hind1 /2/ |@Hind2 /2/] 
221     |#redd (cases (red_d … redd)) #Q * #eqM1 #redA >eqM1
222      @(dummy … i);[@Hind1 /2/ |@Hind2 /2/]
223   ]
224 qed.
225