X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Flattice.ma;h=768c86df1c2af47785bf08b8a762524e72be9426;hb=892992b24f5476c2b4eed13f64e04854ef919020;hp=398b1af89dc0c6c69f7e7dbc3203a0b47773483d;hpb=9791cd146bc0b8df953aee7bb8a3df60553b530c;p=helm.git diff --git a/helm/software/matita/dama/lattice.ma b/helm/software/matita/dama/lattice.ma index 398b1af89..768c86df1 100644 --- a/helm/software/matita/dama/lattice.ma +++ b/helm/software/matita/dama/lattice.ma @@ -14,18 +14,22 @@ set "baseuri" "cic:/matita/lattice/". -include "excedence.ma". +include "excess.ma". record lattice : Type ≝ { l_carr:> apartness; join: l_carr → l_carr → l_carr; meet: l_carr → l_carr → l_carr; + join_refl: ∀x.join x x ≈ x; + meet_refl: ∀x.meet x x ≈ x; join_comm: ∀x,y:l_carr. join x y ≈ join y x; meet_comm: ∀x,y:l_carr. meet x y ≈ meet y x; join_assoc: ∀x,y,z:l_carr. join x (join y z) ≈ join (join x y) z; meet_assoc: ∀x,y,z:l_carr. meet x (meet y z) ≈ meet (meet x y) z; absorbjm: ∀f,g:l_carr. join f (meet f g) ≈ f; - absorbmj: ∀f,g:l_carr. meet f (join f g) ≈ f + absorbmj: ∀f,g:l_carr. meet f (join f g) ≈ f; + strong_extj: ∀x.strong_ext ? (join x); + strong_extm: ∀x.strong_ext ? (meet x) }. interpretation "Lattice meet" 'and a b = @@ -34,60 +38,50 @@ interpretation "Lattice meet" 'and a b = interpretation "Lattice join" 'or a b = (cic:/matita/lattice/join.con _ a b). -(* -include "ordered_set.ma". +definition excl ≝ λl:lattice.λa,b:l.a # (a ∧ b). -record lattice (C:Type) (join,meet:C→C→C) : Prop \def - { (* abelian semigroup properties *) - l_comm_j: symmetric ? join; - l_associative_j: associative ? join; - l_comm_m: symmetric ? meet; - l_associative_m: associative ? meet; - (* other properties *) - l_adsorb_j_m: ∀f,g:C. join f (meet f g) = f; - l_adsorb_m_j: ∀f,g:C. meet f (join f g) = f - }. +lemma excess_of_lattice: lattice → excess. +intro l; apply (mk_excess l (excl l)); +[ intro x; unfold; intro H; unfold in H; apply (ap_coreflexive l x); + apply (ap_rewr ??? (x∧x) (meet_refl l x)); assumption; +| intros 3 (x y z); unfold excl; intro H; + cases (ap_cotransitive ??? (x∧z∧y) H) (H1 H2); [2: + left; apply ap_symmetric; apply (strong_extm ? y); + apply (ap_rewl ???? (meet_comm ???)); + apply (ap_rewr ???? (meet_comm ???)); + assumption] + cases (ap_cotransitive ??? (x∧z) H1) (H2 H3); [left; assumption] + right; apply (strong_extm ? x); apply (ap_rewr ???? (meet_assoc ????)); + assumption] +qed. -record lattice : Type \def - { l_carrier:> Type; - l_join: l_carrier→l_carrier→l_carrier; - l_meet: l_carrier→l_carrier→l_carrier; - l_lattice_properties:> is_lattice ? l_join l_meet - }. +coercion cic:/matita/lattice/excess_of_lattice.con. -definition le \def λL:lattice.λf,g:L. (f ∧ g) = f. +lemma feq_ml: ∀ml:lattice.∀a,b,c:ml. a ≈ b → (c ∧ a) ≈ (c ∧ b). +intros (l a b c H); unfold eq in H ⊢ %; unfold Not in H ⊢ %; +intro H1; apply H; clear H; apply (strong_extm ???? H1); +qed. + +lemma feq_jl: ∀ml:lattice.∀a,b,c:ml. a ≈ b → (c ∨ a) ≈ (c ∨ b). +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. -definition ordered_set_of_lattice: lattice → ordered_set. - intros (L); - apply mk_ordered_set; - [2: apply (le L) - | skip - | apply mk_is_order_relation; - [ unfold reflexive; - intros; - unfold; - rewrite < (l_adsorb_j_m ? ? ? L ? x) in ⊢ (? ? (? ? ? %) ?); - rewrite > l_adsorb_m_j; - [ reflexivity - | apply (l_lattice_properties L) - ] - | intros; - unfold transitive; - unfold le; - intros; - rewrite < H; - rewrite > (l_associative_m ? ? ? L); - rewrite > H1; - reflexivity - | unfold antisimmetric; - unfold le; - intros; - rewrite < H; - rewrite > (l_comm_m ? ? ? L); - assumption - ] - ] +lemma le_to_eqm: ∀ml:lattice.∀a,b:ml. a ≤ b → a ≈ (a ∧ b). +intros (l a b H); + unfold le in H; unfold excess_of_lattice in H; + unfold excl in H; simplify in H; +unfold eq; assumption; qed. -coercion cic:/matita/lattices/ordered_set_of_lattice.con. -*) +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; +apply (Eq≈ ?? (join_comm ???)); +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. + + +