| author | nipkow | 
| Sat, 24 Feb 2024 11:29:30 +0100 | |
| changeset 79714 | 80cb54976c1c | 
| parent 76224 | 64e8d4afcf10 | 
| child 80453 | 7a2d9e3fcdd5 | 
| 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  | 
|
| 
76224
 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 
haftmann 
parents: 
74337 
diff
changeset
 | 
20  | 
Divides  | 
| 65553 | 21  | 
begin  | 
22  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
23  | 
subsection \<open>Namespace cleanup\<close>  | 
| 69275 | 24  | 
|
25  | 
hide_const (open)  | 
|
26  | 
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect  | 
|
27  | 
fsts snds setl setr convol pick_middlep fstOp sndOp csquare relImage relInvImage Succ Shift  | 
|
28  | 
shift proj id_bnf  | 
|
29  | 
||
30  | 
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV  | 
|
31  | 
||
32  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
33  | 
subsection \<open>Syntax cleanup\<close>  | 
| 69275 | 34  | 
|
| 65814 | 35  | 
no_notation  | 
36  | 
ordLeq2 (infix "<=o" 50) and  | 
|
37  | 
ordLeq3 (infix "\<le>o" 50) and  | 
|
38  | 
ordLess2 (infix "<o" 50) and  | 
|
39  | 
ordIso2 (infix "=o" 50) and  | 
|
40  | 
  card_of ("|_|") and
 | 
|
| 69275 | 41  | 
BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and  | 
42  | 
BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and  | 
|
43  | 
BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90) and  | 
|
44  | 
  BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
 | 
|
| 65814 | 45  | 
|
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
46  | 
bundle cardinal_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
47  | 
begin  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
48  | 
|
| 70078 | 49  | 
notation  | 
50  | 
ordLeq2 (infix "<=o" 50) and  | 
|
51  | 
ordLeq3 (infix "\<le>o" 50) and  | 
|
52  | 
ordLess2 (infix "<o" 50) and  | 
|
53  | 
ordIso2 (infix "=o" 50) and  | 
|
54  | 
  card_of ("|_|") and
 | 
|
55  | 
BNF_Cardinal_Arithmetic.csum (infixr "+c" 65) and  | 
|
56  | 
BNF_Cardinal_Arithmetic.cprod (infixr "*c" 80) and  | 
|
57  | 
BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90)  | 
|
58  | 
||
59  | 
alias cinfinite = BNF_Cardinal_Arithmetic.cinfinite  | 
|
60  | 
alias czero = BNF_Cardinal_Arithmetic.czero  | 
|
61  | 
alias cone = BNF_Cardinal_Arithmetic.cone  | 
|
62  | 
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
 | 
63  | 
|
| 70078 | 64  | 
end  | 
65  | 
||
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
66  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
67  | 
subsection \<open>Lattice syntax\<close>  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
68  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
69  | 
bundle lattice_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
70  | 
begin  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
71  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
72  | 
notation  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
73  | 
  bot ("\<bottom>") and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
74  | 
  top ("\<top>") and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
75  | 
inf (infixl "\<sqinter>" 70) and  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
76  | 
sup (infixl "\<squnion>" 65) and  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
77  | 
  Inf  ("\<Sqinter> _" [900] 900) and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
78  | 
  Sup  ("\<Squnion> _" [900] 900)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
79  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
80  | 
syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
81  | 
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
82  | 
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
83  | 
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
84  | 
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
85  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
86  | 
end  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
87  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
88  | 
bundle no_lattice_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
89  | 
begin  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
90  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
91  | 
no_notation  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
92  | 
  bot ("\<bottom>") and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
93  | 
  top ("\<top>") and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
94  | 
inf (infixl "\<sqinter>" 70) and  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
95  | 
sup (infixl "\<squnion>" 65) and  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
96  | 
  Inf  ("\<Sqinter> _" [900] 900) and
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
97  | 
  Sup  ("\<Squnion> _" [900] 900)
 | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
98  | 
|
| 65814 | 99  | 
no_syntax  | 
100  | 
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
|
101  | 
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | 
|
102  | 
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | 
|
103  | 
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | 
|
104  | 
||
| 65553 | 105  | 
end  | 
| 
74334
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
106  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
107  | 
unbundle no_lattice_syntax  | 
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
108  | 
|
| 
 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 
wenzelm 
parents: 
74102 
diff
changeset
 | 
109  | 
end  |