1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/excedence/".
17 include "higher_order_defs/relations.ma".
18 include "nat/plus.ma".
19 include "constructive_connectives.ma".
20 include "constructive_higher_order_relations.ma".
22 record excedence : Type ≝ {
24 exc_relation: exc_carr → exc_carr → Prop;
25 exc_coreflexive: coreflexive ? exc_relation;
26 exc_cotransitive: cotransitive ? exc_relation
29 interpretation "excedence" 'nleq a b =
30 (cic:/matita/excedence/exc_relation.con _ a b).
32 definition le ≝ λE:excedence.λa,b:E. ¬ (a ≰ b).
34 interpretation "ordered sets less or equal than" 'leq a b =
35 (cic:/matita/excedence/le.con _ a b).
37 lemma le_reflexive: ∀E.reflexive ? (le E).
38 intros (E); unfold; cases E; simplify; intros (x); apply (H x);
41 lemma le_transitive: ∀E.transitive ? (le E).
42 intros (E); unfold; cases E; simplify; unfold Not; intros (x y z Rxy Ryz H2);
43 cases (c x z y H2) (H4 H5); clear H2; [exact (Rxy H4)|exact (Ryz H5)]
46 record apartness : Type ≝ {
48 ap_apart: ap_carr → ap_carr → Type;
49 ap_coreflexive: coreflexive ? ap_apart;
50 ap_symmetric: symmetric ? ap_apart;
51 ap_cotransitive: cotransitive ? ap_apart
54 notation "a break # b" non associative with precedence 50 for @{ 'apart $a $b}.
55 interpretation "axiomatic apartness" 'apart x y =
56 (cic:/matita/excedence/ap_apart.con _ x y).
58 definition apart ≝ λE:excedence.λa,b:E. a ≰ b ∨ b ≰ a.
60 definition apart_of_excedence: excedence → apartness.
61 intros (E); apply (mk_apartness E (apart E));
62 [1: unfold; cases E; simplify; clear E; intros (x); unfold;
63 intros (H1); apply (H x); cases H1; assumption;
64 |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption;
65 |3: intros (E); unfold; cases E (T f _ cTf); simplify; intros (x y z Axy);
66 cases Axy (H); lapply (cTf ? ? z H) as H1; cases H1; clear Axy H1;
67 [left; left|right; left|right; right|left; right] assumption]
70 coercion cic:/matita/excedence/apart_of_excedence.con.
72 definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b).
74 notation "a break ≈ b" non associative with precedence 50 for @{ 'napart $a $b}.
75 interpretation "alikeness" 'napart a b =
76 (cic:/matita/excedence/eq.con _ a b).
78 lemma eq_reflexive:∀E. reflexive ? (eq E).
79 intros (E); unfold; intros (x); apply ap_coreflexive;
82 lemma eq_symmetric:∀E.symmetric ? (eq E).
83 intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy;
84 apply ap_symmetric; assumption;
87 lemma eq_symmetric_:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x := eq_symmetric.
89 coercion cic:/matita/excedence/eq_symmetric_.con.
91 lemma eq_transitive_: ∀E.transitive ? (eq E).
92 (* bug. intros k deve fare whd quanto basta *)
93 intros 6 (E x y z Exy Eyz); intro Axy; cases (ap_cotransitive ???y Axy);
94 [apply Exy|apply Eyz] assumption.
97 lemma eq_transitive:∀E:apartness.∀x,y,z:E.x ≈ y → y ≈ z → x ≈ z ≝ eq_transitive_.
99 (* BUG: vedere se ricapita *)
100 lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?).
101 intros 5 (E x y Lxy Lyx); intro H;
102 cases H; [apply Lxy;|apply Lyx] assumption;
105 definition lt ≝ λE:excedence.λa,b:E. a ≤ b ∧ a # b.
107 interpretation "ordered sets less than" 'lt a b =
108 (cic:/matita/excedence/lt.con _ a b).
110 lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
111 intros 2 (E x); intro H; cases H (_ ABS);
112 apply (ap_coreflexive ? x ABS);
115 lemma lt_transitive: ∀E.transitive ? (lt E).
116 intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz);
117 split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2;
118 cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]
119 clear Axy Ayz;lapply (exc_cotransitive E) as c; unfold cotransitive in c;
120 lapply (exc_coreflexive E) as r; unfold coreflexive in r;
121 [1: lapply (c ?? y H1) as H3; cases H3 (H4 H4); [cases (Lxy H4)|cases (r ? H4)]
122 |2: lapply (c ?? x H2) as H3; cases H3 (H4 H4); [right; assumption|cases (Lxy H4)]]
125 theorem lt_to_excede: ∀E:excedence.∀a,b:E. (a < b) → (b ≰ a).
126 intros (E a b Lab); cases Lab (LEab Aab);
127 cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *)
130 theorem le_le_to_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y.
131 intros 6 (E x y L1 L2 H); cases H; [apply (L1 H1)|apply (L2 H1)]
134 lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y.
135 unfold apart_of_excedence; unfold apart; simplify; intros; assumption;
138 lemma le_rewl: ∀E:excedence.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z.
139 intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz);
140 intro Xyz; apply Exy; apply unfold_apart; right; assumption;
143 lemma le_rewr: ∀E:excedence.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y.
144 intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz);
145 intro Xyz; apply Exy; apply unfold_apart; left; assumption;
148 lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z.
149 intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption]
150 cases (Exy (ap_symmetric ??? a));
153 lemma ap_rewr: ∀A:apartness.∀x,z,y:A. x ≈ y → z # y → z # x.
154 intros (A x z y Exy Azy); apply ap_symmetric; apply (ap_rewl ???? Exy);
155 apply ap_symmetric; assumption;