From: Enrico Tassi Date: Mon, 10 Dec 2007 11:39:38 +0000 (+0000) Subject: sandeich lemma makeup X-Git-Tag: make_still_working~5714 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=6067115471521e8b9ea805531cb94a0a80774314;p=helm.git sandeich lemma makeup --- diff --git a/helm/software/matita/dama/divisible_group.ma b/helm/software/matita/dama/divisible_group.ma index 44db35a15..6830ef4e1 100644 --- a/helm/software/matita/dama/divisible_group.ma +++ b/helm/software/matita/dama/divisible_group.ma @@ -83,3 +83,17 @@ intros; elim n; [simplify; apply eq_reflexive] simplify; apply (Eq≈ ? (zero_neutral ??)); assumption; qed. +let rec gpowS (G : abelian_group) (x : G) (n : nat) on n : G ≝ + match n with [ O ⇒ x | S m ⇒ gpowS ? x m + x]. + +lemma gpowS_gpow: ∀G:dgroup.∀e:G.∀n. S n * e ≈ gpowS ? e n. +intros (G e n); elim n; simplify; [ + apply (Eq≈ ? (plus_comm ???));apply zero_neutral] +apply (Eq≈ ?? (plus_comm ???)); +apply (Eq≈ (e+S n1*e) ? H); clear H; simplify; apply eq_reflexive; +qed. + +lemma divpow: ∀G:dgroup.∀e:G.∀n. e ≈ gpowS ? (e/n) n. +intros (G e n); apply (Eq≈ ?? (gpowS_gpow ?(e/n) n)); +apply eq_sym; apply divide_divides; +qed. diff --git a/helm/software/matita/dama/excess.ma b/helm/software/matita/dama/excess.ma index 061d2db9b..d8cf09766 100644 --- a/helm/software/matita/dama/excess.ma +++ b/helm/software/matita/dama/excess.ma @@ -146,22 +146,22 @@ intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz); intro Xyz; apply Exy; apply unfold_apart; right; assumption; qed. -notation > "'Ex'≪" non associative with precedence 50 for - @{'excessrewritel}. - -interpretation "exc_rewl" 'excessrewritel = - (cic:/matita/excess/exc_rewl.con _ _ _). - lemma le_rewr: ∀E:excess.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y. intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz); intro Xyz; apply Exy; apply unfold_apart; left; assumption; qed. -notation > "'Ex'≫" non associative with precedence 50 for - @{'excessrewriter}. +notation > "'Le'≪" non associative with precedence 50 for + @{'lerewritel}. -interpretation "exc_rewr" 'excessrewriter = - (cic:/matita/excess/exc_rewr.con _ _ _). +interpretation "le_rewl" 'lerewritel = + (cic:/matita/excess/le_rewl.con _ _ _). + +notation > "'Le'≫" non associative with precedence 50 for + @{'lerewriter}. + +interpretation "le_rewr" 'lerewriter = + (cic:/matita/excess/le_rewr.con _ _ _). lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z. intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption] @@ -183,6 +183,18 @@ intros (A x z y Exy Azy); elim (exc_cotransitive ???x Azy); [assumption] elim (Exy); left; assumption; qed. +notation > "'Ex'≪" non associative with precedence 50 for + @{'excessrewritel}. + +interpretation "exc_rewl" 'excessrewritel = + (cic:/matita/excess/exc_rewl.con _ _ _). + +notation > "'Ex'≫" non associative with precedence 50 for + @{'excessrewriter}. + +interpretation "exc_rewr" 'excessrewriter = + (cic:/matita/excess/exc_rewr.con _ _ _). + lemma lt_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z < y → z < x. intros (A x y z E H); split; elim H; [apply (le_rewr ???? (eq_sym ??? E));|apply (ap_rewr ???? E)] assumption; @@ -193,6 +205,18 @@ intros (A x y z E H); split; elim H; [apply (le_rewl ???? (eq_sym ??? E));| apply (ap_rewl ???? E);] assumption; qed. +notation > "'Lt'≪" non associative with precedence 50 for + @{'ltrewritel}. + +interpretation "lt_rewl" 'ltrewritel = + (cic:/matita/excess/lt_rewl.con _ _ _). + +notation > "'Lt'≫" non associative with precedence 50 for + @{'ltrewriter}. + +interpretation "lt_rewr" 'ltrewriter = + (cic:/matita/excess/lt_rewr.con _ _ _). + lemma lt_le_transitive: ∀A:excess.∀x,y,z:A.x < y → y ≤ z → x < z. intros (A x y z LT LE); cases LT (LEx APx); split; [apply (le_transitive ???? LEx LE)] whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)] diff --git a/helm/software/matita/dama/sandwich.ma b/helm/software/matita/dama/sandwich.ma index 10b54b3e4..95dd0d3cb 100644 --- a/helm/software/matita/dama/sandwich.ma +++ b/helm/software/matita/dama/sandwich.ma @@ -29,31 +29,6 @@ qed. include "sequence.ma". include "metric_lattice.ma". -lemma sandwich_ineq: ∀G:todgroup.∀e:G.0