]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/s_computation/fqup_weight.ma
partial commit in static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / s_computation / fqup_weight.ma
index 94642d01448a7f092a9d4c22afbd332eef1cd608..ed22b9e12635135eff3e0c7f94acc47da3905f38 100644 (file)
@@ -22,7 +22,7 @@ include "static_2/s_computation/fqup.ma".
 lemma fqup_fwd_fw: ∀b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ →
                    ♯❨G2,L2,T2❩ < ♯❨G1,L1,T1❩.
 #b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-/3 width=3 by fqu_fwd_fw, transitive_lt/
+/3 width=3 by fqu_fwd_fw, nlt_trans/
 qed-.
 
 (* Advanced eliminators *****************************************************)
@@ -31,7 +31,7 @@ lemma fqup_wf_ind: ∀b. ∀Q:relation3 …. (
                       ∀G1,L1,T1. (∀G2,L2,T2. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ → Q G2 L2 T2) →
                       Q G1 L1 T1
                    ) → ∀G1,L1,T1. Q G1 L1 T1.
-#b #Q #HQ @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct
+#b #Q #HQ @(wf3_ind_nlt … fw) #x #IHx #G1 #L1 #T1 #H destruct
 /4 width=2 by fqup_fwd_fw/
 qed-.
 
@@ -39,6 +39,6 @@ lemma fqup_wf_ind_eq: ∀b. ∀Q:relation3 …. (
                          ∀G1,L1,T1. (∀G2,L2,T2. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ → Q G2 L2 T2) →
                          ∀G2,L2,T2. G1 = G2 → L1 = L2 → T1 = T2 → Q G2 L2 T2
                       ) → ∀G1,L1,T1. Q G1 L1 T1.
-#b #Q #HQ @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct
+#b #Q #HQ @(wf3_ind_nlt … fw) #x #IHx #G1 #L1 #T1 #H destruct
 /4 width=7 by fqup_fwd_fw/
 qed-.