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 include "topology/igft.ma".
17 ntheorem axiom_cond: âA:Ax.âa:A.âi:đ a.a â đ a i.
18 #A; #a; #i; @2 i; #x; #H; @; napply H;
21 nlemma hint_auto1 : âA,U,V. (âx.x â U â x â V) â cover_set cover A U V.
25 alias symbol "covers" (instance 1) = "covers".
26 alias symbol "covers" (instance 2) = "covers set".
27 alias symbol "covers" (instance 3) = "covers".
28 ntheorem transitivity: âA:Ax.âa:A.âU,V. a â U â U â V â a â V.
29 #A; #a; #U; #V; #aU; #UV; nelim aU; /3/;
32 ndefinition emptyset: âA.Ί^A â ÎťA.{x | False}.
34 notation "â
" non associative with precedence 90 for @{ 'empty }.
35 interpretation "empty" 'empty = (emptyset ?).
37 naxiom EM : âA:Ax.âa:A.âi_star.(a â đ a i_star) ⨠( a â đ a i_star).
39 alias symbol "covers" = "covers".
41 âA:Ax.âa:A. a â â
â âi. ÂŹ a â đ a i.
44 ##| #b; #i_star; #IH1; #IH2; ncases (EM ⌠b i_star); /2/;
48 ninductive eq1 (A : Type[0]) : Type[0] â CProp[0] â
51 notation "hvbox( a break âź b)" non associative with precedence 40
54 interpretation "eq between types" 'eqT a b = (eq1 a b).
56 ninductive unit : Type[0] â one : unit.
58 nrecord uAx : Type[1] â {
60 with_ : âa:uax_.đ a âź unit
63 ndefinition uax : uAx â Ax.
64 #A; @ (uax_ A) (Îťx.unit); #a; #_;
65 napply (đ a ?); nlapply one; ncases (with_ A a); //;
68 ncoercion uax : âu:uAx. Ax â uax on _u : uAx to Ax.
74 @; ##[ @ A (Îť_.unit) (Îťa,x.S a); ##| #_; @; ##]
77 alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
78 unification hint 0 â ;
80 (* -------------- *) â˘
84 âA:uAx.âa:uax A. a â â
â ÂŹ a â đ a one.
87 ##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 ⌠H3); /2/;
91 ndefinition Z : Ί^axs â { x | x â â
}.
93 ntheorem cover_monotone: âA:Ax.âa:A.âU,V.U â V â a â U â a â V.
94 #A; #a; #U; #V; #HUV; #H; nelim H; /3/;
97 ntheorem th3_1: ÂŹâa:axs.Z â S a ⧠S a â Z.
100 nlapply (axiom_cond ⌠a one); #AxCon; nchange in AxCon with (a â S a);
101 napply (cover_monotone ⌠AxCon); nassumption; ##] #H;
102 ncut (a â â
); ##[ napply (transitivity ⌠H); nwhd in match Z; //; ##] #H1;
103 ncut (ÂŹ a â S a); ##[ napply (col2_4 ⌠H1); ##] #H2;
104 ncut (a â S a); ##[ napply ZSa; napply H1; ##] #H3;
108 include "nat/nat.ma".
110 naxiom phi : nat â nat â nat.
112 notation > "Ď" non associative with precedence 90 for @{ 'phi }.
113 interpretation "phi" 'phi = phi.
115 notation < "Ď a i" non associative with precedence 90 for @{ 'phi2 $a $i}.
116 interpretation "phi2" 'phi2 a i = (phi a i).
117 notation < "Ď a" non associative with precedence 90 for @{ 'phi1 $a }.
118 interpretation "phi2" 'phi1 a = (phi a).
120 ndefinition caxs : uAx.
121 @; ##[ @ nat (Îť_.unit); #a; #_; napply { x | Ď a x = O } ##| #_; @; ##]
125 alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
126 unification hint 0 â ;
128 (* -------------- *) â˘
131 naxiom h : nat â nat.
133 alias symbol "eq" = "leibnitz's equality".
134 alias symbol "eq" = "setoid1 eq".
135 alias symbol "covers" = "covers".
136 alias symbol "eq" = "leibnitz's equality".
137 naxiom Ph : âx.h x = O \liff x â â
.
140 âA:Ax.âU,V.U â V â V â U â âa:A.a â U â a â V.
141 #A; #U; #V; #UV; #VU; #a; #aU; nelim aU; /3/;
144 ntheorem th_ch3: ÂŹâa:caxs.âx.Ď a x = h x.
146 ncut (a â { x | x â â
}); ##[
147 napply (replace_char ⌠{ x | h x = O }); ##[ ##1,2: #x; ncases (Ph x); /2/; ##]
148 napply (replace_char ⌠{ x | Ď a x = O }); ##[##1,2: #x; nrewrite > (H x); //; ##]
149 napply (axiom_cond ⌠a one); ##] #H1;
150 ncut (a â â
); ##[ napply (transitivity ⌠H1); //; ##] #H2;
151 nlapply (col2_4 âŚH2); #H3;
152 ncut (a â đ a one); ##[
153 nnormalize; ncases (Ph a); nrewrite > (H a); /2/; ##] #H4;