X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fvalued_lattice.ma;h=610bf7d359b65650cb84a838a4f3149bbcb973ca;hb=12ebc9483bec78f948de80e7e230c98488890f4d;hp=6ad88424acfb23df13e82c236730c45bba9956ce;hpb=d6c41ecd6b4d6944becbef2f7f6a625c51d8867c;p=helm.git diff --git a/helm/software/matita/dama/valued_lattice.ma b/helm/software/matita/dama/valued_lattice.ma index 6ad88424a..610bf7d35 100644 --- a/helm/software/matita/dama/valued_lattice.ma +++ b/helm/software/matita/dama/valued_lattice.ma @@ -75,6 +75,8 @@ apply eq_reflexive. qed. lemma modularm: ∀R.∀L:vlattice R.∀y,z:L. μ(y∧z) ≈ μy + μz + -μ (y ∨ z). +(* CSC: questa è la causa per cui la hint per cercare i duplicati ci sta 1 mese *) +(* exact modularj; *) intros (R L y z); lapply (modular_mjp ?? y z) as H1; apply (plus_cancl ??? (μ(y ∨ z))); @@ -90,7 +92,7 @@ qed. lemma modularmj: ∀R.∀L:vlattice R.∀x,y,z:L.μ(x∧(y∨z))≈(μx + μ(y ∨ z) + - μ(x∨(y∨z))). intros (R L x y z); lapply (modular_mjp ?? x (y ∨ z)) as H1; -apply (eq_trans ?? (μ(x∨(y∨z))+ μ(x∧(y∨z)) +-μ(x∨(y∨z)))); [2: apply feq_plusr; apply H1;] clear H1; +apply (eq_trans ?? (μ(x∨(y∨z))+ μ(x∧(y∨z)) +-μ(x∨(y∨z))) ?? (feq_plusr ???? H1)); clear H1; apply (eq_trans ?? ? ?? (plus_comm ???)); apply (eq_trans ?? (- μ(x∨(y∨z))+ μ(x∨(y∨z))+ μ(x∧(y∨z))) ?? (plus_assoc ????)); apply (eq_trans ?? (0+μ(x∧(y∨z))) ?? (opp_inverse ??)); @@ -103,8 +105,8 @@ intros (R L x y z); lapply (modular_mjp ?? x (y ∧ z)) as H1; apply (eq_trans ?? (μ(x∧(y∧z))+ μ(x∨(y∧z)) +-μ(x∧(y∧z)))); [2: apply feq_plusr; apply (eq_trans ???? (plus_comm ???)); apply H1] clear H1; apply (eq_trans ?? ? ?? (plus_comm ???)); -apply (eq_trans ?? (- μ(x∧(y∧z))+ μ(x∧(y∧z))+ μ(x∨y∧z))); [2: apply eq_sym; apply plus_assoc] -apply (eq_trans ?? (0+ μ(x∨y∧z))); [2: apply feq_plusr; apply eq_sym; apply opp_inverse] +apply (eq_trans ?? (- μ(x∧(y∧z))+ μ(x∧(y∧z))+ μ(x∨y∧z)) ?? (plus_assoc ????)); +apply (eq_trans ?? (0+ μ(x∨y∧z)) ?? (opp_inverse ??)); apply eq_sym; apply zero_neutral; qed.