author | wenzelm |
Tue, 02 Jun 2015 09:10:05 +0200 | |
changeset 60357 | bc0827281dc1 |
parent 60036 | 218fcc645d22 |
child 60758 | d8d85a8172b5 |
permissions | -rw-r--r-- |
58889 | 1 |
section {* Main HOL *} |
12024 | 2 |
|
15131 | 3 |
theory Main |
60036 | 4 |
imports Predicate_Compile Quickcheck_Narrowing Extraction Coinduction Nitpick BNF_Greatest_Fixpoint Filter |
15131 | 5 |
begin |
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
6 |
|
29304 | 7 |
text {* |
8 |
Classical Higher-order Logic -- only ``Main'', excluding real and |
|
9 |
complex numbers etc. |
|
10 |
*} |
|
11 |
||
58623 | 12 |
text {* See further @{cite "Nipkow-et-al:2002:tutorial"} *} |
25964 | 13 |
|
51112 | 14 |
no_notation |
15 |
bot ("\<bottom>") and |
|
16 |
top ("\<top>") and |
|
17 |
inf (infixl "\<sqinter>" 70) and |
|
18 |
sup (infixl "\<squnion>" 65) and |
|
19 |
Inf ("\<Sqinter>_" [900] 900) and |
|
55057 | 20 |
Sup ("\<Squnion>_" [900] 900) and |
55065 | 21 |
ordLeq2 (infix "<=o" 50) and |
22 |
ordLeq3 (infix "\<le>o" 50) and |
|
23 |
ordLess2 (infix "<o" 50) and |
|
24 |
ordIso2 (infix "=o" 50) and |
|
55078 | 25 |
card_of ("|_|") and |
55057 | 26 |
csum (infixr "+c" 65) and |
27 |
cprod (infixr "*c" 80) and |
|
55086 | 28 |
cexp (infixr "^c" 90) and |
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57208
diff
changeset
|
29 |
convol ("\<langle>(_,/ _)\<rangle>") |
51112 | 30 |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55086
diff
changeset
|
31 |
hide_const (open) |
58152 | 32 |
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect |
33 |
fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift |
|
58353 | 34 |
shift proj id_bnf |
35 |
||
36 |
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55086
diff
changeset
|
37 |
|
51112 | 38 |
no_syntax (xsymbols) |
39 |
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10) |
|
40 |
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10) |
|
41 |
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10) |
|
42 |
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10) |
|
43 |
||
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
44 |
end |