| author | paulson <lp15@cam.ac.uk> | 
| Fri, 28 Feb 2025 13:50:18 +0000 | |
| changeset 82218 | cbf9f856d3e0 | 
| parent 81125 | ec121999a9cb | 
| 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  | 
|
| 66817 | 9  | 
imports  | 
10  | 
Predicate_Compile  | 
|
| 70078 | 11  | 
Quickcheck_Narrowing  | 
| 
73691
 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 
wenzelm 
parents: 
70336 
diff
changeset
 | 
12  | 
Mirabelle  | 
| 66817 | 13  | 
Extraction  | 
14  | 
Nunchaku  | 
|
| 69275 | 15  | 
BNF_Greatest_Fixpoint  | 
16  | 
Filter  | 
|
| 66817 | 17  | 
Conditionally_Complete_Lattices  | 
18  | 
Binomial  | 
|
19  | 
GCD  | 
|
| 65553 | 20  | 
begin  | 
21  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
22  | 
subsection \<open>Namespace cleanup\<close>  | 
| 69275 | 23  | 
|
24  | 
hide_const (open)  | 
|
25  | 
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect  | 
|
26  | 
fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift  | 
|
27  | 
shift proj id_bnf  | 
|
28  | 
||
29  | 
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV  | 
|
30  | 
||
31  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
32  | 
subsection \<open>Syntax cleanup\<close>  | 
| 69275 | 33  | 
|
| 65814 | 34  | 
no_notation  | 
| 
80932
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
35  | 
ordLeq2 (infix \<open><=o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
36  | 
ordLeq3 (infix \<open>\<le>o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
37  | 
ordLess2 (infix \<open><o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
38  | 
ordIso2 (infix \<open>=o\<close> 50) and  | 
| 81125 | 39  | 
card_of (\<open>(\<open>open_block notation=\<open>mixfix card_of\<close>\<close>|_|)\<close>) and  | 
| 
80932
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
40  | 
BNF_Cardinal_Arithmetic.csum (infixr \<open>+c\<close> 65) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
41  | 
BNF_Cardinal_Arithmetic.cprod (infixr \<open>*c\<close> 80) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
42  | 
BNF_Cardinal_Arithmetic.cexp (infixr \<open>^c\<close> 90) and  | 
| 81125 | 43  | 
BNF_Def.convol (\<open>(\<open>indent=1 notation=\<open>mixfix convol\<close>\<close>\<langle>_,/ _\<rangle>)\<close>)  | 
| 65814 | 44  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
45  | 
bundle cardinal_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
46  | 
begin  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
47  | 
|
| 70078 | 48  | 
notation  | 
| 
80932
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
49  | 
ordLeq2 (infix \<open><=o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
50  | 
ordLeq3 (infix \<open>\<le>o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
51  | 
ordLess2 (infix \<open><o\<close> 50) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
52  | 
ordIso2 (infix \<open>=o\<close> 50) and  | 
| 81125 | 53  | 
card_of (\<open>(\<open>open_block notation=\<open>mixfix card_of\<close>\<close>|_|)\<close>) and  | 
| 
80932
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
54  | 
BNF_Cardinal_Arithmetic.csum (infixr \<open>+c\<close> 65) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
55  | 
BNF_Cardinal_Arithmetic.cprod (infixr \<open>*c\<close> 80) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
56  | 
BNF_Cardinal_Arithmetic.cexp (infixr \<open>^c\<close> 90)  | 
| 70078 | 57  | 
|
58  | 
alias cinfinite = BNF_Cardinal_Arithmetic.cinfinite  | 
|
59  | 
alias czero = BNF_Cardinal_Arithmetic.czero  | 
|
60  | 
alias cone = BNF_Cardinal_Arithmetic.cone  | 
|
61  | 
alias ctwo = BNF_Cardinal_Arithmetic.ctwo  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
62  | 
|
| 70078 | 63  | 
end  | 
64  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
65  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
66  | 
subsection \<open>Lattice syntax\<close>  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
67  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
68  | 
bundle lattice_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
69  | 
begin  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
70  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
71  | 
notation  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
80934 
diff
changeset
 | 
72  | 
bot (\<open>\<bottom>\<close>) and  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
80934 
diff
changeset
 | 
73  | 
top (\<open>\<top>\<close>) and  | 
| 
80932
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
74  | 
inf (infixl \<open>\<sqinter>\<close> 70) and  | 
| 
 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 
wenzelm 
parents: 
80453 
diff
changeset
 | 
75  | 
sup (infixl \<open>\<squnion>\<close> 65) and  | 
| 81125 | 76  | 
Inf (\<open>(\<open>open_block notation=\<open>prefix \<Sqinter>\<close>\<close>\<Sqinter> _)\<close> [900] 900) and  | 
77  | 
Sup (\<open>(\<open>open_block notation=\<open>prefix \<Squnion>\<close>\<close>\<Squnion> _)\<close> [900] 900)  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
78  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
79  | 
syntax  | 
| 80934 | 80  | 
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Sqinter>\<close>\<close>\<Sqinter>_./ _)\<close> [0, 10] 10)  | 
81  | 
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Sqinter>\<close>\<close>\<Sqinter>_\<in>_./ _)\<close> [0, 0, 10] 10)  | 
|
82  | 
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Squnion>\<close>\<close>\<Squnion>_./ _)\<close> [0, 10] 10)  | 
|
83  | 
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Squnion>\<close>\<close>\<Squnion>_\<in>_./ _)\<close> [0, 0, 10] 10)  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
84  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
85  | 
end  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
86  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
80934 
diff
changeset
 | 
87  | 
unbundle no lattice_syntax  | 
| 65814 | 88  | 
|
| 65553 | 89  | 
end  |