]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/topology/cantor.ma
6a79dd2d808df2478f6a4ad25568e41e1eb68894
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "topology/igft.ma".
16
17 ntheorem axiom_cond: âˆ€A:Ax.∀a:A.∀i:𝐈 a.a â—ƒ đ‚ a i.
18 #A; #a; #i; @2 i; #x; #H; @; napply H;
19 nqed.
20
21 nlemma hint_auto1 : âˆ€A,U,V. (∀x.x âˆˆ U â†’ x â—ƒ V) â†’ cover_set cover A U V.
22 nnormalize; /2/.
23 nqed.
24
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/;
30 nqed.
31
32 ndefinition emptyset: âˆ€A.Ί^A â‰ ÎťA.{x | False}.
33
34 notation "∅" non associative with precedence 90 for @{ 'empty }.
35 interpretation "empty" 'empty = (emptyset ?).
36
37 naxiom EM : âˆ€A:Ax.∀a:A.∀i_star.(a âˆˆ đ‚ a i_star) âˆ¨ ÂŹ( a âˆˆ đ‚ a i_star).
38
39 alias symbol "covers" = "covers".
40 ntheorem th2_3 :
41   âˆ€A:Ax.∀a:A. a â—ƒ âˆ… â†’ âˆƒi. ÂŹ a âˆˆ đ‚ a i.
42 #A; #a; #H; nelim H;
43 ##[ #n; *;
44 ##| #b; #i_star; #IH1; #IH2; ncases (EM â€Ś b i_star); /3/;
45 ##] 
46 nqed.
47
48 ninductive eq1 (A : Type[0]) : Type[0] â†’ CProp[0] â‰ 
49 | refl1 : eq1 A A.
50
51 notation "hvbox( a break âˆź b)" non associative with precedence 40 
52 for @{ 'eqT $a $b }.
53
54 interpretation "eq between types" 'eqT a b = (eq1 a b).
55
56 ninductive unit : Type[0] â‰ one : unit.
57
58 nrecord uAx : Type[1] â‰ {
59   uax_ : Ax;
60   with_ : âˆ€a:uax_.𝐈 a âˆź unit
61 }.
62
63 ndefinition uax : uAx â†’ Ax.
64 #A; @ (uax_ A) (Îťx.unit); #a; #_; 
65 napply (𝐂 a ?);  nlapply one; ncases (with_ A a); //; 
66 nqed.
67
68 ncoercion uax : âˆ€u:uAx. Ax â‰ uax on _u : uAx to Ax.
69
70 naxiom A: Type[0].
71 naxiom S: A â†’ ÎŠ^A.
72
73 ndefinition axs: uAx.
74 @; ##[ @ A (Îť_.unit) (Îťa,x.S a); ##| #_; @; ##]
75 nqed.
76  
77 alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
78 unification hint 0 â‰” ;
79          x â‰Ÿ axs  
80   (* -------------- *) âŠ˘
81          S (uax x) â‰Ą A. (* XXX: bug coercions/ disamb multipasso che ne fa 1 solo*) 
82
83 ntheorem col2_4 :
84   âˆ€A:uAx.∀a:uax A. a â—ƒ âˆ… â†’ ÂŹ a âˆˆ đ‚ a one.
85 #A; #a; #H; nelim H;
86 ##[ #n; *;
87 ##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 â€Ś H3); /2/;
88 ##]
89 nqed.
90
91 (* bug interpretazione non aggiunta per âˆ… *)
92 ndefinition Z : ÎŠ^axs â‰ { x | x â—ƒ (emptyset ?) }.
93
94 ntheorem cover_monotone: âˆ€A:Ax.∀a:A.∀U,V.U âŠ† V â†’ a â—ƒ U â†’ a â—ƒ V.
95 #A; #a; #U; #V; #HUV; #H; nelim H; /3/; 
96 nqed.
97
98 ntheorem th3_1: ÂŹâˆƒa:axs.Z âŠ† S a âˆ§ S a âŠ† Z.
99 *; #a; *; #ZSa; #SaZ; 
100 ncut (a â—ƒ Z); ##[
101   nlapply (axiom_cond â€Ś a one); #AxCon; nchange in AxCon with (a â—ƒ S a);
102   napply (cover_monotone â€Ś AxCon); nassumption; ##] #H; 
103 ncut (a â—ƒ âˆ…); ##[ napply (transitivity â€Ś H); nwhd in match Z; //; ##] #H1;
104 ncut (ÂŹ a âˆˆ S a); ##[ napply (col2_4 â€Ś H1); ##] #H2;
105 ncut (a âˆˆ S a); ##[ napply ZSa; napply H1; ##] #H3;
106 /2/;
107 nqed.
108
109 include "nat/nat.ma".
110
111 naxiom phi : nat â†’ nat â†’ nat.
112
113 notation > "ϕ" non associative with precedence 90 for @{ 'phi }.
114 interpretation "phi" 'phi = phi.
115  
116 notation < "ϕ a i" non associative with precedence 90 for @{ 'phi2 $a $i}.
117 interpretation "phi2" 'phi2 a i = (phi a i). 
118 notation < "ϕ a" non associative with precedence 90 for @{ 'phi1 $a }.
119 interpretation "phi2" 'phi1 a = (phi a). 
120
121 ndefinition caxs : uAx. 
122 @; ##[ @ nat (Îť_.unit); #a; #_; napply { x | Ď• a x = O } ##| #_; @; ##]
123 nqed.
124
125
126 alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
127 unification hint 0 â‰” ;
128          x â‰Ÿ caxs  
129   (* -------------- *) âŠ˘
130          S (uax x) â‰Ą nat. 
131
132 naxiom h : nat â†’ nat. 
133
134 alias symbol "eq" = "leibnitz's equality".
135 alias symbol "eq" = "setoid1 eq".
136 alias symbol "covers" = "covers".
137 alias symbol "eq" = "leibnitz's equality".
138 naxiom Ph :  âˆ€x.h x = O \liff  x â—ƒ âˆ….
139
140 nlemma replace_char:
141   âˆ€A:Ax.∀U,V.U âŠ† V â†’ V âŠ† U â†’ âˆ€a:A.a â—ƒ U â†’ a â—ƒ V.
142 #A; #U; #V;  #UV; #VU; #a; #aU; nelim aU; /3/;
143 nqed.
144
145 ntheorem th_ch3: ÂŹâˆƒa:caxs.∀x.ϕ a x = h x.
146 *; #a; #H;
147 ncut (a â—ƒ { x | x â—ƒ âˆ…}); ##[
148   napply (replace_char â€Ś { x | h x = O }); ##[ ##1,2: #x; ncases (Ph x); /2/; ##]
149   napply (replace_char â€Ś { x | Ď• a x = O }); ##[##1,2: #x; nrewrite > (H x); //; ##]
150   napply (axiom_cond â€Ś a one); ##] #H1;
151 ncut (a â—ƒ âˆ…); ##[ napply (transitivity â€Ś H1); //; ##] #H2;
152 nlapply (col2_4 â€ŚH2); #H3;
153 ncut (a âˆˆ đ‚ a one); ##[
154   nnormalize; ncases (Ph a); nrewrite > (H a); /2/; ##] #H4;
155 /2/;
156 nqed.