From b0b85f3cad753caba19e785e09cc10ff8a6c00d9 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Tue, 13 Nov 2007 16:40:43 +0000 Subject: [PATCH] snapshot --- .../constructive_higher_order_relations.ma | 6 + matita/dama/excedence.ma | 61 ++++---- matita/dama/groups.ma | 133 ++++++++++-------- matita/dama/ordered_groups.ma | 37 ++--- matita/dama/ordered_sets.ma | 4 + 5 files changed, 125 insertions(+), 116 deletions(-) diff --git a/matita/dama/constructive_higher_order_relations.ma b/matita/dama/constructive_higher_order_relations.ma index b66ba6843..cf58e2640 100644 --- a/matita/dama/constructive_higher_order_relations.ma +++ b/matita/dama/constructive_higher_order_relations.ma @@ -29,3 +29,9 @@ definition symmetric ≝ definition transitive ≝ λC:Type.λle:C→C→Type.∀x,y,z:C.le x y → le y z → le x z. + +definition associative ≝ + λC:Type.λop:C→C→C.λeq:C→C→Type.∀x,y,z. eq (op x (op y z)) (op (op x y) z). + +definition commutative ≝ + λC:Type.λop:C→C→C.λeq:C→C→Type.∀x,y. eq (op x y) (op y x). diff --git a/matita/dama/excedence.ma b/matita/dama/excedence.ma index a78d61184..d91c61170 100644 --- a/matita/dama/excedence.ma +++ b/matita/dama/excedence.ma @@ -43,51 +43,55 @@ intros (E); unfold; cases E; simplify; unfold Not; intros (x y z Rxy Ryz H2); cases (c x z y H2) (H4 H5); clear H2; [exact (Rxy H4)|exact (Ryz H5)] qed. -definition apart ≝ λE:excedence.λa,b:E. a ≰ b ∨ b ≰ a. +record apartness : Type ≝ { + ap_carr:> Type; + ap_apart: ap_carr → ap_carr → Type; + ap_coreflexive: coreflexive ? ap_apart; + ap_symmetric: symmetric ? ap_apart; + ap_cotransitive: cotransitive ? ap_apart +}. notation "a # b" non associative with precedence 50 for @{ 'apart $a $b}. -interpretation "apartness" 'apart a b = (cic:/matita/excedence/apart.con _ a b). +interpretation "axiomatic apartness" 'apart x y = + (cic:/matita/excedence/ap_apart.con _ x y). -lemma apart_coreflexive: ∀E.coreflexive ? (apart E). -intros (E); unfold; cases E; simplify; clear E; intros (x); unfold; -intros (H1); apply (H x); cases H1; assumption; -qed. +definition apart ≝ λE:excedence.λa,b:E. a ≰ b ∨ b ≰ a. -lemma apart_symmetric: ∀E.symmetric ? (apart E). -intros (E); unfold; intros(x y H); cases H; clear H; [right|left] assumption; +definition apart_of_excedence: excedence → apartness. +intros (E); apply (mk_apartness E (apart E)); +[1: unfold; cases E; simplify; clear E; intros (x); unfold; + intros (H1); apply (H x); cases H1; assumption; +|2: unfold; intros(x y H); cases H; clear H; [right|left] assumption; +|3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy); + cases Axy (H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; + [left; left|right; left|right; right|left; right] assumption] qed. -lemma apart_cotrans: ∀E. cotransitive ? (apart E). -intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy); -cases Axy (H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; -[left; left|right; left|right; right|left; right] assumption. -qed. +coercion cic:/matita/excedence/apart_of_excedence.con. -definition eq ≝ λE:excedence.λa,b:E. ¬ (a # b). +definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b). notation "a ≈ b" non associative with precedence 50 for @{ 'napart $a $b}. interpretation "alikeness" 'napart a b = (cic:/matita/excedence/eq.con _ a b). lemma eq_reflexive:∀E. reflexive ? (eq E). -intros (E); unfold; cases E (T f cRf _); simplify; unfold Not; intros (x H); -apply (cRf x); cases H; assumption; +intros (E); unfold; intros (x); apply ap_coreflexive; qed. lemma eq_symmetric:∀E.symmetric ? (eq E). -intros (E); unfold; unfold eq; unfold Not; -intros (x y H1 H2); apply H1; cases H2; [right|left] assumption; +intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy; +apply ap_symmetric; assumption; qed. lemma eq_transitive: ∀E.transitive ? (eq E). -intros (E); unfold; cases E (T f _ cTf); simplify; unfold Not; -intros (x y z H1 H2 H3); cases H3 (H4 H4); clear E H3; lapply (cTf ? ? y H4) as H5; -cases H5; clear H5 H4 cTf; [1,4: apply H1|*:apply H2] clear H1 H2; -[1,3:left|*:right] assumption; +(* bug. intros k deve fare whd quanto basta *) +intros 6 (E x y z Exy Eyz); intro Axy; cases (ap_cotransitive ???y Axy); +[apply Exy|apply Eyz] assumption. qed. - -lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq E). -intros (E); unfold; intros (x y Lxy Lyx); unfold; unfold; intros (H); +(* BUG: vedere se ricapita *) +lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?). +intros 5 (E x y Lxy Lyx); intro H; cases H; [apply Lxy;|apply Lyx] assumption; qed. @@ -97,10 +101,11 @@ interpretation "ordered sets less than" 'lt a b = (cic:/matita/excedence/lt.con _ a b). lemma lt_coreflexive: ∀E.coreflexive ? (lt E). -intros (E); unfold; unfold Not; intros (x H); cases H (_ ABS); -apply (apart_coreflexive ? x ABS); +intros 2 (E x); intro H; cases H (_ ABS); +apply (ap_coreflexive ? x ABS); qed. +(* lemma lt_transitive: ∀E.transitive ? (lt E). intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz); split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2; @@ -115,3 +120,5 @@ theorem lt_to_excede: ∀E:excedence.∀a,b:E. (a < b) → (b ≰ a). intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *) qed. + +*) \ No newline at end of file diff --git a/matita/dama/groups.ma b/matita/dama/groups.ma index 699bd73fc..a35b43d3b 100644 --- a/matita/dama/groups.ma +++ b/matita/dama/groups.ma @@ -14,47 +14,36 @@ set "baseuri" "cic:/matita/groups/". -include "higher_order_defs/functions.ma". -include "nat/nat.ma". -include "nat/orders.ma". -include "constructive_connectives.ma". +include "excedence.ma". -definition left_neutral \def λC,op.λe:C. ∀x:C. op e x = x. - -definition right_neutral \def λC,op. λe:C. ∀x:C. op x e=x. - -definition left_inverse \def λC,op.λe:C.λinv:C→C. ∀x:C. op (inv x) x = e. - -definition right_inverse \def λC,op.λe:C.λ inv: C→ C. ∀x:C. op x (inv x)=e. +definition left_neutral ≝ λC:apartness.λop.λe:C. ∀x:C. op e x ≈ x. +definition right_neutral ≝ λC:apartness.λop. λe:C. ∀x:C. op x e ≈ x. +definition left_inverse ≝ λC:apartness.λop.λe:C.λinv:C→C. ∀x:C. op (inv x) x ≈ e. +definition right_inverse ≝ λC:apartness.λop.λe:C.λ inv: C→ C. ∀x:C. op x (inv x) ≈ e. +definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y. +(* ALLOW DEFINITION WITH SOME METAS *) definition distributive_left ≝ - λA:Type.λf:A→A→A.λg:A→A→A. - ∀x,y,z. f x (g y z) = g (f x y) (f x z). + λA:apartness.λf:A→A→A.λg:A→A→A. + ∀x,y,z. f x (g y z) ≈ g (f x y) (f x z). definition distributive_right ≝ - λA:Type.λf:A→A→A.λg:A→A→A. - ∀x,y,z. f (g x y) z = g (f x z) (f y z). - -record is_abelian_group (C:Type) (plus:C→C→C) (zero:C) (opp:C→C) : Prop \def - { (* abelian additive semigroup properties *) - plus_assoc_: associative ? plus; - plus_comm_: symmetric ? plus; - (* additive monoid properties *) - zero_neutral_: left_neutral ? plus zero; - (* additive group properties *) - opp_inverse_: left_inverse ? plus zero opp - }. - -record abelian_group : Type \def - { carrier:> Type; - plus: carrier → carrier → carrier; - zero: carrier; - opp: carrier → carrier; - ag_abelian_group_properties: is_abelian_group ? plus zero opp - }. + λA:apartness.λf:A→A→A.λg:A→A→A. + ∀x,y,z. f (g x y) z ≈ g (f x z) (f y z). + +record abelian_group : Type ≝ + { carr:> apartness; + plus: carr → carr → carr; + zero: carr; + opp: carr → carr; + plus_assoc: associative ? plus (eq carr); + plus_comm: commutative ? plus (eq carr); + zero_neutral: left_neutral ? plus zero; + opp_inverse: left_inverse ? plus zero opp; + plus_strong_ext: ∀z.strong_ext ? (plus z) +}. -notation "0" with precedence 89 -for @{ 'zero }. +notation "0" with precedence 89 for @{ 'zero }. interpretation "Abelian group zero" 'zero = (cic:/matita/groups/zero.con _). @@ -71,37 +60,57 @@ definition minus ≝ interpretation "Abelian group minus" 'minus a b = (cic:/matita/groups/minus.con _ a b). -theorem plus_assoc: ∀G:abelian_group. associative ? (plus G). - intro; - apply (plus_assoc_ ? ? ? ? (ag_abelian_group_properties G)). +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] +cases (Exy (ap_symmetric ??? a)); qed. - -theorem plus_comm: ∀G:abelian_group. symmetric ? (plus G). - intro; - apply (plus_comm_ ? ? ? ? (ag_abelian_group_properties G)). + +lemma ap_rewr: ∀A:apartness.∀x,z,y:A. x ≈ y → z # y → z # x. +intros (A x z y Exy Azy); apply ap_symmetric; apply (ap_rewl ???? Exy); +apply ap_symmetric; assumption; qed. -theorem zero_neutral: ∀G:abelian_group. left_neutral ? (plus G) 0. - intro; - apply (zero_neutral_ ? ? ? ? (ag_abelian_group_properties G)). +definition ext ≝ λA:apartness.λop:A→A. ∀x,y. x ≈ y → op x ≈ op y. + +lemma strong_ext_to_ext: ∀A:apartness.∀op:A→A. strong_ext ? op → ext ? op. +intros 6 (A op SEop x y Exy); intro Axy; apply Exy; apply SEop; assumption; +qed. + +lemma f_plusl: ∀G:abelian_group.∀x,y,z:G. y ≈ z → x+y ≈ x+z. +intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_ext ? x)); +assumption; +qed. + +lemma plus_strong_extr: ∀G:abelian_group.∀z:G.strong_ext ? (λx.x + z). +intros 5 (G z x y A); simplify in A; +lapply (plus_comm ? z x) as E1; lapply (plus_comm ? z y) as E2; +lapply (ap_rewl ???? E1 A) as A1; lapply (ap_rewr ???? E2 A1) as A2; +apply (plus_strong_ext ???? A2); qed. - -theorem opp_inverse: ∀G:abelian_group. left_inverse ? (plus G) 0 (opp G). - intro; - apply (opp_inverse_ ? ? ? ? (ag_abelian_group_properties G)). + +lemma feq_plusr: ∀G:abelian_group.∀x,y,z:G. y ≈ z → y+x ≈ z+x. +intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_extr ? x)); +assumption; +qed. + +lemma fap_plusl: ∀G:abelian_group.∀x,y,z:G. y # z → x+y # x+z. +intros (G x y z Ayz); apply (plus_strong_ext ? (-x)); +apply (ap_rewl ??? ((-x + x) + y)); +[1: apply plus_assoc; +|2: apply (ap_rewr ??? ((-x +x) +z)); + [1: apply plus_assoc; + |2: apply (ap_rewl ??? (0 + y)); + [1: apply (feq_plusr ???? (opp_inverse ??)); + |2: apply (ap_rewl ???? (zero_neutral ? y)); apply (ap_rewr ??? (0 + z)); + [1: apply (feq_plusr ???? (opp_inverse ??)); + |2: apply (ap_rewr ???? (zero_neutral ? z)); assumption;]]]] qed. -lemma cancellationlaw: ∀G:abelian_group.∀x,y,z:G. x+y=x+z → y=z. -intros; -generalize in match (eq_f ? ? (λa.-x +a) ? ? H); -intros; clear H; -rewrite < plus_assoc in H1; -rewrite < plus_assoc in H1; -rewrite > opp_inverse in H1; -rewrite > zero_neutral in H1; -rewrite > zero_neutral in H1; -assumption. -qed. +lemma plus_canc: ∀G:abelian_group.∀x,y,z:G. x+y ≈ x+z → y ≈ z. +intros 6 (G x y z E Ayz); apply E; apply fap_plusl; assumption; +qed. + +(* theorem eq_opp_plus_plus_opp_opp: ∀G:abelian_group.∀x,y:G. -(x+y) = -x + -y. intros; @@ -137,4 +146,6 @@ theorem eq_zero_opp_zero: ∀G:abelian_group.0=-0. rewrite > zero_neutral; reflexivity ]. -qed. \ No newline at end of file +qed. + +*) \ No newline at end of file diff --git a/matita/dama/ordered_groups.ma b/matita/dama/ordered_groups.ma index c9cab27f8..4d2e18e28 100644 --- a/matita/dama/ordered_groups.ma +++ b/matita/dama/ordered_groups.ma @@ -18,38 +18,19 @@ include "groups.ma". include "ordered_sets.ma". record pre_ordered_abelian_group : Type ≝ - { og_abelian_group:> abelian_group; - og_tordered_set_: tordered_set; - og_with: exc_carr og_tordered_set_ = og_abelian_group + { og_abelian_group_: abelian_group; + og_tordered_set:> tordered_set; + og_with: carr og_abelian_group_ = og_tordered_set }. -lemma og_tordered_set: pre_ordered_abelian_group → tordered_set. -intro G; apply mk_tordered_set; -[1: apply mk_pordered_set; - [1: apply (mk_excedence G); - [1: cases G; clear G; simplify; rewrite < H; clear H; - cases og_tordered_set_; clear og_tordered_set_; simplify; - cases tos_poset; simplify; cases pos_carr; simplify; assumption; - |2: cases G; simplify; cases H; simplify; clear H; - cases og_tordered_set_; simplify; clear og_tordered_set_; - cases tos_poset; simplify; cases pos_carr; simplify; - intros; apply H; - |3: cases G; simplify; cases H; simplify; cases og_tordered_set_; simplify; - cases tos_poset; simplify; cases pos_carr; simplify; - intros; apply c; assumption] - |2: cases G; simplify; - cases H; simplify; clear H; cases og_tordered_set_; simplify; - cases tos_poset; simplify; assumption;] -|2: simplify; (* SLOW, senza la simplify il widget muore *) - cases G; simplify; - generalize in match (tos_totality og_tordered_set_); - unfold total_order_property; - cases H; simplify; cases og_tordered_set_; simplify; - cases tos_poset; simplify; cases pos_carr; simplify; - intros; apply f; assumption;] +lemma og_abelian_group: pre_ordered_abelian_group → abelian_group. +intro G; apply (mk_abelian_group G); [1,2,3: rewrite < (og_with G)] +[apply (plus (og_abelian_group_ G));|apply zero;|apply opp] +unfold apartness_OF_pre_ordered_abelian_group; cases (og_with G); simplify; +[apply plus_assoc|apply plus_comm|apply zero_neutral|apply opp_inverse|apply plus_strong_ext] qed. -coercion cic:/matita/ordered_groups/og_tordered_set.con. +coercion cic:/matita/ordered_groups/og_abelian_group.con. definition is_ordered_abelian_group ≝ λG:pre_ordered_abelian_group. ∀f,g,h:G. f≤g → f+h≤g+h. diff --git a/matita/dama/ordered_sets.ma b/matita/dama/ordered_sets.ma index a7f772f34..9cd9f58a9 100644 --- a/matita/dama/ordered_sets.ma +++ b/matita/dama/ordered_sets.ma @@ -132,6 +132,8 @@ cases E (T f cRf cTf); simplify; |2: intros (x y z); apply Or_symmetric; apply cTf; assumption;] qed. +(* + definition reverse_pordered_set: pordered_set → pordered_set. intros (p); apply (mk_pordered_set (reverse_excedence p)); generalize in match (reverse_excedence p); intros (E); cases E (T f cRf cTf); @@ -212,6 +214,8 @@ record cotransitively_ordered_set: Type := }. *) +*) + definition total_order_property : ∀E:excedence. Type ≝ λE:excedence. ∀a,b:E. a ≰ b → a < b. -- 2.39.2