From 0e93f77172427eed198b974e32c7f3e80d2c0251 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Fri, 30 Nov 2007 12:40:56 +0000 Subject: [PATCH] 3.11 !!! --- helm/software/matita/dama/lattice.ma | 31 ++++++++++- helm/software/matita/dama/metric_lattice.ma | 59 ++++++++++++++++++--- helm/software/matita/dama/ordered_group.ma | 10 ++++ helm/software/matita/dama/sequence.ma | 2 + 4 files changed, 93 insertions(+), 9 deletions(-) diff --git a/helm/software/matita/dama/lattice.ma b/helm/software/matita/dama/lattice.ma index 46662b588..9e91376d1 100644 --- a/helm/software/matita/dama/lattice.ma +++ b/helm/software/matita/dama/lattice.ma @@ -55,4 +55,33 @@ intro l; apply (mk_excedence l (excl l)); assumption] qed. -coercion cic:/matita/lattice/excedence_of_lattice.con. \ No newline at end of file +coercion cic:/matita/lattice/excedence_of_lattice.con. + +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. + +lemma le_to_eqm: ∀ml:lattice.∀a,b:ml. a ≤ b → a ≈ (a ∧ b). +intros (l a b H); + unfold le in H; unfold excedence_of_lattice in H; + unfold excl in H; simplify in H; +unfold eq; assumption; +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; +apply (eq_trans ????? (join_comm ???)); +apply (eq_trans ?? (b∨a∧b) ?? H2); clear H1 H2 H; +apply (eq_trans ?? (b∨(b∧a)) ?? (feq_jl ???? (meet_comm ???))); +apply eq_sym; apply absorbjm; +qed. + + + diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index 6028ada59..917e2a199 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -43,31 +43,74 @@ record mlattice (R : ogroup) : Type ≝ { ml_props:> is_mlattice R ml_carr }. -lemma eq_to_zero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0. +lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b. +intros (R ml a b E); intro H; apply E; apply (ml_prop1 ?? ml); +assumption; +qed. + +lemma eq_to_dzero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0. intros (R ml x y H); intro H1; apply H; clear H; apply (ml_prop1 ?? ml); split [apply mpositive] apply ap_symmetric; assumption; qed. -lemma meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y. +(* +lemma lt_to_dpos: ∀R.∀ml:mlattice R.∀x,y:ml.x < y → 0 < δ x y. +intros 4; repeat (unfold in ⊢ (? % ? ?→?)); simplify; unfold excl; +intro H; elim H (H1 H2); elim H2 (H3 H3); [cases (H1 H3)] +split; [apply mpositive] +*) + +lemma meq_l: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y. intros (R ml x y z); apply le_le_eq; [ apply (le_transitive ???? (mtineq ???y z)); - apply (le_rewl ??? (0+δz y) (eq_to_zero ???? H)); + apply (le_rewl ??? (0+δz y) (eq_to_dzero ???? H)); apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive; | apply (le_transitive ???? (mtineq ???y x)); - apply (le_rewl ??? (0+δx y) (eq_to_zero ??z x H)); + apply (le_rewl ??? (0+δx y) (eq_to_dzero ??z x H)); apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;] qed. -lemma meq_joinr: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z. +(* 3.3 *) +lemma meq_r: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z. intros; apply (eq_trans ???? (msymmetric ??y x)); -apply (eq_trans ????? (msymmetric ??z y)); apply meq_joinl; assumption; +apply (eq_trans ????? (msymmetric ??z y)); apply meq_l; assumption; +qed. + +lemma ap_le_to_lt: ∀O:ogroup.∀a,c:O.c # a → c ≤ a → c < a. +intros (R a c A L); split; assumption; qed. +lemma dap_to_lt: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → 0 < δ x y. +intros; split [apply mpositive] apply ap_symmetric; assumption; +qed. + +lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y. +intros (R ml x y H); apply (ml_prop1 ?? ml); split; [apply mpositive;] +apply ap_symmetric; assumption; +qed. + (* 3.11 *) lemma le_mtri: ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z. intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq] apply (le_transitive ????? (ml_prop2 ?? ml (y) ??)); -(* auto type. assert failure *) -whd; + cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [ + apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive] + lapply (le_to_eqm ??? Lxy) as Dxm; + lapply (le_to_eqm ??? Lyz) as Dym; + lapply (le_to_eqj ??? Lxy) as Dxj; + lapply (le_to_eqj ??? Lyz) as Dyj; clear Lxy Lyz; + apply (eq_trans ?? (δ(x∧y) y + δy z)); [apply feq_plusr; apply (meq_l ????? Dxm);] + apply (eq_trans ?? (δ(x∧y) (y∧z) + δy z)); [apply feq_plusr; apply (meq_r ????? Dym);] + apply (eq_trans ?? (δ(x∧y) (y∧z) + δ(x∨y) z)); [apply feq_plusl; apply (meq_l ????? Dxj);] + apply (eq_trans ?? (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [apply feq_plusl; apply (meq_r ????? Dyj);] + apply (eq_trans ?? ? ? (plus_comm ???)); + apply (eq_trans ?? (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z))); [apply feq_plusr; apply (meq_l ????? (join_comm ???));] + apply feq_plusl; + apply (eq_trans ?? (δ(y∧x) (y∧z))); [apply (meq_l ????? (meet_comm ???));] + apply eq_reflexive; +qed. + + + diff --git a/helm/software/matita/dama/ordered_group.ma b/helm/software/matita/dama/ordered_group.ma index 439a6e262..5b0f0aa9b 100644 --- a/helm/software/matita/dama/ordered_group.ma +++ b/helm/software/matita/dama/ordered_group.ma @@ -185,6 +185,16 @@ apply (le_rewr ??? b (zero_neutral ??)); assumption; qed. +lemma le_le0plus: + ∀G:ogroup.∀a,b:G. 0 ≤ a → 0 ≤ b → 0 ≤ a + b. +intros (G a b L1 L2); apply (le_transitive ???? L1); +apply (plus_cancl_le ??? (-a)); +apply (le_rewl ??? 0 (opp_inverse ??)); +apply (le_rewr ??? (-a + a + b) (plus_assoc ????)); +apply (le_rewr ??? (0 + b) (opp_inverse ??)); +apply (le_rewr ??? b (zero_neutral ??)); +assumption; +qed. diff --git a/helm/software/matita/dama/sequence.ma b/helm/software/matita/dama/sequence.ma index 1350545ae..52baef839 100644 --- a/helm/software/matita/dama/sequence.ma +++ b/helm/software/matita/dama/sequence.ma @@ -72,6 +72,8 @@ definition decreasing ≝ λO:pordered_set.λa:sequence O.∀n:nat.a (S n) ≤ a n. + + (* definition is_upper_bound ≝ λO:pordered_set.λa:sequence O.λu:O.∀n:nat.a n ≤ u. -- 2.39.2