author | wenzelm |
Wed, 23 Jan 2019 23:12:40 +0100 | |
changeset 69729 | 4591221824f6 |
parent 69275 | 9bbd5497befd |
child 69745 | aec42cee2521 |
permissions | -rw-r--r-- |
65553 | 1 |
section \<open>Main HOL\<close> |
2 |
||
3 |
text \<open> |
|
4 |
Classical Higher-order Logic -- only ``Main'', excluding real and |
|
5 |
complex numbers etc. |
|
6 |
\<close> |
|
7 |
||
8 |
theory Main |
|
66817 | 9 |
imports |
10 |
Predicate_Compile |
|
11 |
Quickcheck_Narrowing |
|
12 |
Extraction |
|
13 |
Nunchaku |
|
69275 | 14 |
BNF_Greatest_Fixpoint |
15 |
Filter |
|
66817 | 16 |
Conditionally_Complete_Lattices |
17 |
Binomial |
|
18 |
GCD |
|
65553 | 19 |
begin |
20 |
||
69275 | 21 |
text \<open>Legacy\<close> |
22 |
||
23 |
context Inf |
|
24 |
begin |
|
25 |
||
26 |
abbreviation (input) INFIMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
27 |
where "INFIMUM A f \<equiv> \<Sqinter>(f ` A)" |
|
28 |
||
29 |
end |
|
30 |
||
31 |
context Sup |
|
32 |
begin |
|
33 |
||
34 |
abbreviation (input) SUPREMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
35 |
where "SUPREMUM A f \<equiv> \<Squnion>(f ` A)" |
|
36 |
||
37 |
end |
|
38 |
||
39 |
abbreviation (input) INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" |
|
40 |
where "INTER \<equiv> INFIMUM" |
|
41 |
||
42 |
abbreviation (input) UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" |
|
43 |
where "UNION \<equiv> SUPREMUM" |
|
44 |
||
45 |
||
46 |
text \<open>Namespace cleanup\<close> |
|
47 |
||
48 |
hide_const (open) |
|
49 |
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect |
|
50 |
fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift |
|
51 |
shift proj id_bnf |
|
52 |
||
53 |
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV |
|
54 |
||
55 |
||
56 |
text \<open>Syntax cleanup\<close> |
|
57 |
||
65814 | 58 |
no_notation |
59 |
bot ("\<bottom>") and |
|
60 |
top ("\<top>") and |
|
61 |
inf (infixl "\<sqinter>" 70) and |
|
62 |
sup (infixl "\<squnion>" 65) and |
|
68980
5717fbc55521
added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents:
67385
diff
changeset
|
63 |
Inf ("\<Sqinter> _" [900] 900) and |
5717fbc55521
added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents:
67385
diff
changeset
|
64 |
Sup ("\<Squnion> _" [900] 900) and |
65814 | 65 |
ordLeq2 (infix "<=o" 50) and |
66 |
ordLeq3 (infix "\<le>o" 50) and |
|
67 |
ordLess2 (infix "<o" 50) and |
|
68 |
ordIso2 (infix "=o" 50) and |
|
69 |
card_of ("|_|") and |
|
69275 | 70 |
BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and |
71 |
BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and |
|
72 |
BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90) and |
|
73 |
BNF_Def.convol ("\<langle>(_,/ _)\<rangle>") |
|
65814 | 74 |
|
75 |
no_syntax |
|
76 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
|
77 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
|
78 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
|
79 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
|
80 |
||
65553 | 81 |
end |