| 
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
  | 
| 
 | 
    14  | 
    BNF_Greatest_Fixpoint Filter
  | 
| 
 | 
    15  | 
    Conditionally_Complete_Lattices
  | 
| 
 | 
    16  | 
    Binomial
  | 
| 
 | 
    17  | 
    GCD
  | 
| 
65553
 | 
    18  | 
begin
  | 
| 
 | 
    19  | 
  | 
| 
65814
 | 
    20  | 
no_notation
  | 
| 
 | 
    21  | 
  bot ("\<bottom>") and
 | 
| 
 | 
    22  | 
  top ("\<top>") and
 | 
| 
 | 
    23  | 
  inf (infixl "\<sqinter>" 70) and
  | 
| 
 | 
    24  | 
  sup (infixl "\<squnion>" 65) and
  | 
| 
 | 
    25  | 
  Inf ("\<Sqinter>_" [900] 900) and
 | 
| 
 | 
    26  | 
  Sup ("\<Squnion>_" [900] 900) and
 | 
| 
 | 
    27  | 
  ordLeq2 (infix "<=o" 50) and
  | 
| 
 | 
    28  | 
  ordLeq3 (infix "\<le>o" 50) and
  | 
| 
 | 
    29  | 
  ordLess2 (infix "<o" 50) and
  | 
| 
 | 
    30  | 
  ordIso2 (infix "=o" 50) and
  | 
| 
 | 
    31  | 
  card_of ("|_|") and
 | 
| 
 | 
    32  | 
  csum (infixr "+c" 65) and
  | 
| 
 | 
    33  | 
  cprod (infixr "*c" 80) and
  | 
| 
 | 
    34  | 
  cexp (infixr "^c" 90) and
  | 
| 
 | 
    35  | 
  convol ("\<langle>(_,/ _)\<rangle>")
 | 
| 
 | 
    36  | 
  | 
| 
 | 
    37  | 
hide_const (open)
  | 
| 
 | 
    38  | 
  czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect
  | 
| 
 | 
    39  | 
  fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift
  | 
| 
 | 
    40  | 
  shift proj id_bnf
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV
  | 
| 
 | 
    43  | 
  | 
| 
 | 
    44  | 
no_syntax
  | 
| 
 | 
    45  | 
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
| 
 | 
    46  | 
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
 | 
    47  | 
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | 
| 
 | 
    48  | 
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
 | 
    49  | 
  | 
| 
65553
 | 
    50  | 
end
  |