author | blanchet |
Mon Feb 17 13:31:42 2014 +0100 (2014-02-17) | |
changeset 55534 | b18bdcbda41b |
parent 55128 | 6e16d2dd4f14 |
child 55579 | 207538943038 |
permissions | -rw-r--r-- |
wenzelm@12024 | 1 |
header {* Main HOL *} |
wenzelm@12024 | 2 |
|
nipkow@15131 | 3 |
theory Main |
blanchet@55534 | 4 |
imports Predicate_Compile Extraction Lifting_Sum List_Prefix Coinduction Nitpick BNF_GFP |
nipkow@15131 | 5 |
begin |
wenzelm@9650 | 6 |
|
wenzelm@29304 | 7 |
text {* |
wenzelm@29304 | 8 |
Classical Higher-order Logic -- only ``Main'', excluding real and |
wenzelm@29304 | 9 |
complex numbers etc. |
wenzelm@29304 | 10 |
*} |
wenzelm@29304 | 11 |
|
haftmann@27367 | 12 |
text {* See further \cite{Nipkow-et-al:2002:tutorial} *} |
haftmann@25964 | 13 |
|
haftmann@51112 | 14 |
no_notation |
haftmann@51112 | 15 |
bot ("\<bottom>") and |
haftmann@51112 | 16 |
top ("\<top>") and |
haftmann@51112 | 17 |
inf (infixl "\<sqinter>" 70) and |
haftmann@51112 | 18 |
sup (infixl "\<squnion>" 65) and |
haftmann@51112 | 19 |
Inf ("\<Sqinter>_" [900] 900) and |
blanchet@55057 | 20 |
Sup ("\<Squnion>_" [900] 900) and |
blanchet@55065 | 21 |
ordLeq2 (infix "<=o" 50) and |
blanchet@55065 | 22 |
ordLeq3 (infix "\<le>o" 50) and |
blanchet@55065 | 23 |
ordLess2 (infix "<o" 50) and |
blanchet@55065 | 24 |
ordIso2 (infix "=o" 50) and |
blanchet@55078 | 25 |
card_of ("|_|") and |
blanchet@55057 | 26 |
csum (infixr "+c" 65) and |
blanchet@55057 | 27 |
cprod (infixr "*c" 80) and |
blanchet@55086 | 28 |
cexp (infixr "^c" 90) and |
blanchet@55086 | 29 |
convol ("<_ , _>") |
haftmann@51112 | 30 |
|
blanchet@55087 | 31 |
hide_const (open) |
blanchet@55128 | 32 |
czero cinfinite cfinite csum cone ctwo Csum cprod cexp |
blanchet@55087 | 33 |
image2 image2p vimage2p Gr Grp collect fsts snds setl setr |
blanchet@55087 | 34 |
convol pick_middlep fstOp sndOp csquare inver relImage relInvImage |
blanchet@55087 | 35 |
prefCl PrefCl Succ Shift shift proj |
blanchet@55087 | 36 |
|
haftmann@51112 | 37 |
no_syntax (xsymbols) |
haftmann@51112 | 38 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
haftmann@51112 | 39 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
haftmann@51112 | 40 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
haftmann@51112 | 41 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
haftmann@51112 | 42 |
|
wenzelm@9650 | 43 |
end |