]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/dama/excedence.ma
moved to pkg-ocaml-maint
[helm.git] / matita / dama / excedence.ma
index 6d1fc4dd7cf1c292541e963c1b15c051ef564060..3b2afd5dc982b86c8a8308018b72bc599bc04fc2 100644 (file)
@@ -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; (* Big bug: era in Prop!!! *)
   exc_coreflexive: coreflexive ? exc_relation;
   exc_cotransitive: cotransitive ? exc_relation 
 }.
@@ -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).
 
@@ -63,7 +63,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 +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). 
 
@@ -79,16 +79,23 @@ 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,y,z:E.x ≈ y → y ≈ z → x ≈ z ≝ eq_trans_.
+
 (* BUG: vedere se ricapita *)
 lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?).
 intros 5 (E x y Lxy Lyx); intro H;
@@ -119,3 +126,33 @@ 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.
+
+(* CSC: lo avevi gia' dimostrato; ho messo apply! *)
+theorem le_le_to_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y.
+apply le_antisymmetric;
+qed.
+
+(* CSC: perche' quel casino: bastava intros; assumption! *)
+lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y.
+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.
+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.
\ No newline at end of file