]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma
- strongly normalizing terms form a candidate of reducibility
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_fqus.ma
1 include "basic_2/computation/cpxs_lift.ma".
2 include "basic_2/multiple/fqus_fqus.ma".
3
4 (* Advanced properties ******************************************************)
5
6 lemma lstas_cpxs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
7                   ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ⬈*[h, o] T2.
8 #h #o #G #L #T1 #T2 #d2 #H elim H -G -L -T1 -T2 -d2 //
9 [ /3 width=3 by cpxs_sort, da_inv_sort/
10 | #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
11   elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
12   lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpxs_delta/
13 | #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
14   elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
15   lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
16   #HV1 #H destruct lapply (le_plus_to_le_c … Hd21) -Hd21
17   /3 width=7 by cpxs_delta/
18 | /4 width=3 by cpxs_bind_dx, da_inv_bind/
19 | /4 width=3 by cpxs_flat_dx, da_inv_flat/
20 | /4 width=3 by cpxs_eps, da_inv_flat/
21 ]
22 qed.
23
24 (* Properties on supclosure *************************************************)
25
26 lemma fqu_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 →
27                       ∀T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
28                       ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
29 #h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
30 #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqu_cpx_trans … HT1 … HT2) -T
31 #T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
32 qed-.
33
34 lemma fquq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 →
35                        ∀T1. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
36                        ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
37 #h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fquq_inv_gen … H) -H
38 [ #HT12 elim (fqu_cpxs_trans … HTU2 … HT12) /3 width=3 by fqu_fquq, ex2_intro/
39 | * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
40 ]
41 qed-.
42
43 lemma fquq_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
44                         ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
45                         ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
46                         ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
47 /3 width=5 by fquq_cpxs_trans, lstas_cpxs/ qed-.
48
49 lemma fqup_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 →
50                        ∀T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
51                        ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
52 #h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
53 #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqup_cpx_trans … HT1 … HT2) -T
54 #U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
55 qed-.
56
57 lemma fqus_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 →
58                        ∀T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
59                        ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
60 #h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fqus_inv_gen … H) -H
61 [ #HT12 elim (fqup_cpxs_trans … HTU2 … HT12) /3 width=3 by fqup_fqus, ex2_intro/
62 | * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
63 ]
64 qed-.
65
66 lemma fqus_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
67                         ∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
68                         ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
69                         ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
70 /3 width=6 by fqus_cpxs_trans, lstas_cpxs/ qed-.
71
72 (* Properties on supclosure *************************************************)
73
74 lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
75                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
76                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
77 #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
78 [ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
79   #U2 #HVU2 @(ex3_intro … U2)
80   [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
81   | #H destruct 
82     lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
83   ]
84 | #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
85   [1,3: /2 width=4 by fqu_pair_sn, cpxs_pair_sn/
86   | #H0 destruct /2 width=1 by/
87   ]
88 | #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
89   [1,3: /2 width=4 by fqu_bind_dx, cpxs_bind/
90   | #H0 destruct /2 width=1 by/
91   ]
92 | #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
93   [1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
94   | #H0 destruct /2 width=1 by/
95   ]
96 | #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
97   #U2 #HTU2 @(ex3_intro … U2)
98   [1,3: /2 width=10 by cpxs_lift, fqu_drop/
99   | #H0 destruct /3 width=5 by lift_inj/
100 ]
101 qed-.
102
103 lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
104                            ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
105                            ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
106 #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
107 [ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
108   /3 width=4 by fqu_fquq, ex3_intro/
109 | * #HG #HL #HT destruct /3 width=4 by ex3_intro/
110 ]
111 qed-.
112
113 lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
114                            ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
115                            ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
116 #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
117 [ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
118   /3 width=4 by fqu_fqup, ex3_intro/
119 | #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
120   #U1 #HTU1 #H #H12 elim (fqu_cpxs_trans_neq … H1 … HTU1 H) -T1
121   /3 width=8 by fqup_strap2, ex3_intro/
122 ]
123 qed-.
124
125 lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
126                            ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
127                            ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
128 #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
129 [ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
130   /3 width=4 by fqup_fqus, ex3_intro/
131 | * #HG #HL #HT destruct /3 width=4 by ex3_intro/
132 ]
133 qed-.