author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 66614 | 1f1c5d85d232 |
child 66817 | 0b12755ccbb2 |
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 |
|
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
65814
diff
changeset
|
9 |
imports Predicate_Compile Quickcheck_Narrowing Extraction Nunchaku BNF_Greatest_Fixpoint Filter Conditionally_Complete_Lattices Binomial GCD |
65553 | 10 |
begin |
11 |
||
65814 | 12 |
text \<open> |
13 |
Classical Higher-order Logic -- only ``Main'', excluding real and |
|
14 |
complex numbers etc. |
|
15 |
\<close> |
|
16 |
||
17 |
no_notation |
|
18 |
bot ("\<bottom>") and |
|
19 |
top ("\<top>") and |
|
20 |
inf (infixl "\<sqinter>" 70) and |
|
21 |
sup (infixl "\<squnion>" 65) and |
|
22 |
Inf ("\<Sqinter>_" [900] 900) and |
|
23 |
Sup ("\<Squnion>_" [900] 900) and |
|
24 |
ordLeq2 (infix "<=o" 50) and |
|
25 |
ordLeq3 (infix "\<le>o" 50) and |
|
26 |
ordLess2 (infix "<o" 50) and |
|
27 |
ordIso2 (infix "=o" 50) and |
|
28 |
card_of ("|_|") and |
|
29 |
csum (infixr "+c" 65) and |
|
30 |
cprod (infixr "*c" 80) and |
|
31 |
cexp (infixr "^c" 90) and |
|
32 |
convol ("\<langle>(_,/ _)\<rangle>") |
|
33 |
||
34 |
hide_const (open) |
|
35 |
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect |
|
36 |
fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift |
|
37 |
shift proj id_bnf |
|
38 |
||
39 |
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV |
|
40 |
||
41 |
no_syntax |
|
42 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
|
43 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
|
44 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
|
45 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
|
46 |
||
65553 | 47 |
end |