]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/topology/igft3.ma
New paramod tac.
[helm.git] / helm / software / matita / nlibrary / topology / igft3.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 "arithmetics/nat.ma".
16 include "datatypes/bool.ma".
17
18 ndefinition two ≝ S (S O).
19 ndefinition natone ≝ S O.
20 ndefinition four ≝ two * two.
21 ndefinition eight ≝ two * four.
22 ndefinition natS ≝ S.
23
24 include "topology/igft.ma".
25
26 nlemma hint_auto2 : ∀T.∀U,V:Ω^T.(∀x.x ∈ U → x ∈ V) → U ⊆ V.
27 nnormalize; nauto;
28 nqed.
29
30 ninductive Sigma (A: Type[0]) (P: A → CProp[0]) : Type[0] ≝
31  mk_Sigma: ∀a:A. P a → Sigma A P.
32
33 (*<< To be moved in igft.ma *)
34 ninductive ncover (A : nAx) (U : Ω^A) : A → CProp[0] ≝
35 | ncreflexivity : ∀a. a ∈ U → ncover A U a
36 | ncinfinity    : ∀a. ∀i. (∀y.Sigma ? (λj.y = 𝐝 a i j) → ncover A U y) → ncover A U a.
37
38 interpretation "ncovers" 'covers a U = (ncover ? U a).
39
40 ntheorem ncover_cover_ok: ∀A:nAx.∀U.∀a:A. a ◃ U → cover (Ax_of_nAx A) U a.
41  #A; #U; #a; #H; nelim H
42   [ #n; #H1; @1; nassumption
43   | #a; #i; #IH; #H; @2 [ napply i ]
44     nnormalize; #y; *; #j; #E; nrewrite > E;
45     napply H;
46     /2/ ]
47 nqed.
48
49 ntheorem cover_ncover_ok: ∀A:Ax.∀U.∀a:A. a ◃ U → ncover (nAx_of_Ax A) U a.
50  #A; #U; #a; #H; nelim H
51   [ #n; #H1; @1; nassumption
52   | #a; #i; #IH; #H; @2 [ napply i ] #y; *; #j; #E; nrewrite > E; ncases j; #x; #K;
53     napply H; nnormalize; nassumption.
54 nqed.
55
56 ndefinition ncoverage : ∀A:nAx.∀U:Ω^A.Ω^A ≝ λA,U.{ a | a ◃ U }.
57
58 interpretation "ncoverage cover" 'coverage U = (ncoverage ? U).
59
60 (*>> To be moved in igft.ma *)
61
62 (*XXX
63 nlemma ncover_ind':
64  ∀A:nAx.∀U,P:Ω^A.
65    (U ⊆ P) → (∀a:A.∀i:𝐈 a.(∀j. 𝐝 a i j ◃ U) → (∀j. 𝐝 a i j ∈ P) → a ∈ P) →
66     ◃ U ⊆ P.
67  #A; #U; #P; #refl; #infty; #a; #H; nelim H
68   [ nauto | (*nauto depth=4;*) #b; #j; #K1; #K2; 
69             napply infty; nauto; ##] 
70 nqed.
71
72 alias symbol "covers" (instance 3) = "ncovers".
73 nlemma cover_ind'':
74  ∀A:nAx.∀U:Ω^A.∀P:A → CProp[0].
75   (∀a. a ∈ U → P a) → (∀a:A.∀i:𝐈 a.(∀j. 𝐝 a i j ◃ U) → (∀j. P (𝐝 a i j)) → P a) →
76    ∀b. b ◃ U → P b.
77  #A; #U; #P; nletin V ≝ {x | P x}; napply (ncover_ind' … V).
78 nqed.
79 *)
80
81 (*********** from Cantor **********)
82 ninductive eq1 (A : Type[0]) : Type[0] → CProp[0] ≝
83 | refl1 : eq1 A A.
84
85 notation "hvbox( a break ∼ b)" non associative with precedence 40
86 for @{ 'eqT $a $b }.
87
88 interpretation "eq between types" 'eqT a b = (eq1 a b).
89
90 ninductive unit : Type[0] ≝ one : unit.
91
92 ninductive option (A: Type[0]) : Type[0] ≝
93    None: option A
94  | Some: A → option A
95  | Twice: A → A → option A.
96
97 nrecord uuAx : Type[1] ≝ {
98   uuS : Type[0];
99   uuC : uuS → option uuS
100 }.
101
102 ndefinition uuax : uuAx → nAx.
103 #A; @ (uuS A)
104   [ #a; ncases (uuC … a) [ napply False | #_; napply unit | #_; #_; napply unit]
105 ##| #a; ncases (uuC … a); nnormalize
106      [ #H; napply (False_rect_Type1 … H)
107      | #_; #_; napply unit
108      | #_; #_; #_; napply bool ]
109 ##| #a; ncases (uuC … a); nnormalize
110      [ #H; napply (False_rect_Type1 … H)
111      | #b; #_; #_; napply b
112      | #b1; #b2; #_; * [ napply b1 | napply b2]##]##]
113 nqed.
114
115 ncoercion uuax : ∀u:uuAx. nAx ≝ uuax on _u : uuAx to nAx.
116
117 nlemma eq_rect_Type0_r':
118  ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → P x p.
119  #A; #a; #x; #p; ncases p; nauto;
120 nqed.
121
122 nlemma eq_rect_Type0_r:
123  ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
124  #A; #a; #P; #p; #x0; #p0; napply (eq_rect_Type0_r' ??? p0); nassumption.
125 nqed.
126
127 nrecord memdec (A: Type[0]) (U:Ω^A) : Type[0] ≝
128  { decide_mem:> A → bool;
129    decide_mem_ok: ∀x. decide_mem x = true → x ∈ U;
130    decide_mem_ko: ∀x. decide_mem x = false → ¬ (x ∈ U)
131  }.
132
133 (*********** end from Cantor ********)
134
135 nlemma csc_sym_eq: ∀A,x,y. eq A x y → eq A y x.
136  #A; #x; #y; #H; ncases H; @1.
137 nqed.
138
139 nlemma csc_eq_rect_CProp0_r':
140  ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. CProp[0]. P a → P x.
141  #A; #a; #x; #p; #P; #H;
142  napply (match csc_sym_eq ??? p return λa.λ_.P a with [ refl ⇒ H ]).
143 nqed.
144  
145 nlet rec cover_rect
146  (A:uuAx) (U:Ω^(uuax A)) (memdec: memdec … U) (P:uuax A → Type[0])
147   (refl: ∀a:uuax A. a ∈ U → P a)
148   (infty: ∀a:uuax A.∀i: 𝐈 a.(∀j. 𝐝 a i j ◃ U) → (∀j.P (𝐝 a i j)) → P a)
149    (b:uuax A) (p: b ◃ U) on p : P b
150 ≝ ?.
151  nlapply (decide_mem_ok … memdec b); nlapply (decide_mem_ko … memdec b);
152  ncases (decide_mem … memdec b)
153   [ #_; #H; napply refl; nauto
154   | #H; #_; ncut (uuC … b=uuC … b) [nauto] ncases (uuC … b) in ⊢ (???% → ?)
155     [ #E; napply False_rect_Type0; ncut (b=b) [nauto] ncases p in ⊢ (???% → ?)
156       [ #a; #K; #E2; napply H [ nauto | nrewrite > E2; nauto ]
157     ##| #a; #i; #K; #E2; nrewrite < E2 in i; nnormalize; nrewrite > E; nnormalize;
158         nauto]
159   ##| #a; #E;
160       ncut (a ◃ U)
161        [ nlapply E; nlapply (H ?) [nauto] ncases p
162           [ #x; #Hx; #K1; #_; ncases (K1 Hx)
163         ##| #x; #i; #Hx; #K1; #E2; napply Hx; ngeneralize in match i; nnormalize;
164             nrewrite > E2; nnormalize; /2/ ]##]
165       #Hcut;
166       nlapply (infty b); nnormalize; nrewrite > E; nnormalize; #H2;
167       napply (H2 one); #y
168        [ napply Hcut
169      ##| napply (cover_rect A U memdec P refl infty a); nauto ]
170   ##| #a; #a1; #E;
171       ncut (a ◃ U)
172        [ nlapply E; nlapply (H ?) [nauto] ncases p
173           [ #x; #Hx; #K1; #_; ncases (K1 Hx)
174         ##| #x; #i; #Hx; #K1; #E2; napply Hx; ngeneralize in match i; nnormalize;
175             nrewrite > E2; nnormalize; #_; @1 (true); /2/ ]##]
176       #Hcut;
177       ncut (a1 ◃ U)
178        [ nlapply E; nlapply (H ?) [nauto] ncases p
179           [ #x; #Hx; #K1; #_; ncases (K1 Hx)
180         ##| #x; #i; #Hx; #K1; #E2; napply Hx; ngeneralize in match i; nnormalize;
181             nrewrite > E2; nnormalize; #_; @1 (false); /2/ ]##]
182       #Hcut1;
183       nlapply (infty b); nnormalize; nrewrite > E; nnormalize; #H2;
184       napply (H2 one); #y; ncases y; nnormalize
185        [##1,2: nassumption
186        | napply (cover_rect A U memdec P refl infty a); nauto
187        | napply (cover_rect A U memdec P refl infty a1); nauto]
188 nqed.
189
190 (********* Esempio:
191    let rec skip n =
192      match n with
193       [ O ⇒ 1
194       | S m ⇒
195          match m with
196           [ O ⇒ skipfact O
197           | S _ ⇒ S m * skipfact (pred m) * skipfact (pred m) ]]
198 **)
199
200 ntheorem psym_plus: ∀n,m. n + m = m + n.
201  #n; nelim n
202   [ #m; nelim m; //; #n0; #H;
203     nchange with (natS n0 = natS (n0 + O));
204     nrewrite < H; //
205   | #n0; #H; #m; nchange with (S (n0 + m) = m + S n0);
206     nrewrite > (H …);
207     nelim m; //;
208     #n1; #E; nrewrite > E; //]
209 nqed.
210
211 nlemma easy1: ∀n:nat. two * (S n) = two + two * n.
212  #n; nelim n;//;
213  #n0; #H; nnormalize;
214  nrewrite > (psym_plus ??);
215  nrewrite > H; nnormalize;
216  nrewrite > (psym_plus ??);
217  //.
218 nqed.
219
220 ndefinition skipfact_dom: uuAx.
221  @ nat; #n; ncases n [ napply None | #m; ncases m [ napply (Some … O) | #_; napply (Twice … (pred m) (pred m)) ]
222 nqed.
223
224 ntheorem skipfact_base_dec:
225  memdec (uuax skipfact_dom) (mk_powerclass ? (λx: uuax skipfact_dom. x=O)).
226  nnormalize; @ (λx. match x with [ O ⇒ true | S _ ⇒ false ]); #n; nelim n;
227   nnormalize; //; #X; ndestruct; #Y; #Z; ndestruct; #W; ndestruct.
228 nqed.
229
230 ntheorem skipfact_partial:
231  ∀n: uuax skipfact_dom. two * n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x=O).
232  #n; nelim n
233   [ @1; nnormalize; @1
234   | #m; nelim m; nnormalize
235      [ #H; @2; nnormalize
236        [ //
237        | #y; *; #a; #E; nrewrite > E; ncases a; nnormalize; // ]
238    ##| #p; #H1; #H2; @2; nnormalize
239        [ //
240        | #y; *; #a; #E; nrewrite > E; ncases a; nnormalize;
241          nrewrite < (plus_n_Sm …); // ]##]
242 nqed.
243
244 ndefinition skipfact: ∀n:nat. n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x=O) → nat.
245  #n; #D; napply (cover_rect … skipfact_base_dec … n D)
246   [ #a; #_; napply natone
247   | #a; ncases a
248     [ nnormalize; #i; nelim i
249     | #m; ncases m
250        [ nnormalize; #_; #_; #H; napply H; @1
251        | #p; #i; nnormalize in i; #K;
252          #H; nnormalize in H;
253          napply (S m * H true * H false) ]
254 nqed.
255
256 nlemma test: skipfact four ? = four * two * two. ##[##2: napply (skipfact_partial two)]
257  nnormalize; //.
258 nqed.