]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/lattice.ma
snapshot
[helm.git] / helm / software / matita / dama / lattice.ma
index 9b163378e80e8f6092a33f076123b413b5edadfd..fcad3fdc54dc227e4ea3cbd9c3ac7665db07b831 100644 (file)
@@ -133,6 +133,11 @@ intros (l a b c H); unfold eq in H ⊢ %; unfold Not in H ⊢ %;
 intro H1; apply H; clear H; apply (strong_extj ???? H1);
 qed.
 
+lemma feq_jr: ∀ml:lattice.∀a,b,c:ml. a ≈ b → (a ∨ c) ≈ (b ∨ c).
+intros (l a b c H); apply (Eq≈ ? (join_comm ???)); apply (Eq≈ ?? (join_comm ???));
+apply (feq_jl ???? H);
+qed.
+
 lemma le_to_eqj: ∀ml:lattice.∀a,b:ml. a ≤ b → b ≈ (a ∨ b).
 intros (l a b H); lapply (le_to_eqm ??? H) as H1;
 lapply (feq_jl ??? b H1) as H2;
@@ -141,3 +146,8 @@ apply (Eq≈ (b∨a∧b) ? H2); clear H1 H2 H;
 apply (Eq≈ (b∨(b∧a)) ? (feq_jl ???? (meet_comm ???)));
 apply eq_sym; apply absorbjm;
 qed.
+
+lemma lej: ∀l:lattice.∀x,y:l.x ≤ (x ∨ y).
+intros (l x y); 
+apply (Le≪ ? (absorbmj ? x y)); apply lem;
+qed.
\ No newline at end of file