]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma
some renaming and some typos corrected ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lsubr_lsubr.ma
index 937f6f78c30893dd323b3cdbaea9ca9a3b4d667f..f0e171e0b2198c2564fbcfe41061d6fe2499a12e 100644 (file)
@@ -18,7 +18,7 @@ include "basic_2/static/lsubr.ma".
 
 (* Auxiliary inversion lemmas ***********************************************)
 
-fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
+fact lsubr_inv_pair1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
                           ∨∨ L2 = ⋆
                            | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
                            | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
@@ -30,12 +30,12 @@ fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
 ]
 qed-.
 
-lemma lsubr_inv_bind1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 →
+lemma lsubr_inv_pair1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 →
                        ∨∨ L2 = ⋆
                         | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
                         | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
                                     I = Abbr & X = ⓝW.V.
-/2 width=3 by lsubr_inv_bind1_aux/ qed-.
+/2 width=3 by lsubr_inv_pair1_aux/ qed-.
 
 (* Main properties **********************************************************)
 
@@ -44,10 +44,10 @@ theorem lsubr_trans: Transitive … lsubr.
 [ #L1 #X #H
   lapply (lsubr_inv_atom1 … H) -H //
 | #I #L1 #L #V #_ #IHL1 #X #H
-  elim (lsubr_inv_bind1 … H) -H // *
-  #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_bind, lsubr_abst/
+  elim (lsubr_inv_pair1 … H) -H // *
+  #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_pair, lsubr_beta/
 | #L1 #L #V1 #W #_ #IHL1 #X #H
   elim (lsubr_inv_abst1 … H) -H // *
-  #L2 #HL2 #H destruct /3 width=1 by lsubr_abst/
+  #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/
 ]
 qed-.