X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fexcedence.ma;h=fa6848e120314b987eaa0f61bf81cec57d42dea6;hb=c0f86a886451a0df3b42a1435e21b5def9f34792;hp=59b8baa4ed1668cae080c58dcaec9cb1fcc99f39;hpb=624a7c13a2ed22ed2535690074c7a08e18de7f13;p=helm.git diff --git a/helm/software/matita/dama/excedence.ma b/helm/software/matita/dama/excedence.ma index 59b8baa4e..fa6848e12 100644 --- a/helm/software/matita/dama/excedence.ma +++ b/helm/software/matita/dama/excedence.ma @@ -21,7 +21,7 @@ include "constructive_higher_order_relations.ma". record excedence : Type ≝ { exc_carr:> Type; - exc_relation: exc_carr → exc_carr → Prop; + exc_relation: exc_carr → exc_carr → Type; exc_coreflexive: coreflexive ? exc_relation; exc_cotransitive: cotransitive ? exc_relation }. @@ -51,10 +51,12 @@ record apartness : Type ≝ { ap_cotransitive: cotransitive ? ap_apart }. -notation "a # b" non associative with precedence 50 for @{ 'apart $a $b}. -interpretation "axiomatic apartness" 'apart x y = +notation "a break # b" non associative with precedence 50 for @{ 'apart $a $b}. +interpretation "apartness" 'apart x y = (cic:/matita/excedence/ap_apart.con _ x y). +definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y. + definition apart ≝ λE:excedence.λa,b:E. a ≰ b ∨ b ≰ a. definition apart_of_excedence: excedence → apartness. @@ -63,7 +65,7 @@ intros (E); apply (mk_apartness E (apart E)); 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; + cases Axy (H H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1; [left; left|right; left|right; right|left; right] assumption] qed. @@ -71,7 +73,7 @@ coercion cic:/matita/excedence/apart_of_excedence.con. definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b). -notation "a ≈ b" non associative with precedence 50 for @{ 'napart $a $b}. +notation "a break ≈ b" non associative with precedence 50 for @{ 'napart $a $b}. interpretation "alikeness" 'napart a b = (cic:/matita/excedence/eq.con _ a b). @@ -79,17 +81,32 @@ lemma eq_reflexive:∀E. reflexive ? (eq E). intros (E); unfold; intros (x); apply ap_coreflexive; qed. -lemma eq_symmetric:∀E.symmetric ? (eq E). +lemma eq_sym_:∀E.symmetric ? (eq E). 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). +lemma eq_sym:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_. + +coercion cic:/matita/excedence/eq_sym.con. + +lemma eq_trans_: ∀E.transitive ? (eq E). (* 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 eq_trans:∀E:apartness.∀x,z,y:E.x ≈ y → y ≈ z → x ≈ z ≝ + λE,x,y,z.eq_trans_ E x z y. + +notation > "'Eq'≈" non associative with precedence 50 for + @{'eqrewrite}. + +interpretation "eq_rew" 'eqrewrite = + (cic:/matita/excedence/eq_trans.con _ _ _). + (* BUG: vedere se ricapita *) +alias id "antisymmetric" = "cic:/matita/constructive_higher_order_relations/antisymmetric.con". lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?). intros 5 (E x y Lxy Lyx); intro H; cases H; [apply Lxy;|apply Lyx] assumption; @@ -120,12 +137,8 @@ intros (E a b Lab); cases Lab (LEab Aab); cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *) qed. -theorem le_le_to_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y. -intros 6 (E x y L1 L2 H); cases H; [apply (L1 H1)|apply (L2 H1)] -qed. - lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y. -unfold apart_of_excedence; unfold apart; simplify; intros; assumption; +intros; assumption; qed. lemma le_rewl: ∀E:excedence.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z. @@ -133,7 +146,80 @@ 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 + @{'excedencerewritel}. + +interpretation "exc_rewl" 'excedencerewritel = + (cic:/matita/excedence/exc_rewl.con _ _ _). + lemma le_rewr: ∀E:excedence.∀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 + @{'excedencerewriter}. + +interpretation "exc_rewr" 'excedencerewriter = + (cic:/matita/excedence/exc_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] +cases (Exy (ap_symmetric ??? a)); +qed. + +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. + +lemma exc_rewl: ∀A:excedence.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z. +intros (A x z y Exy Ayz); elim (exc_cotransitive ???x Ayz); [2:assumption] +cases Exy; right; assumption; +qed. + +lemma exc_rewr: ∀A:excedence.∀x,z,y:A. x ≈ y → z ≰ y → z ≰ x. +intros (A x z y Exy Azy); elim (exc_cotransitive ???x Azy); [assumption] +elim (Exy); left; assumption; +qed. + +lemma lt_rewr: ∀A:excedence.∀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; +qed. + +lemma lt_rewl: ∀A:excedence.∀x,z,y:A. x ≈ y → y < z → x < z. +intros (A x y z E H); split; elim H; +[apply (le_rewl ???? (eq_sym ??? E));| apply (ap_rewl ???? E);] assumption; +qed. + +lemma lt_le_transitive: ∀A:excedence.∀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)] +cases (exc_cotransitive ??? z EXx) (EXz EXz); [cases (LE EXz)] +right; assumption; +qed. + +lemma le_lt_transitive: ∀A:excedence.∀x,y,z:A.x ≤ y → y < z → x < z. +intros (A x y z LE LT); cases LT (LEx APx); split; [apply (le_transitive ???? LE LEx)] +whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)] +cases (exc_cotransitive ??? x EXx) (EXz EXz); [right; assumption] +cases LE; assumption; +qed. + +lemma le_le_eq: ∀E:excedence.∀a,b:E. a ≤ b → b ≤ a → a ≈ b. +intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption; +qed. + +lemma eq_le_le: ∀E:excedence.∀a,b:E. a ≈ b → a ≤ b ∧ b ≤ a. +intros (E x y H); unfold apart_of_excedence in H; unfold apart in H; +simplify in H; split; intro; apply H; [left|right] assumption. +qed. + +lemma ap_le_to_lt: ∀E:excedence.∀a,c:E.c # a → c ≤ a → c < a. +intros; split; assumption; +qed. + +definition total_order_property : ∀E:excedence. Type ≝ + λE:excedence. ∀a,b:E. a ≰ b → b < a. +