C: âa:S. I a â Ί ^ S
}.
+notation "đ term 90 a" non associative with precedence 70 for @{ 'I $a }.
+notation "đ term 90 a term 90 i" non associative with precedence 70 for @{ 'C $a $i }.
+
+interpretation "I" 'I a = (I ? a).
+interpretation "C" 'C a i = (C ? a i).
+
ndefinition cover_set â Îťc:âA:axiom_set.Ί^A â A â CProp[0].ÎťA,C,U.
ây.y â C â c A U y.
ninductive cover (A : axiom_set) (U : Ί^A) : A â CProp[0] â
| creflexivity : âa:A. a â U â cover ? U a
-| cinfinity : âa:A. âi:I ? a. (C ? a i â U) â cover ? U a.
+| cinfinity : âa:A. âi:?. ((C ? a i) â U) â cover ? U a.
napply cover;
nqed.
interpretation "fish set temp" 'fish A U = (fish_set ?? U A).
ncoinductive fish (A : axiom_set) (F : Ί^A) : A â CProp[0] â
-| cfish : âa. a â F â (âi:I ? a.C A a i â F) â fish A F a.
+| cfish : âa. a â F â (âi:đ a.C A a i â F) â fish A F a.
napply fish;
nqed.
nlet corec fish_rec (A:axiom_set) (U: Ί^A)
(P: Ί^A) (H1: P â U)
- (H2: âa:A. a â P â âj: I ? a. C ? a j ⏠P):
+ (H2: âa:A. a â P â âj: đ a. C ? a j ⏠P):
âa:A. âp: a â P. a â U â ?.
#a; #p; napply cfish;
##[ napply H1; napply p;
interpretation "coverage cover" 'coverage U = (coverage ? U).
ndefinition cover_equation : âA:axiom_set.âU,X:Ί^A.CProp[0] â ÎťA,U,X.
- âa.a â X â (a â U ⨠âi:I ? a.ây.y â C ? a i â y â X).
+ âa.a â X â (a â U ⨠âi:đ a.ây.y â đ a i â y â X).
ntheorem coverage_cover_equation : âA,U. cover_equation A U (âU).
#A; #U; #a; @; #H;
interpretation "fished fish" 'fished F = (fished ? F).
ndefinition fish_equation : âA:axiom_set.âF,X:Ί^A.CProp[0] â ÎťA,F,X.
- âa. a â X â a â F ⧠âi:I ? a.ây.y â C ? a i ⧠y â X.
+ âa. a â X â a â F ⧠âi:đ a.ây.y â đ a i ⧠y â X.
ntheorem fised_fish_equation : âA,F. fish_equation A F (âF).
#A; #F; #a; @; (* bug, fare case sotto diverso da farlo sopra *) #H; ncases H;
##]
nqed.
-ntheorem fised_min_fish_equation : âA,F,G. fish_equation A F G â G â âF.
+ntheorem fised_max_fish_equation : âA,F,G. fish_equation A F G â G â âF.
#A; #F; #G; #H; #a; #aG; napply (fish_rec ⌠aG);
#b; ncases (H b); #H1; #_; #bG; ncases (H1 bG); #E1; #E2; nassumption;
nqed.
-STOP
-
-(*
-alias symbol "covers" (instance 0) = "covers".
-alias symbol "covers" (instance 2) = "covers".
-alias symbol "covers" (instance 1) = "covers set".
-ntheorem covers_elim2:
- âA: axiom_set. âU:Ί^A.âP: A â CProp[0].
- (âa:A. a â U â P a) â
- (âa:A.âV:Ί^A. a â V â V â U â (ây. y â V â P y) â P a) â
- âa:A. a â U â P a.
-#A; #U; #P; #H1; #H2; #a; #aU; nelim aU;
-##[ #b; #H; napply H1; napply H;
-##| #b; #i; #CaiU; #H; napply H2;
- ##[ napply (C ? b i);
- ##| napply cinfinity; ##[ napply i ] nwhd; #y; #H3; napply creflexivity; ##]
- nassumption;
-##]
-nqed.
-*)
-
-alias symbol "fish" (instance 1) = "fish set".
-alias symbol "covers" = "covers".
-ntheorem compatibility: âA:axiom_set.âa:A.âU,V. a â V â a â U â U â V.
-#A; #a; #U; #V; #aV; #aU; ngeneralize in match aV; (* clear aV *)
-nelim aU;
-##[ #b; #bU; #bV; @; ##[ napply b] @; nassumption;
-##| #b; #i; #CaiU; #H; #bV; ncases bV in i CaiU H;
- #c; #cV; #CciV; #i; #CciU; #H; ncases (CciV i); #x; *; #xCci; #xV;
- napply H; nassumption;
-##]
-nqed.
-
-
-
-
-STOP
-
-definition leq â ÎťA:axiom_set.Îťa,b:A. a â {y|b=y}.
-
-interpretation "covered by one" 'leq a b = (leq ? a b).
-
-theorem leq_refl: âA:axiom_set.âa:A. a ⤠a.
- intros;
- apply refl;
- normalize;
- reflexivity.
-qed.
-
-theorem leq_trans: âA:axiom_set.âa,b,c:A. a ⤠b â b ⤠c â a ⤠c.
- intros;
- unfold in H H1 ⢠%;
- apply (transitivity ???? H);
- constructor 1;
- intros;
- normalize in H2;
- rewrite < H2;
- assumption.
-qed.
-
-definition uparrow â ÎťA:axiom_set.Îťa:A.mk_powerset ? (Îťb:A. a ⤠b).
-
-interpretation "uparrow" 'uparrow a = (uparrow ? a).
-
-definition downarrow â ÎťA:axiom_set.ÎťU:Ί \sup A.mk_powerset ? (Îťa:A. (âa) ⏠U).
-
-interpretation "downarrow" 'downarrow a = (downarrow ? a).
-
-definition fintersects â ÎťA:axiom_set.ÎťU,V:Ί \sup A.âU ⊠âV.
-
-interpretation "fintersects" 'fintersects U V = (fintersects ? U V).
-
-record convergent_generated_topology : Type â
- { AA:> axiom_set;
- convergence: âa:AA.âU,V:Ί \sup AA. a â U â a â V â a â (U â V)
- }.