]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/topology/igft.ma
9bff8502fba897a233caa3f7f8148c54a8fc190e
[helm.git] / helm / software / matita / nlibrary / topology / igft.ma
1 include "sets/sets.ma".
2
3 nrecord axiom_set : Type[1] ≝ { 
4   S:> Type[0];
5   I: S → Type[0];
6   C: ∀a:S. I a → Ω ^ S
7 }.
8
9 ndefinition cover_set ≝ λc:∀A:axiom_set.Ω^A → A → CProp[0].λA,C,U.
10   ∀y.y ∈ C → c A U y.
11
12 (* a \ltri b *)
13 notation "hvbox(a break ◃ b)" non associative with precedence 45
14 for @{ 'covers $a $b }.
15
16 interpretation "covers set temp" 'covers C U = (cover_set ?? C U).
17
18 ninductive cover (A : axiom_set) (U : Ω^A) : A → CProp[0] ≝ 
19 | creflexivity : ∀a:A. a ∈ U → cover ? U a
20 | cinfinity    : ∀a:A. ∀i:I ? a. (C ? a i ◃ U) → cover ? U a.
21 napply cover;
22 nqed.
23
24 interpretation "covers" 'covers a U = (cover ? U a).
25 interpretation "covers set" 'covers a U = (cover_set cover ? a U).
26
27 ndefinition fish_set ≝ λf:∀A:axiom_set.Ω^A → A → CProp[0].
28  λA,U,V.
29   ∃a.a ∈ V ∧ f A U a.
30
31 (* a \ltimes b *)
32 notation "hvbox(a break ⋉ b)" non associative with precedence 45
33 for @{ 'fish $a $b }. 
34
35 interpretation "fish set temp" 'fish A U = (fish_set ?? U A).
36
37 ncoinductive fish (A : axiom_set) (F : Ω^A) : A → CProp[0] ≝ 
38 | cfish : ∀a. a ∈ F → (∀i:I ? a.C A a i ⋉ F) → fish A F a.
39 napply fish;
40 nqed.
41
42 interpretation "fish set" 'fish A U = (fish_set fish ? U A).
43 interpretation "fish" 'fish a U = (fish ? U a).
44
45 nlet corec fish_rec (A:axiom_set) (U: Ω^A)
46  (P: Ω^A) (H1: P ⊆ U)
47   (H2: ∀a:A. a ∈ P → ∀j: I ? a. C ? a j ≬ P):
48    ∀a:A. ∀p: a ∈ P. a ⋉ U ≝ ?.
49 #a; #p; napply cfish;
50 ##[ napply H1; napply p;
51 ##| #i; ncases (H2 a p i); #x; *; #xC; #xP; @; ##[napply x]
52     @; ##[ napply xC ] napply (fish_rec ? U P); nassumption;
53 ##]
54 nqed.
55
56 notation "◃U" non associative with precedence 55
57 for @{ 'coverage $U }.
58
59 ndefinition coverage : ∀A:axiom_set.∀U:Ω^A.Ω^A ≝ λA,U.{ a | a ◃ U }.
60
61 interpretation "coverage cover" 'coverage U = (coverage ? U).
62
63 ndefinition cover_equation : ∀A:axiom_set.∀U,X:Ω^A.CProp[0] ≝  λA,U,X. 
64   ∀a.a ∈ X ↔ (a ∈ U ∨ ∃i:I ? a.∀y.y ∈ C ? a i → y ∈ X).  
65
66 ntheorem coverage_cover_equation : ∀A,U. cover_equation A U (◃U).
67 #A; #U; #a; @; #H;
68 ##[ nelim H; #b; (* manca clear *)
69     ##[ #bU; @1; nassumption;
70     ##| #i; #CaiU; #IH; @2; @ i; #c; #cCbi; ncases (IH ? cCbi);
71         ##[ #E; @; napply E;
72         ##| #_; napply CaiU; nassumption; ##] ##]
73 ##| ncases H; ##[ #E; @; nassumption]
74     *; #j; #Hj; @2 j; #w; #wC; napply Hj; nassumption;
75 ##]
76 nqed. 
77
78 ntheorem coverage_min_cover_equation : ∀A,U,W. cover_equation A U W → ◃U ⊆ W.
79 #A; #U; #W; #H; #a; #aU; nelim aU; #b;
80 ##[ #bU; ncases (H b); #_; #H1; napply H1; @1; nassumption;
81 ##| #i; #CbiU; #IH; ncases (H b); #_; #H1; napply H1; @2; @i; napply IH;
82 ##]
83 nqed.
84
85 notation "⋉F" non associative with precedence 55
86 for @{ 'fished $F }.
87
88 ndefinition fished : ∀A:axiom_set.∀F:Ω^A.Ω^A ≝ λA,F.{ a | a ⋉ F }.
89
90 interpretation "fished fish" 'fished F = (fished ? F).
91
92 ndefinition fish_equation : ∀A:axiom_set.∀F,X:Ω^A.CProp[0] ≝ λA,F,X.
93   ∀a. a ∈ X ↔ a ∈ F ∧ ∀i:I ? a.∃y.y ∈ C ? a i ∧ y ∈ X. 
94   
95 ntheorem fised_fish_equation : ∀A,F. fish_equation A F (⋉F).
96 #A; #F; #a; @; (* bug, fare case sotto diverso da farlo sopra *) #H; ncases H;
97 ##[ #b; #bF; #H2; @ bF; #i; ncases (H2 i); #c; *; #cC; #cF; @c; @ cC;
98     napply cF;  
99 ##| #aF; #H1; @ aF; napply H1;
100 ##]
101 nqed.
102
103 ntheorem fised_min_fish_equation : ∀A,F,G. fish_equation A F G → G ⊆ ⋉F.
104 #A; #F; #G; #H; #a; #aG; napply (fish_rec … aG);
105 #b; ncases (H b); #H1; #_; #bG; ncases (H1 bG); #E1; #E2; nassumption; 
106 nqed. 
107
108 STOP
109
110 (*
111 alias symbol "covers" (instance 0) = "covers".
112 alias symbol "covers" (instance 2) = "covers".
113 alias symbol "covers" (instance 1) = "covers set".
114 ntheorem covers_elim2:
115  ∀A: axiom_set. ∀U:Ω^A.∀P: A → CProp[0].
116   (∀a:A. a ∈ U → P a) →
117    (∀a:A.∀V:Ω^A. a ◃ V → V ◃ U → (∀y. y ∈ V → P y) → P a) →
118      ∀a:A. a ◃ U → P a.
119 #A; #U; #P; #H1; #H2; #a; #aU; nelim aU;
120 ##[ #b; #H; napply H1; napply H;
121 ##| #b; #i; #CaiU; #H; napply H2; 
122     ##[ napply (C ? b i);
123     ##| napply cinfinity; ##[ napply i ] nwhd; #y; #H3; napply creflexivity; ##]
124     nassumption; 
125 ##]
126 nqed.
127 *)
128
129 alias symbol "fish" (instance 1) = "fish set".
130 alias symbol "covers" = "covers".
131 ntheorem compatibility: ∀A:axiom_set.∀a:A.∀U,V. a ⋉ V → a ◃ U → U ⋉ V.
132 #A; #a; #U; #V; #aV; #aU; ngeneralize in match aV; (* clear aV *)
133 nelim aU;
134 ##[ #b; #bU; #bV; @; ##[ napply b] @; nassumption;
135 ##| #b; #i; #CaiU; #H; #bV; ncases bV in i CaiU H;
136     #c; #cV; #CciV; #i; #CciU; #H; ncases (CciV i); #x; *; #xCci; #xV;
137     napply H; nassumption;
138 ##]
139 nqed.
140
141
142
143
144 STOP
145
146 definition leq ≝ λA:axiom_set.λa,b:A. a ◃ {y|b=y}.
147
148 interpretation "covered by one" 'leq a b = (leq ? a b).
149
150 theorem leq_refl: ∀A:axiom_set.∀a:A. a ≤ a.
151  intros;
152  apply refl;
153  normalize;
154  reflexivity.
155 qed.
156
157 theorem leq_trans: ∀A:axiom_set.∀a,b,c:A. a ≤ b → b ≤ c → a ≤ c.
158  intros;
159  unfold in H H1 ⊢ %;
160  apply (transitivity ???? H);
161  constructor 1;
162  intros;
163  normalize in H2;
164  rewrite < H2;
165  assumption.
166 qed.
167
168 definition uparrow ≝ λA:axiom_set.λa:A.mk_powerset ? (λb:A. a ≤ b).
169
170 interpretation "uparrow" 'uparrow a = (uparrow ? a).
171
172 definition downarrow ≝ λA:axiom_set.λU:Ω \sup A.mk_powerset ? (λa:A. (↑a) ≬ U).
173
174 interpretation "downarrow" 'downarrow a = (downarrow ? a).
175
176 definition fintersects ≝ λA:axiom_set.λU,V:Ω \sup A.↓U ∩ ↓V.
177
178 interpretation "fintersects" 'fintersects U V = (fintersects ? U V).
179
180 record convergent_generated_topology : Type ≝
181  { AA:> axiom_set;
182    convergence: ∀a:AA.∀U,V:Ω \sup AA. a ◃ U → a ◃ V → a ◃ (U ↓ V)
183  }.