X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fexcedence.ma;h=4931e70d2ab2476131c78fac4d845dcbc4951df4;hb=8de247e8d14940d9f619b6d35107cc5fae11b412;hp=d91c61170f10862ec14abf62f38ccb26f51e54b4;hpb=475120bb8ccfe5877bbe294006170ed7ab0b1fcd;p=helm.git diff --git a/helm/software/matita/dama/excedence.ma b/helm/software/matita/dama/excedence.ma index d91c61170..4931e70d2 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). @@ -104,8 +104,7 @@ lemma lt_coreflexive: ∀E.coreflexive ? (lt E). 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; @@ -121,4 +120,20 @@ 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 +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; +qed. + +lemma le_rewl: ∀E:excedence.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z. +intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz); +intro Xyz; apply Exy; apply unfold_apart; right; assumption; +qed. + +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.