X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fexcedence.ma;h=84a033c3a8ebb071bf2b56abca7b4fa1bbf8a873;hb=eb9bc52a705fd347a5e1906ec32fd12f86507fe9;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..84a033c3a 100644 --- a/helm/software/matita/dama/excedence.ma +++ b/helm/software/matita/dama/excedence.ma @@ -51,7 +51,7 @@ record apartness : Type ≝ { ap_cotransitive: cotransitive ? ap_apart }. -notation "a # b" non associative with precedence 50 for @{ 'apart $a $b}. +notation "a break # b" non associative with precedence 50 for @{ 'apart $a $b}. interpretation "axiomatic apartness" 'apart x y = (cic:/matita/excedence/ap_apart.con _ x y). @@ -71,7 +71,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). @@ -84,11 +84,16 @@ intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy; apply ap_symmetric; assumption; qed. +lemma eq_symmetric_:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x := eq_symmetric. + +coercion cic:/matita/excedence/eq_symmetric_.con. + lemma eq_transitive: ∀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. + (* BUG: vedere se ricapita *) lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?). intros 5 (E x y Lxy Lyx); intro H; @@ -137,3 +142,13 @@ 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. + +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.