| 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
 | 
| 70078 |     11 |     Quickcheck_Narrowing
 | 
| 66817 |     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>Namespace cleanup\<close>
 | 
|  |     22 | 
 | 
|  |     23 | hide_const (open)
 | 
|  |     24 |   czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect
 | 
|  |     25 |   fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift
 | 
|  |     26 |   shift proj id_bnf
 | 
|  |     27 | 
 | 
|  |     28 | hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV
 | 
|  |     29 | 
 | 
|  |     30 | 
 | 
|  |     31 | text \<open>Syntax cleanup\<close>
 | 
|  |     32 | 
 | 
| 65814 |     33 | no_notation
 | 
|  |     34 |   bot ("\<bottom>") and
 | 
|  |     35 |   top ("\<top>") and
 | 
|  |     36 |   inf (infixl "\<sqinter>" 70) and
 | 
|  |     37 |   sup (infixl "\<squnion>" 65) and
 | 
| 69745 |     38 |   Inf ("\<Sqinter>") and
 | 
|  |     39 |   Sup ("\<Squnion>") and
 | 
| 65814 |     40 |   ordLeq2 (infix "<=o" 50) and
 | 
|  |     41 |   ordLeq3 (infix "\<le>o" 50) and
 | 
|  |     42 |   ordLess2 (infix "<o" 50) and
 | 
|  |     43 |   ordIso2 (infix "=o" 50) and
 | 
|  |     44 |   card_of ("|_|") and
 | 
| 69275 |     45 |   BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and
 | 
|  |     46 |   BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and
 | 
|  |     47 |   BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90) and
 | 
|  |     48 |   BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
 | 
| 65814 |     49 | 
 | 
| 70078 |     50 | bundle cardinal_syntax begin
 | 
|  |     51 | notation
 | 
|  |     52 |   ordLeq2 (infix "<=o" 50) and
 | 
|  |     53 |   ordLeq3 (infix "\<le>o" 50) and
 | 
|  |     54 |   ordLess2 (infix "<o" 50) and
 | 
|  |     55 |   ordIso2 (infix "=o" 50) and
 | 
|  |     56 |   card_of ("|_|") and
 | 
|  |     57 |   BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and
 | 
|  |     58 |   BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and
 | 
|  |     59 |   BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90)
 | 
|  |     60 | 
 | 
|  |     61 | alias cinfinite = BNF_Cardinal_Arithmetic.cinfinite
 | 
|  |     62 | alias czero = BNF_Cardinal_Arithmetic.czero
 | 
|  |     63 | alias cone = BNF_Cardinal_Arithmetic.cone
 | 
|  |     64 | alias ctwo = BNF_Cardinal_Arithmetic.ctwo
 | 
|  |     65 | end
 | 
|  |     66 | 
 | 
| 65814 |     67 | no_syntax
 | 
|  |     68 |   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
|  |     69 |   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | 
|  |     70 |   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | 
|  |     71 |   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | 
|  |     72 | 
 | 
| 65553 |     73 | end
 |