From: Enrico Tassi Date: Tue, 4 Dec 2007 18:40:03 +0000 (+0000) Subject: refactoring of some lemmas, shorter proofs X-Git-Tag: make_still_working~5729 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=db3104b6a36b957d7be4c79e50f22fd0ffd90f0d;p=helm.git refactoring of some lemmas, shorter proofs --- diff --git a/helm/software/matita/dama/divisible_group.ma b/helm/software/matita/dama/divisible_group.ma index 7d2eeed45..44db35a15 100644 --- a/helm/software/matita/dama/divisible_group.ma +++ b/helm/software/matita/dama/divisible_group.ma @@ -41,7 +41,7 @@ intro G; cases G; unfold divide; intros (x n); simplify; cases (f x n); simplify; exact H; qed. -lemma feq_pow: ∀G:dgroup.∀x,y:G.∀n.x≈y → n * x ≈ n * y. +lemma feq_mul: ∀G:dgroup.∀x,y:G.∀n.x≈y → n * x ≈ n * y. intros (G x y n H); elim n; [apply eq_reflexive] simplify; apply (Eq≈ (x + (n1 * y)) H1); apply (Eq≈ (y+n1*y) H (eq_reflexive ??)); @@ -54,7 +54,7 @@ apply (ap_rewl ???? (plus_comm ???)); apply (ap_rewl ???w (zero_neutral ??)); assumption; qed. -lemma appow_ap: ∀G:dgroup.∀x,y:G.∀n.S n * x # S n * y → x # y. +lemma apmul_ap: ∀G:dgroup.∀x,y:G.∀n.S n * x # S n * y → x # y. intros 4 (G x y n); elim n; [2: simplify in a; cases (applus ????? a); [assumption] @@ -62,7 +62,7 @@ intros 4 (G x y n); elim n; [2: apply (plus_cancr_ap ??? 0); assumption; qed. -lemma pluspow: ∀G:dgroup.∀x,y:G.∀n.n * (x+y) ≈ n * x + n * y. +lemma plusmul: ∀G:dgroup.∀x,y:G.∀n.n * (x+y) ≈ n * x + n * y. intros (G x y n); elim n; [ simplify; apply (Eq≈ 0 ? (zero_neutral ? 0)); apply eq_reflexive] simplify; apply eq_sym; apply (Eq≈ (x+y+(n1*x+n1*y))); [ @@ -78,7 +78,7 @@ simplify; apply eq_sym; apply (Eq≈ (x+y+(n1*x+n1*y))); [ apply feq_plusl; apply eq_sym; assumption; qed. -lemma powzero: ∀G:dgroup.∀n.n*0 ≈ 0. [2: apply dg_carr; apply G] +lemma mulzero: ∀G:dgroup.∀n.n*0 ≈ 0. [2: apply dg_carr; apply G] intros; elim n; [simplify; apply eq_reflexive] simplify; apply (Eq≈ ? (zero_neutral ??)); assumption; qed. diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index cca643098..b5261b9dc 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -92,7 +92,8 @@ lapply (le_to_eqj ??? Lxy) as Dxj; lapply (le_to_eqj ??? Lyz) as Dyj; clear Lxy apply (Eq≈ (δ(x∧y) y + δy z) (meq_l ????? Dxm)); apply (Eq≈ (δ(x∧y) (y∧z) + δy z) (meq_r ????? Dym)); apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) z) (meq_l ????? Dxj)); -apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z)) (meq_r ????? Dyj)); +apply (Eq≈ (δ(x∧y) (y∧z) + δ(x∨y) (y∨z))); [ + apply (feq_plusl ? (δ(x∧y) (y∧z)) ?? (meq_r ??? (x∨y) ? Dyj));] apply (Eq≈ ? (plus_comm ???)); apply (Eq≈ (δ(y∨x) (y∨z)+ δ(x∧y) (y∧z)) (meq_l ????? (join_comm ?x y))); apply feq_plusl; diff --git a/helm/software/matita/dama/ordered_divisible_group.ma b/helm/software/matita/dama/ordered_divisible_group.ma index cae552114..a6c13e189 100644 --- a/helm/software/matita/dama/ordered_divisible_group.ma +++ b/helm/software/matita/dama/ordered_divisible_group.ma @@ -32,19 +32,19 @@ qed. coercion cic:/matita/ordered_divisible_group/todg_division.con. -lemma pow_ge: ∀G:todgroup.∀x:G.∀n.0 ≤ x → 0 ≤ n * x. +lemma mul_ge: ∀G:todgroup.∀x:G.∀n.0 ≤ x → 0 ≤ n * x. intros (G x n); elim n; simplify; [apply le_reflexive] apply (le_transitive ???? H1); apply (le_rewl ??? (0+(n1*x)) (zero_neutral ??)); apply fle_plusr; assumption; qed. -lemma lt_ltpow: ∀G:todgroup.∀x,y:G.∀n. x < y → S n * x < S n * y. +lemma lt_ltmul: ∀G:todgroup.∀x,y:G.∀n. x < y → S n * x < S n * y. intros; elim n; [simplify; apply flt_plusr; assumption] simplify; apply (ltplus); [assumption] assumption; qed. -lemma ltpow_lt: ∀G:todgroup.∀x,y:G.∀n. S n * x < S n * y → x < y. +lemma ltmul_lt: ∀G:todgroup.∀x,y:G.∀n. S n * x < S n * y → x < y. intros 4; elim n; [apply (plus_cancr_lt ??? 0); assumption] simplify in l; cases (ltplus_orlt ????? l); [assumption] apply f; assumption; @@ -53,16 +53,16 @@ qed. lemma divide_preserves_lt: ∀G:todgroup.∀e:G.∀n.0 sym_plus; simplify; apply (lt_rewl ??? (0+(y+n*y))); [ apply eq_sym; apply zero_neutral] diff --git a/helm/software/matita/dama/ordered_group.ma b/helm/software/matita/dama/ordered_group.ma index 9b70c4f62..e6b6f55ef 100644 --- a/helm/software/matita/dama/ordered_group.ma +++ b/helm/software/matita/dama/ordered_group.ma @@ -181,6 +181,14 @@ apply (lt_rewl ??? x (zero_neutral ??)); assumption; qed. +lemma lt_opp_x_zero_to_lt_zero_x: + ∀G:pogroup.∀x:G. -x < 0 → 0 < x. +intros (G x Lx0); apply (plus_cancr_lt ??? (-x)); +apply (lt_rewl ??? (-x) (zero_neutral ??)); +apply (lt_rewr ??? (-x+x) (plus_comm ???)); +apply (lt_rewr ??? 0 (opp_inverse ??)); +assumption; +qed. lemma lt0plus_orlt: ∀G:pogroup. ∀x,y:G. 0 ≤ x → 0 ≤ y → 0 < x + y → 0 < x ∨ 0 < y. diff --git a/helm/software/matita/dama/sandwich.ma b/helm/software/matita/dama/sandwich.ma index c83136c77..10b54b3e4 100644 --- a/helm/software/matita/dama/sandwich.ma +++ b/helm/software/matita/dama/sandwich.ma @@ -36,13 +36,13 @@ cases (dg_prop (mk_todgroup todg_order todg_division_ H) e 3) 0; cases (dg_prop (mk_todgroup todg_order todg_division_ H) e 2) 0; simplify; intro H3; cut (0