]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/relocation/sex_sex.ma
some restyling ...
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / relocation / sex_sex.ma
index 571379e2f49b00b576659753838654ce5a5dfee0..c1954b1a789f00b113efbbd08973972c36a51afd 100644 (file)
@@ -23,9 +23,9 @@ theorem sex_trans_gen (RN1) (RP1) (RN2) (RP2) (RN) (RP):
                       ∀L1,f.
                       (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → sex_transitive RN1 RN2 RN RN1 RP1 g K I) →
                       (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ⫯g = ⫱*[n] f → sex_transitive RP1 RP2 RP RN1 RP1 g K I) →
-                      ∀L0. L1 ⪤[RN1, RP1, f] L0 →
-                      ∀L2. L0 ⪤[RN2, RP2, f] L2 →
-                      L1 ⪤[RN, RP, f] L2.
+                      ∀L0. L1 ⪤[RN1,RP1,f] L0 →
+                      ∀L2. L0 ⪤[RN2,RP2,f] L2 →
+                      L1 ⪤[RN,RP,f] L2.
 #RN1 #RP1 #RN2 #RP2 #RN #RP #L1 elim L1 -L1
 [ #f #_ #_ #L0 #H1 #L2 #H2
   lapply (sex_inv_atom1 … H1) -H1 #H destruct
@@ -50,8 +50,8 @@ theorem sex_trans (RN) (RP) (f): (∀g,I,K. sex_transitive RN RN RN RN RP g K I)
                                  Transitive … (sex RN RP f).
 /2 width=9 by sex_trans_gen/ qed-.
 
-theorem sex_trans_id_cfull: ∀R1,R2,R3,L1,L,f. L1 ⪤[R1, cfull, f] L → 𝐈⦃f⦄ →
-                            ∀L2. L ⪤[R2, cfull, f] L2 → L1 ⪤[R3, cfull, f] L2.
+theorem sex_trans_id_cfull: ∀R1,R2,R3,L1,L,f. L1 ⪤[R1,cfull,f] L → 𝐈⦃f⦄ →
+                            ∀L2. L ⪤[R2,cfull,f] L2 → L1 ⪤[R3,cfull,f] L2.
 #R1 #R2 #R3 #L1 #L #f #H elim H -L1 -L -f
 [ #f #Hf #L2 #H >(sex_inv_atom1 … H) -L2 // ]
 #f #I1 #I #K1 #K #HK1 #_ #IH #Hf #L2 #H
@@ -93,9 +93,9 @@ theorem sex_canc_dx: ∀RN,RP,f. Transitive … (sex RN RP f) →
 /3 width=3 by/ qed-.
 
 lemma sex_meet: ∀RN,RP,L1,L2.
-                ∀f1. L1 ⪤[RN, RP, f1] L2 →
-                ∀f2. L1 ⪤[RN, RP, f2] L2 →
-                ∀f. f1 ⋒ f2 ≘ f → L1 ⪤[RN, RP, f] L2.
+                ∀f1. L1 ⪤[RN,RP,f1] L2 →
+                ∀f2. L1 ⪤[RN,RP,f2] L2 →
+                ∀f. f1 ⋒ f2 ≘ f → L1 ⪤[RN,RP,f] L2.
 #RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
 #f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf
 elim (pn_split f2) * #g2 #H2 destruct
@@ -106,9 +106,9 @@ try elim (sex_inv_push … H) try elim (sex_inv_next … H) -H
 qed-.
 
 lemma sex_join: ∀RN,RP,L1,L2.
-                ∀f1. L1 ⪤[RN, RP, f1] L2 →
-                ∀f2. L1 ⪤[RN, RP, f2] L2 →
-                ∀f. f1 ⋓ f2 ≘ f → L1 ⪤[RN, RP, f] L2.
+                ∀f1. L1 ⪤[RN,RP,f1] L2 →
+                ∀f2. L1 ⪤[RN,RP,f2] L2 →
+                ∀f. f1 ⋓ f2 ≘ f → L1 ⪤[RN,RP,f] L2.
 #RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
 #f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf
 elim (pn_split f2) * #g2 #H2 destruct