| author | desharna | 
| Wed, 12 Mar 2025 19:26:59 +0100 | |
| changeset 82249 | bdefffffd05f | 
| parent 80453 | 7a2d9e3fcdd5 | 
| permissions | -rw-r--r-- | 
| 30293 | 1 | (*<*) | 
| 30401 | 2 | theory Main_Doc | 
| 30293 | 3 | imports Main | 
| 4 | begin | |
| 5 | ||
| 61996 | 6 | setup \<open> | 
| 73761 | 7 | Document_Output.antiquotation_pretty_source \<^binding>\<open>term_type_only\<close> (Args.term -- Args.typ_abbrev) | 
| 67463 | 8 | (fn ctxt => fn (t, T) => | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42361diff
changeset | 9 | (if fastype_of t = Sign.certify_typ (Proof_Context.theory_of ctxt) T then () | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42361diff
changeset | 10 | else error "term_type_only: type mismatch"; | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67463diff
changeset | 11 | Syntax.pretty_typ ctxt T)) | 
| 61996 | 12 | \<close> | 
| 13 | setup \<open> | |
| 73761 | 14 | Document_Output.antiquotation_pretty_source \<^binding>\<open>expanded_typ\<close> Args.typ | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67463diff
changeset | 15 | Syntax.pretty_typ | 
| 61996 | 16 | \<close> | 
| 30293 | 17 | (*>*) | 
| 61996 | 18 | text\<open> | 
| 30293 | 19 | |
| 20 | \begin{abstract}
 | |
| 74720 | 21 | This document lists the main types, functions and syntax provided by theory \<^theory>\<open>Main\<close>. It is meant as a quick overview of what is available. For infix operators and their precedences see the final section. The sophisticated class structure is only hinted at. For details see \<^url>\<open>https://isabelle.in.tum.de/library/HOL/HOL\<close>. | 
| 30293 | 22 | \end{abstract}
 | 
| 23 | ||
| 50581 | 24 | \section*{HOL}
 | 
| 30293 | 25 | |
| 69597 | 26 | The basic logic: \<^prop>\<open>x = y\<close>, \<^const>\<open>True\<close>, \<^const>\<open>False\<close>, \<^prop>\<open>\<not> P\<close>, \<^prop>\<open>P \<and> Q\<close>, | 
| 27 | \<^prop>\<open>P \<or> Q\<close>, \<^prop>\<open>P \<longrightarrow> Q\<close>, \<^prop>\<open>\<forall>x. P\<close>, \<^prop>\<open>\<exists>x. P\<close>, \<^prop>\<open>\<exists>! x. P\<close>, | |
| 28 | \<^term>\<open>THE x. P\<close>. | |
| 61996 | 29 | \<^smallskip> | 
| 30440 | 30 | |
| 31 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 32 | \<^const>\<open>HOL.undefined\<close> & \<^typeof>\<open>HOL.undefined\<close>\\ | 
| 33 | \<^const>\<open>HOL.default\<close> & \<^typeof>\<open>HOL.default\<close>\\ | |
| 30440 | 34 | \end{tabular}
 | 
| 35 | ||
| 36 | \subsubsection*{Syntax}
 | |
| 30293 | 37 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 38 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 39 | \<^term>\<open>\<not> (x = y)\<close> & @{term[source]"\<not> (x = y)"} & (\<^verbatim>\<open>~=\<close>)\\
 | 
| 40 | @{term[source]"P \<longleftrightarrow> Q"} & \<^term>\<open>P \<longleftrightarrow> Q\<close> \\
 | |
| 41 | \<^term>\<open>If x y z\<close> & @{term[source]"If x y z"}\\
 | |
| 42 | \<^term>\<open>Let e\<^sub>1 (\<lambda>x. e\<^sub>2)\<close> & @{term[source]"Let e\<^sub>1 (\<lambda>x. e\<^sub>2)"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 43 | \end{tabular}
 | 
| 30440 | 44 | |
| 45 | ||
| 50581 | 46 | \section*{Orderings}
 | 
| 30440 | 47 | |
| 48 | A collection of classes defining basic orderings: | |
| 49 | preorder, partial order, linear order, dense linear order and wellorder. | |
| 61996 | 50 | \<^smallskip> | 
| 30293 | 51 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 52 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 53 | \<^const>\<open>Orderings.less_eq\<close> & \<^typeof>\<open>Orderings.less_eq\<close> & (\<^verbatim>\<open><=\<close>)\\ | 
| 54 | \<^const>\<open>Orderings.less\<close> & \<^typeof>\<open>Orderings.less\<close>\\ | |
| 55 | \<^const>\<open>Orderings.Least\<close> & \<^typeof>\<open>Orderings.Least\<close>\\ | |
| 56 | \<^const>\<open>Orderings.Greatest\<close> & \<^typeof>\<open>Orderings.Greatest\<close>\\ | |
| 57 | \<^const>\<open>Orderings.min\<close> & \<^typeof>\<open>Orderings.min\<close>\\ | |
| 58 | \<^const>\<open>Orderings.max\<close> & \<^typeof>\<open>Orderings.max\<close>\\ | |
| 59 | @{const[source] top} & \<^typeof>\<open>Orderings.top\<close>\\
 | |
| 60 | @{const[source] bot} & \<^typeof>\<open>Orderings.bot\<close>\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 61 | \end{tabular}
 | 
| 30293 | 62 | |
| 63 | \subsubsection*{Syntax}
 | |
| 64 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 65 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 66 | @{term[source]"x \<ge> y"} & \<^term>\<open>x \<ge> y\<close> & (\<^verbatim>\<open>>=\<close>)\\
 | 
| 67 | @{term[source]"x > y"} & \<^term>\<open>x > y\<close>\\
 | |
| 68 | \<^term>\<open>\<forall>x\<le>y. P\<close> & @{term[source]"\<forall>x. x \<le> y \<longrightarrow> P"}\\
 | |
| 69 | \<^term>\<open>\<exists>x\<le>y. P\<close> & @{term[source]"\<exists>x. x \<le> y \<and> P"}\\
 | |
| 30440 | 70 | \multicolumn{2}{@ {}l@ {}}{Similarly for $<$, $\ge$ and $>$}\\
 | 
| 69597 | 71 | \<^term>\<open>LEAST x. P\<close> & @{term[source]"Least (\<lambda>x. P)"}\\
 | 
| 72 | \<^term>\<open>GREATEST x. P\<close> & @{term[source]"Greatest (\<lambda>x. P)"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 73 | \end{tabular}
 | 
| 30293 | 74 | |
| 30401 | 75 | |
| 50581 | 76 | \section*{Lattices}
 | 
| 30401 | 77 | |
| 78 | Classes semilattice, lattice, distributive lattice and complete lattice (the | |
| 69597 | 79 | latter in theory \<^theory>\<open>HOL.Set\<close>). | 
| 30401 | 80 | |
| 81 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 82 | \<^const>\<open>Lattices.inf\<close> & \<^typeof>\<open>Lattices.inf\<close>\\ | 
| 83 | \<^const>\<open>Lattices.sup\<close> & \<^typeof>\<open>Lattices.sup\<close>\\ | |
| 84 | \<^const>\<open>Complete_Lattices.Inf\<close> & @{term_type_only Complete_Lattices.Inf "'a set \<Rightarrow> 'a::Inf"}\\
 | |
| 85 | \<^const>\<open>Complete_Lattices.Sup\<close> & @{term_type_only Complete_Lattices.Sup "'a set \<Rightarrow> 'a::Sup"}\\
 | |
| 30401 | 86 | \end{tabular}
 | 
| 87 | ||
| 88 | \subsubsection*{Syntax}
 | |
| 89 | ||
| 74334 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 wenzelm parents: 
73761diff
changeset | 90 | Available via \<^theory_text>\<open>unbundle lattice_syntax\<close>. | 
| 30401 | 91 | |
| 74719 | 92 | \begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | 
| 69597 | 93 | @{text[source]"x \<sqsubseteq> y"} & \<^term>\<open>x \<le> y\<close>\\
 | 
| 94 | @{text[source]"x \<sqsubset> y"} & \<^term>\<open>x < y\<close>\\
 | |
| 95 | @{text[source]"x \<sqinter> y"} & \<^term>\<open>inf x y\<close>\\
 | |
| 96 | @{text[source]"x \<squnion> y"} & \<^term>\<open>sup x y\<close>\\
 | |
| 97 | @{text[source]"\<Sqinter>A"} & \<^term>\<open>Inf A\<close>\\
 | |
| 98 | @{text[source]"\<Squnion>A"} & \<^term>\<open>Sup A\<close>\\
 | |
| 30440 | 99 | @{text[source]"\<top>"} & @{term[source] top}\\
 | 
| 100 | @{text[source]"\<bottom>"} & @{term[source] bot}\\
 | |
| 74719 | 101 | \end{supertabular}
 | 
| 30401 | 102 | |
| 103 | ||
| 50581 | 104 | \section*{Set}
 | 
| 30293 | 105 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 106 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 107 | \<^const>\<open>Set.empty\<close> & @{term_type_only "Set.empty" "'a set"}\\
 | 
| 108 | \<^const>\<open>Set.insert\<close> & @{term_type_only insert "'a\<Rightarrow>'a set\<Rightarrow>'a set"}\\
 | |
| 109 | \<^const>\<open>Collect\<close> & @{term_type_only Collect "('a\<Rightarrow>bool)\<Rightarrow>'a set"}\\
 | |
| 110 | \<^const>\<open>Set.member\<close> & @{term_type_only Set.member "'a\<Rightarrow>'a set\<Rightarrow>bool"} & (\<^verbatim>\<open>:\<close>)\\
 | |
| 111 | \<^const>\<open>Set.union\<close> & @{term_type_only Set.union "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\<^verbatim>\<open>Un\<close>)\\
 | |
| 112 | \<^const>\<open>Set.inter\<close> & @{term_type_only Set.inter "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\<^verbatim>\<open>Int\<close>)\\
 | |
| 113 | \<^const>\<open>Union\<close> & @{term_type_only Union "'a set set\<Rightarrow>'a set"}\\
 | |
| 114 | \<^const>\<open>Inter\<close> & @{term_type_only Inter "'a set set\<Rightarrow>'a set"}\\
 | |
| 115 | \<^const>\<open>Pow\<close> & @{term_type_only Pow "'a set \<Rightarrow>'a set set"}\\
 | |
| 116 | \<^const>\<open>UNIV\<close> & @{term_type_only UNIV "'a set"}\\
 | |
| 117 | \<^const>\<open>image\<close> & @{term_type_only image "('a\<Rightarrow>'b)\<Rightarrow>'a set\<Rightarrow>'b set"}\\
 | |
| 118 | \<^const>\<open>Ball\<close> & @{term_type_only Ball "'a set\<Rightarrow>('a\<Rightarrow>bool)\<Rightarrow>bool"}\\
 | |
| 119 | \<^const>\<open>Bex\<close> & @{term_type_only Bex "'a set\<Rightarrow>('a\<Rightarrow>bool)\<Rightarrow>bool"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 120 | \end{tabular}
 | 
| 30293 | 121 | |
| 122 | \subsubsection*{Syntax}
 | |
| 123 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 124 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 61996 | 125 | \<open>{a\<^sub>1,\<dots>,a\<^sub>n}\<close> & \<open>insert a\<^sub>1 (\<dots> (insert a\<^sub>n {})\<dots>)\<close>\\
 | 
| 69597 | 126 | \<^term>\<open>a \<notin> A\<close> & @{term[source]"\<not>(x \<in> A)"}\\
 | 
| 127 | \<^term>\<open>A \<subseteq> B\<close> & @{term[source]"A \<le> B"}\\
 | |
| 128 | \<^term>\<open>A \<subset> B\<close> & @{term[source]"A < B"}\\
 | |
| 30293 | 129 | @{term[source]"A \<supseteq> B"} & @{term[source]"B \<le> A"}\\
 | 
| 130 | @{term[source]"A \<supset> B"} & @{term[source]"B < A"}\\
 | |
| 69597 | 131 | \<^term>\<open>{x. P}\<close> & @{term[source]"Collect (\<lambda>x. P)"}\\
 | 
| 61996 | 132 | \<open>{t | x\<^sub>1 \<dots> x\<^sub>n. P}\<close> & \<open>{v. \<exists>x\<^sub>1 \<dots> x\<^sub>n. v = t \<and> P}\<close>\\
 | 
| 69313 | 133 | @{term[source]"\<Union>x\<in>I. A"} & @{term[source]"\<Union>((\<lambda>x. A) ` I)"} & (\texttt{UN})\\
 | 
| 134 | @{term[source]"\<Union>x. A"} & @{term[source]"\<Union>((\<lambda>x. A) ` UNIV)"}\\
 | |
| 135 | @{term[source]"\<Inter>x\<in>I. A"} & @{term[source]"\<Inter>((\<lambda>x. A) ` I)"} & (\texttt{INT})\\
 | |
| 136 | @{term[source]"\<Inter>x. A"} & @{term[source]"\<Inter>((\<lambda>x. A) ` UNIV)"}\\
 | |
| 69597 | 137 | \<^term>\<open>\<forall>x\<in>A. P\<close> & @{term[source]"Ball A (\<lambda>x. P)"}\\
 | 
| 138 | \<^term>\<open>\<exists>x\<in>A. P\<close> & @{term[source]"Bex A (\<lambda>x. P)"}\\
 | |
| 139 | \<^term>\<open>range f\<close> & @{term[source]"f ` UNIV"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 140 | \end{tabular}
 | 
| 30293 | 141 | |
| 142 | ||
| 50581 | 143 | \section*{Fun}
 | 
| 30293 | 144 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 145 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 146 | \<^const>\<open>Fun.id\<close> & \<^typeof>\<open>Fun.id\<close>\\ | 
| 147 | \<^const>\<open>Fun.comp\<close> & \<^typeof>\<open>Fun.comp\<close> & (\texttt{o})\\
 | |
| 148 | \<^const>\<open>Fun.inj_on\<close> & @{term_type_only Fun.inj_on "('a\<Rightarrow>'b)\<Rightarrow>'a set\<Rightarrow>bool"}\\
 | |
| 149 | \<^const>\<open>Fun.inj\<close> & \<^typeof>\<open>Fun.inj\<close>\\ | |
| 150 | \<^const>\<open>Fun.surj\<close> & \<^typeof>\<open>Fun.surj\<close>\\ | |
| 151 | \<^const>\<open>Fun.bij\<close> & \<^typeof>\<open>Fun.bij\<close>\\ | |
| 152 | \<^const>\<open>Fun.bij_betw\<close> & @{term_type_only Fun.bij_betw "('a\<Rightarrow>'b)\<Rightarrow>'a set\<Rightarrow>'b set\<Rightarrow>bool"}\\
 | |
| 76054 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 153 | \<^const>\<open>Fun.monotone_on\<close> & \<^typeof>\<open>Fun.monotone_on\<close>\\ | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 154 | \<^const>\<open>Fun.monotone\<close> & \<^typeof>\<open>Fun.monotone\<close>\\ | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 155 | \<^const>\<open>Fun.mono_on\<close> & \<^typeof>\<open>Fun.mono_on\<close>\\ | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 156 | \<^const>\<open>Fun.mono\<close> & \<^typeof>\<open>Fun.mono\<close>\\ | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 157 | \<^const>\<open>Fun.strict_mono_on\<close> & \<^typeof>\<open>Fun.strict_mono_on\<close>\\ | 
| 
a4b47c684445
moved mono and strict_mono to Fun and redefined them as abbreviations
 desharna parents: 
74720diff
changeset | 158 | \<^const>\<open>Fun.strict_mono\<close> & \<^typeof>\<open>Fun.strict_mono\<close>\\ | 
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
76054diff
changeset | 159 | \<^const>\<open>Fun.antimono\<close> & \<^typeof>\<open>Fun.antimono\<close>\\ | 
| 69597 | 160 | \<^const>\<open>Fun.fun_upd\<close> & \<^typeof>\<open>Fun.fun_upd\<close>\\ | 
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 161 | \end{tabular}
 | 
| 30293 | 162 | |
| 163 | \subsubsection*{Syntax}
 | |
| 164 | ||
| 165 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 69597 | 166 | \<^term>\<open>fun_upd f x y\<close> & @{term[source]"fun_upd f x y"}\\
 | 
| 61996 | 167 | \<open>f(x\<^sub>1:=y\<^sub>1,\<dots>,x\<^sub>n:=y\<^sub>n)\<close> & \<open>f(x\<^sub>1:=y\<^sub>1)\<dots>(x\<^sub>n:=y\<^sub>n)\<close>\\ | 
| 30293 | 168 | \end{tabular}
 | 
| 169 | ||
| 170 | ||
| 50581 | 171 | \section*{Hilbert\_Choice}
 | 
| 33019 | 172 | |
| 69597 | 173 | Hilbert's selection ($\varepsilon$) operator: \<^term>\<open>SOME x. P\<close>. | 
| 61996 | 174 | \<^smallskip> | 
| 33019 | 175 | |
| 176 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 177 | \<^const>\<open>Hilbert_Choice.inv_into\<close> & @{term_type_only Hilbert_Choice.inv_into "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"}
 | 
| 33019 | 178 | \end{tabular}
 | 
| 179 | ||
| 180 | \subsubsection*{Syntax}
 | |
| 181 | ||
| 182 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 69597 | 183 | \<^term>\<open>inv\<close> & @{term[source]"inv_into UNIV"}
 | 
| 33019 | 184 | \end{tabular}
 | 
| 185 | ||
| 50581 | 186 | \section*{Fixed Points}
 | 
| 30293 | 187 | |
| 69597 | 188 | Theory: \<^theory>\<open>HOL.Inductive\<close>. | 
| 30293 | 189 | |
| 69597 | 190 | Least and greatest fixed points in a complete lattice \<^typ>\<open>'a\<close>: | 
| 30293 | 191 | |
| 192 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 193 | \<^const>\<open>Inductive.lfp\<close> & \<^typeof>\<open>Inductive.lfp\<close>\\ | 
| 194 | \<^const>\<open>Inductive.gfp\<close> & \<^typeof>\<open>Inductive.gfp\<close>\\ | |
| 30293 | 195 | \end{tabular}
 | 
| 196 | ||
| 69597 | 197 | Note that in particular sets (\<^typ>\<open>'a \<Rightarrow> bool\<close>) are complete lattices. | 
| 30293 | 198 | |
| 50581 | 199 | \section*{Sum\_Type}
 | 
| 30293 | 200 | |
| 61996 | 201 | Type constructor \<open>+\<close>. | 
| 30293 | 202 | |
| 203 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 204 | \<^const>\<open>Sum_Type.Inl\<close> & \<^typeof>\<open>Sum_Type.Inl\<close>\\ | 
| 205 | \<^const>\<open>Sum_Type.Inr\<close> & \<^typeof>\<open>Sum_Type.Inr\<close>\\ | |
| 206 | \<^const>\<open>Sum_Type.Plus\<close> & @{term_type_only Sum_Type.Plus "'a set\<Rightarrow>'b set\<Rightarrow>('a+'b)set"}
 | |
| 30293 | 207 | \end{tabular}
 | 
| 208 | ||
| 209 | ||
| 50581 | 210 | \section*{Product\_Type}
 | 
| 30293 | 211 | |
| 69597 | 212 | Types \<^typ>\<open>unit\<close> and \<open>\<times>\<close>. | 
| 30293 | 213 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 214 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 215 | \<^const>\<open>Product_Type.Unity\<close> & \<^typeof>\<open>Product_Type.Unity\<close>\\ | 
| 216 | \<^const>\<open>Pair\<close> & \<^typeof>\<open>Pair\<close>\\ | |
| 217 | \<^const>\<open>fst\<close> & \<^typeof>\<open>fst\<close>\\ | |
| 218 | \<^const>\<open>snd\<close> & \<^typeof>\<open>snd\<close>\\ | |
| 219 | \<^const>\<open>case_prod\<close> & \<^typeof>\<open>case_prod\<close>\\ | |
| 220 | \<^const>\<open>curry\<close> & \<^typeof>\<open>curry\<close>\\ | |
| 221 | \<^const>\<open>Product_Type.Sigma\<close> & @{term_type_only Product_Type.Sigma "'a set\<Rightarrow>('a\<Rightarrow>'b set)\<Rightarrow>('a*'b)set"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 222 | \end{tabular}
 | 
| 30293 | 223 | |
| 224 | \subsubsection*{Syntax}
 | |
| 225 | ||
| 30440 | 226 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} ll @ {}}
 | 
| 69597 | 227 | \<^term>\<open>Pair a b\<close> & @{term[source]"Pair a b"}\\
 | 
| 228 | \<^term>\<open>case_prod (\<lambda>x y. t)\<close> & @{term[source]"case_prod (\<lambda>x y. t)"}\\
 | |
| 229 | \<^term>\<open>A \<times> B\<close> & \<open>Sigma A (\<lambda>\<^latex>\<open>\_\<close>. B)\<close> | |
| 30293 | 230 | \end{tabular}
 | 
| 231 | ||
| 232 | Pairs may be nested. Nesting to the right is printed as a tuple, | |
| 69597 | 233 | e.g.\ \mbox{\<^term>\<open>(a,b,c)\<close>} is really \mbox{\<open>(a, (b, c))\<close>.}
 | 
| 30293 | 234 | Pattern matching with pairs and tuples extends to all binders, | 
| 69597 | 235 | e.g.\ \mbox{\<^prop>\<open>\<forall>(x,y)\<in>A. P\<close>,} \<^term>\<open>{(x,y). P}\<close>, etc.
 | 
| 30293 | 236 | |
| 237 | ||
| 50581 | 238 | \section*{Relation}
 | 
| 30293 | 239 | |
| 74719 | 240 | \begin{supertabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 241 | \<^const>\<open>Relation.converse\<close> & @{term_type_only Relation.converse "('a * 'b)set \<Rightarrow> ('b*'a)set"}\\
 | 
| 242 | \<^const>\<open>Relation.relcomp\<close> & @{term_type_only Relation.relcomp "('a*'b)set\<Rightarrow>('b*'c)set\<Rightarrow>('a*'c)set"}\\
 | |
| 243 | \<^const>\<open>Relation.Image\<close> & @{term_type_only Relation.Image "('a*'b)set\<Rightarrow>'a set\<Rightarrow>'b set"}\\
 | |
| 244 | \<^const>\<open>Relation.inv_image\<close> & @{term_type_only Relation.inv_image "('a*'a)set\<Rightarrow>('b\<Rightarrow>'a)\<Rightarrow>('b*'b)set"}\\
 | |
| 245 | \<^const>\<open>Relation.Id_on\<close> & @{term_type_only Relation.Id_on "'a set\<Rightarrow>('a*'a)set"}\\
 | |
| 246 | \<^const>\<open>Relation.Id\<close> & @{term_type_only Relation.Id "('a*'a)set"}\\
 | |
| 247 | \<^const>\<open>Relation.Domain\<close> & @{term_type_only Relation.Domain "('a*'b)set\<Rightarrow>'a set"}\\
 | |
| 248 | \<^const>\<open>Relation.Range\<close> & @{term_type_only Relation.Range "('a*'b)set\<Rightarrow>'b set"}\\
 | |
| 249 | \<^const>\<open>Relation.Field\<close> & @{term_type_only Relation.Field "('a*'a)set\<Rightarrow>'a set"}\\
 | |
| 250 | \<^const>\<open>Relation.refl_on\<close> & @{term_type_only Relation.refl_on "'a set\<Rightarrow>('a*'a)set\<Rightarrow>bool"}\\
 | |
| 251 | \<^const>\<open>Relation.refl\<close> & @{term_type_only Relation.refl "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 252 | \<^const>\<open>Relation.sym\<close> & @{term_type_only Relation.sym "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 253 | \<^const>\<open>Relation.antisym\<close> & @{term_type_only Relation.antisym "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 254 | \<^const>\<open>Relation.trans\<close> & @{term_type_only Relation.trans "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 255 | \<^const>\<open>Relation.irrefl\<close> & @{term_type_only Relation.irrefl "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 256 | \<^const>\<open>Relation.total_on\<close> & @{term_type_only Relation.total_on "'a set\<Rightarrow>('a*'a)set\<Rightarrow>bool"}\\
 | |
| 257 | \<^const>\<open>Relation.total\<close> & @{term_type_only Relation.total "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 74719 | 258 | \end{supertabular}
 | 
| 30293 | 259 | |
| 260 | \subsubsection*{Syntax}
 | |
| 261 | ||
| 30440 | 262 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 263 | \<^term>\<open>converse r\<close> & @{term[source]"converse r"} & (\<^verbatim>\<open>^-1\<close>)
 | 
| 30293 | 264 | \end{tabular}
 | 
| 61996 | 265 | \<^medskip> | 
| 47187 | 266 | |
| 267 | \noindent | |
| 69597 | 268 | Type synonym \ \<^typ>\<open>'a rel\<close> \<open>=\<close> @{expanded_typ "'a rel"}
 | 
| 30293 | 269 | |
| 50581 | 270 | \section*{Equiv\_Relations}
 | 
| 30293 | 271 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 272 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 273 | \<^const>\<open>Equiv_Relations.equiv\<close> & @{term_type_only Equiv_Relations.equiv "'a set \<Rightarrow> ('a*'a)set\<Rightarrow>bool"}\\
 | 
| 274 | \<^const>\<open>Equiv_Relations.quotient\<close> & @{term_type_only Equiv_Relations.quotient "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set set"}\\
 | |
| 275 | \<^const>\<open>Equiv_Relations.congruent\<close> & @{term_type_only Equiv_Relations.congruent "('a*'a)set\<Rightarrow>('a\<Rightarrow>'b)\<Rightarrow>bool"}\\
 | |
| 276 | \<^const>\<open>Equiv_Relations.congruent2\<close> & @{term_type_only Equiv_Relations.congruent2 "('a*'a)set\<Rightarrow>('b*'b)set\<Rightarrow>('a\<Rightarrow>'b\<Rightarrow>'c)\<Rightarrow>bool"}\\
 | |
| 30293 | 277 | %@ {const Equiv_Relations.} & @ {term_type_only Equiv_Relations. ""}\\
 | 
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 278 | \end{tabular}
 | 
| 30293 | 279 | |
| 280 | \subsubsection*{Syntax}
 | |
| 281 | ||
| 282 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 69597 | 283 | \<^term>\<open>congruent r f\<close> & @{term[source]"congruent r f"}\\
 | 
| 284 | \<^term>\<open>congruent2 r r f\<close> & @{term[source]"congruent2 r r f"}\\
 | |
| 30293 | 285 | \end{tabular}
 | 
| 286 | ||
| 287 | ||
| 50581 | 288 | \section*{Transitive\_Closure}
 | 
| 30293 | 289 | |
| 290 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 291 | \<^const>\<open>Transitive_Closure.rtrancl\<close> & @{term_type_only Transitive_Closure.rtrancl "('a*'a)set\<Rightarrow>('a*'a)set"}\\
 | 
| 292 | \<^const>\<open>Transitive_Closure.trancl\<close> & @{term_type_only Transitive_Closure.trancl "('a*'a)set\<Rightarrow>('a*'a)set"}\\
 | |
| 293 | \<^const>\<open>Transitive_Closure.reflcl\<close> & @{term_type_only Transitive_Closure.reflcl "('a*'a)set\<Rightarrow>('a*'a)set"}\\
 | |
| 294 | \<^const>\<open>Transitive_Closure.acyclic\<close> & @{term_type_only Transitive_Closure.acyclic "('a*'a)set\<Rightarrow>bool"}\\
 | |
| 295 | \<^const>\<open>compower\<close> & @{term_type_only "(^^) :: ('a*'a)set\<Rightarrow>nat\<Rightarrow>('a*'a)set" "('a*'a)set\<Rightarrow>nat\<Rightarrow>('a*'a)set"}\\
 | |
| 30293 | 296 | \end{tabular}
 | 
| 297 | ||
| 298 | \subsubsection*{Syntax}
 | |
| 299 | ||
| 30440 | 300 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 301 | \<^term>\<open>rtrancl r\<close> & @{term[source]"rtrancl r"} & (\<^verbatim>\<open>^*\<close>)\\
 | 
| 302 | \<^term>\<open>trancl r\<close> & @{term[source]"trancl r"} & (\<^verbatim>\<open>^+\<close>)\\
 | |
| 303 | \<^term>\<open>reflcl r\<close> & @{term[source]"reflcl r"} & (\<^verbatim>\<open>^=\<close>)
 | |
| 30293 | 304 | \end{tabular}
 | 
| 305 | ||
| 306 | ||
| 50581 | 307 | \section*{Algebra}
 | 
| 30293 | 308 | |
| 80453 
7a2d9e3fcdd5
moved transitional theory Divides to HOL-Library
 haftmann parents: 
79766diff
changeset | 309 | Theories \<^theory>\<open>HOL.Groups\<close>, \<^theory>\<open>HOL.Rings\<close>, | 
| 
7a2d9e3fcdd5
moved transitional theory Divides to HOL-Library
 haftmann parents: 
79766diff
changeset | 310 | \<^theory>\<open>HOL.Euclidean_Rings\<close> and \<^theory>\<open>HOL.Fields\<close> | 
| 
7a2d9e3fcdd5
moved transitional theory Divides to HOL-Library
 haftmann parents: 
79766diff
changeset | 311 | define a large collection of classes describing common algebraic | 
| 30440 | 312 | structures from semigroups up to fields. Everything is done in terms of | 
| 313 | overloaded operators: | |
| 314 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 315 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 316 | \<open>0\<close> & \<^typeof>\<open>zero\<close>\\ | 
| 317 | \<open>1\<close> & \<^typeof>\<open>one\<close>\\ | |
| 318 | \<^const>\<open>plus\<close> & \<^typeof>\<open>plus\<close>\\ | |
| 319 | \<^const>\<open>minus\<close> & \<^typeof>\<open>minus\<close>\\ | |
| 320 | \<^const>\<open>uminus\<close> & \<^typeof>\<open>uminus\<close> & (\<^verbatim>\<open>-\<close>)\\ | |
| 321 | \<^const>\<open>times\<close> & \<^typeof>\<open>times\<close>\\ | |
| 322 | \<^const>\<open>inverse\<close> & \<^typeof>\<open>inverse\<close>\\ | |
| 323 | \<^const>\<open>divide\<close> & \<^typeof>\<open>divide\<close>\\ | |
| 324 | \<^const>\<open>abs\<close> & \<^typeof>\<open>abs\<close>\\ | |
| 325 | \<^const>\<open>sgn\<close> & \<^typeof>\<open>sgn\<close>\\ | |
| 326 | \<^const>\<open>Rings.dvd\<close> & \<^typeof>\<open>Rings.dvd\<close>\\ | |
| 327 | \<^const>\<open>divide\<close> & \<^typeof>\<open>divide\<close>\\ | |
| 328 | \<^const>\<open>modulo\<close> & \<^typeof>\<open>modulo\<close>\\ | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 329 | \end{tabular}
 | 
| 30440 | 330 | |
| 331 | \subsubsection*{Syntax}
 | |
| 332 | ||
| 333 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 69597 | 334 | \<^term>\<open>\<bar>x\<bar>\<close> & @{term[source] "abs x"}
 | 
| 30440 | 335 | \end{tabular}
 | 
| 30293 | 336 | |
| 337 | ||
| 50581 | 338 | \section*{Nat}
 | 
| 30293 | 339 | |
| 69597 | 340 | \<^datatype>\<open>nat\<close> | 
| 61996 | 341 | \<^bigskip> | 
| 30293 | 342 | |
| 343 | \begin{tabular}{@ {} lllllll @ {}}
 | |
| 69597 | 344 | \<^term>\<open>(+) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | 
| 345 | \<^term>\<open>(-) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | |
| 346 | \<^term>\<open>(*) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | |
| 347 | \<^term>\<open>(^) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | |
| 348 | \<^term>\<open>(div) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close>& | |
| 349 | \<^term>\<open>(mod) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close>& | |
| 350 | \<^term>\<open>(dvd) :: nat \<Rightarrow> nat \<Rightarrow> bool\<close>\\ | |
| 351 | \<^term>\<open>(\<le>) :: nat \<Rightarrow> nat \<Rightarrow> bool\<close> & | |
| 352 | \<^term>\<open>(<) :: nat \<Rightarrow> nat \<Rightarrow> bool\<close> & | |
| 353 | \<^term>\<open>min :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | |
| 354 | \<^term>\<open>max :: nat \<Rightarrow> nat \<Rightarrow> nat\<close> & | |
| 355 | \<^term>\<open>Min :: nat set \<Rightarrow> nat\<close> & | |
| 356 | \<^term>\<open>Max :: nat set \<Rightarrow> nat\<close>\\ | |
| 30293 | 357 | \end{tabular}
 | 
| 358 | ||
| 359 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 360 | \<^const>\<open>Nat.of_nat\<close> & \<^typeof>\<open>Nat.of_nat\<close>\\ | 
| 361 | \<^term>\<open>(^^) :: ('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> &
 | |
| 67399 | 362 |   @{term_type_only "(^^) :: ('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"}
 | 
| 30293 | 363 | \end{tabular}
 | 
| 364 | ||
| 50581 | 365 | \section*{Int}
 | 
| 30293 | 366 | |
| 69597 | 367 | Type \<^typ>\<open>int\<close> | 
| 61996 | 368 | \<^bigskip> | 
| 30293 | 369 | |
| 370 | \begin{tabular}{@ {} llllllll @ {}}
 | |
| 69597 | 371 | \<^term>\<open>(+) :: int \<Rightarrow> int \<Rightarrow> int\<close> & | 
| 372 | \<^term>\<open>(-) :: int \<Rightarrow> int \<Rightarrow> int\<close> & | |
| 373 | \<^term>\<open>uminus :: int \<Rightarrow> int\<close> & | |
| 374 | \<^term>\<open>(*) :: int \<Rightarrow> int \<Rightarrow> int\<close> & | |
| 375 | \<^term>\<open>(^) :: int \<Rightarrow> nat \<Rightarrow> int\<close> & | |
| 376 | \<^term>\<open>(div) :: int \<Rightarrow> int \<Rightarrow> int\<close>& | |
| 377 | \<^term>\<open>(mod) :: int \<Rightarrow> int \<Rightarrow> int\<close>& | |
| 378 | \<^term>\<open>(dvd) :: int \<Rightarrow> int \<Rightarrow> bool\<close>\\ | |
| 379 | \<^term>\<open>(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool\<close> & | |
| 380 | \<^term>\<open>(<) :: int \<Rightarrow> int \<Rightarrow> bool\<close> & | |
| 381 | \<^term>\<open>min :: int \<Rightarrow> int \<Rightarrow> int\<close> & | |
| 382 | \<^term>\<open>max :: int \<Rightarrow> int \<Rightarrow> int\<close> & | |
| 383 | \<^term>\<open>Min :: int set \<Rightarrow> int\<close> & | |
| 384 | \<^term>\<open>Max :: int set \<Rightarrow> int\<close>\\ | |
| 385 | \<^term>\<open>abs :: int \<Rightarrow> int\<close> & | |
| 386 | \<^term>\<open>sgn :: int \<Rightarrow> int\<close>\\ | |
| 30293 | 387 | \end{tabular}
 | 
| 388 | ||
| 30440 | 389 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 390 | \<^const>\<open>Int.nat\<close> & \<^typeof>\<open>Int.nat\<close>\\ | 
| 391 | \<^const>\<open>Int.of_int\<close> & \<^typeof>\<open>Int.of_int\<close>\\ | |
| 392 | \<^const>\<open>Int.Ints\<close> & @{term_type_only Int.Ints "'a::ring_1 set"} & (\<^verbatim>\<open>Ints\<close>)
 | |
| 30293 | 393 | \end{tabular}
 | 
| 394 | ||
| 395 | \subsubsection*{Syntax}
 | |
| 396 | ||
| 397 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 69597 | 398 | \<^term>\<open>of_nat::nat\<Rightarrow>int\<close> & @{term[source]"of_nat"}\\
 | 
| 30293 | 399 | \end{tabular}
 | 
| 400 | ||
| 401 | ||
| 50581 | 402 | \section*{Finite\_Set}
 | 
| 30401 | 403 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 404 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 405 | \<^const>\<open>Finite_Set.finite\<close> & @{term_type_only Finite_Set.finite "'a set\<Rightarrow>bool"}\\
 | 
| 406 | \<^const>\<open>Finite_Set.card\<close> & @{term_type_only Finite_Set.card "'a set \<Rightarrow> nat"}\\
 | |
| 407 | \<^const>\<open>Finite_Set.fold\<close> & @{term_type_only Finite_Set.fold "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 408 | \end{tabular}
 | 
| 64281 
bfc2e92d9b4c
restored document structure after theory refactoring
 haftmann parents: 
64272diff
changeset | 409 | |
| 
bfc2e92d9b4c
restored document structure after theory refactoring
 haftmann parents: 
64272diff
changeset | 410 | |
| 65952 | 411 | \section*{Lattices\_Big}
 | 
| 412 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 413 | \begin{tabular}{@ {} l @ {~::~} l l @ {}}
 | 
| 69597 | 414 | \<^const>\<open>Lattices_Big.Min\<close> & \<^typeof>\<open>Lattices_Big.Min\<close>\\ | 
| 415 | \<^const>\<open>Lattices_Big.Max\<close> & \<^typeof>\<open>Lattices_Big.Max\<close>\\ | |
| 416 | \<^const>\<open>Lattices_Big.arg_min\<close> & \<^typeof>\<open>Lattices_Big.arg_min\<close>\\ | |
| 417 | \<^const>\<open>Lattices_Big.is_arg_min\<close> & \<^typeof>\<open>Lattices_Big.is_arg_min\<close>\\ | |
| 418 | \<^const>\<open>Lattices_Big.arg_max\<close> & \<^typeof>\<open>Lattices_Big.arg_max\<close>\\ | |
| 419 | \<^const>\<open>Lattices_Big.is_arg_max\<close> & \<^typeof>\<open>Lattices_Big.is_arg_max\<close>\\ | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 420 | \end{tabular}
 | 
| 65952 | 421 | |
| 422 | \subsubsection*{Syntax}
 | |
| 423 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 424 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 425 | \<^term>\<open>ARG_MIN f x. P\<close> & @{term[source]"arg_min f (\<lambda>x. P)"}\\
 | 
| 426 | \<^term>\<open>ARG_MAX f x. P\<close> & @{term[source]"arg_max f (\<lambda>x. P)"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 427 | \end{tabular}
 | 
| 65952 | 428 | |
| 429 | ||
| 64281 
bfc2e92d9b4c
restored document structure after theory refactoring
 haftmann parents: 
64272diff
changeset | 430 | \section*{Groups\_Big}
 | 
| 
bfc2e92d9b4c
restored document structure after theory refactoring
 haftmann parents: 
64272diff
changeset | 431 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 432 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 433 | \<^const>\<open>Groups_Big.sum\<close> & @{term_type_only Groups_Big.sum "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b::comm_monoid_add"}\\
 | 
| 434 | \<^const>\<open>Groups_Big.prod\<close> & @{term_type_only Groups_Big.prod "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b::comm_monoid_mult"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 435 | \end{tabular}
 | 
| 30401 | 436 | |
| 437 | ||
| 438 | \subsubsection*{Syntax}
 | |
| 439 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 440 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
 | 
| 69597 | 441 | \<^term>\<open>sum (\<lambda>x. x) A\<close> & @{term[source]"sum (\<lambda>x. x) A"} & (\<^verbatim>\<open>SUM\<close>)\\
 | 
| 442 | \<^term>\<open>sum (\<lambda>x. t) A\<close> & @{term[source]"sum (\<lambda>x. t) A"}\\
 | |
| 443 | @{term[source] "\<Sum>x|P. t"} & \<^term>\<open>\<Sum>x|P. t\<close>\\
 | |
| 61996 | 444 | \multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Prod>\<close> instead of \<open>\<Sum>\<close>} & (\<^verbatim>\<open>PROD\<close>)\\
 | 
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 445 | \end{tabular}
 | 
| 30401 | 446 | |
| 447 | ||
| 50581 | 448 | \section*{Wellfounded}
 | 
| 30293 | 449 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 450 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 451 | \<^const>\<open>Wellfounded.wf\<close> & @{term_type_only Wellfounded.wf "('a*'a)set\<Rightarrow>bool"}\\
 | 
| 452 | \<^const>\<open>Wellfounded.acc\<close> & @{term_type_only Wellfounded.acc "('a*'a)set\<Rightarrow>'a set"}\\
 | |
| 453 | \<^const>\<open>Wellfounded.measure\<close> & @{term_type_only Wellfounded.measure "('a\<Rightarrow>nat)\<Rightarrow>('a*'a)set"}\\
 | |
| 454 | \<^const>\<open>Wellfounded.lex_prod\<close> & @{term_type_only Wellfounded.lex_prod "('a*'a)set\<Rightarrow>('b*'b)set\<Rightarrow>(('a*'b)*('a*'b))set"}\\
 | |
| 455 | \<^const>\<open>Wellfounded.mlex_prod\<close> & @{term_type_only Wellfounded.mlex_prod "('a\<Rightarrow>nat)\<Rightarrow>('a*'a)set\<Rightarrow>('a*'a)set"}\\
 | |
| 456 | \<^const>\<open>Wellfounded.less_than\<close> & @{term_type_only Wellfounded.less_than "(nat*nat)set"}\\
 | |
| 457 | \<^const>\<open>Wellfounded.pred_nat\<close> & @{term_type_only Wellfounded.pred_nat "(nat*nat)set"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 458 | \end{tabular}
 | 
| 30293 | 459 | |
| 460 | ||
| 69597 | 461 | \section*{Set\_Interval} % \<^theory>\<open>HOL.Set_Interval\<close>
 | 
| 30321 | 462 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 463 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | 
| 69597 | 464 | \<^const>\<open>lessThan\<close> & @{term_type_only lessThan "'a::ord \<Rightarrow> 'a set"}\\
 | 
| 465 | \<^const>\<open>atMost\<close> & @{term_type_only atMost "'a::ord \<Rightarrow> 'a set"}\\
 | |
| 466 | \<^const>\<open>greaterThan\<close> & @{term_type_only greaterThan "'a::ord \<Rightarrow> 'a set"}\\
 | |
| 467 | \<^const>\<open>atLeast\<close> & @{term_type_only atLeast "'a::ord \<Rightarrow> 'a set"}\\
 | |
| 468 | \<^const>\<open>greaterThanLessThan\<close> & @{term_type_only greaterThanLessThan "'a::ord \<Rightarrow> 'a \<Rightarrow> 'a set"}\\
 | |
| 469 | \<^const>\<open>atLeastLessThan\<close> & @{term_type_only atLeastLessThan "'a::ord \<Rightarrow> 'a \<Rightarrow> 'a set"}\\
 | |
| 470 | \<^const>\<open>greaterThanAtMost\<close> & @{term_type_only greaterThanAtMost "'a::ord \<Rightarrow> 'a \<Rightarrow> 'a set"}\\
 | |
| 471 | \<^const>\<open>atLeastAtMost\<close> & @{term_type_only atLeastAtMost "'a::ord \<Rightarrow> 'a \<Rightarrow> 'a set"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 472 | \end{tabular}
 | 
| 30321 | 473 | |
| 474 | \subsubsection*{Syntax}
 | |
| 475 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 476 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | 
| 69597 | 477 | \<^term>\<open>lessThan y\<close> & @{term[source] "lessThan y"}\\
 | 
| 478 | \<^term>\<open>atMost y\<close> & @{term[source] "atMost y"}\\
 | |
| 479 | \<^term>\<open>greaterThan x\<close> & @{term[source] "greaterThan x"}\\
 | |
| 480 | \<^term>\<open>atLeast x\<close> & @{term[source] "atLeast x"}\\
 | |
| 481 | \<^term>\<open>greaterThanLessThan x y\<close> & @{term[source] "greaterThanLessThan x y"}\\
 | |
| 482 | \<^term>\<open>atLeastLessThan x y\<close> & @{term[source] "atLeastLessThan x y"}\\
 | |
| 483 | \<^term>\<open>greaterThanAtMost x y\<close> & @{term[source] "greaterThanAtMost x y"}\\
 | |
| 484 | \<^term>\<open>atLeastAtMost x y\<close> & @{term[source] "atLeastAtMost x y"}\\
 | |
| 61995 
74709e9c4f17
clarified print modes: Isabelle symbols are used by default, but "latex" mode needs to be for some syntax forms;
 wenzelm parents: 
61943diff
changeset | 485 | @{term[source] "\<Union>i\<le>n. A"} & @{term[source] "\<Union>i \<in> {..n}. A"}\\
 | 
| 
74709e9c4f17
clarified print modes: Isabelle symbols are used by default, but "latex" mode needs to be for some syntax forms;
 wenzelm parents: 
61943diff
changeset | 486 | @{term[source] "\<Union>i<n. A"} & @{term[source] "\<Union>i \<in> {..<n}. A"}\\
 | 
| 61996 | 487 | \multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Inter>\<close> instead of \<open>\<Union>\<close>}\\
 | 
| 69597 | 488 | \<^term>\<open>sum (\<lambda>x. t) {a..b}\<close> & @{term[source] "sum (\<lambda>x. t) {a..b}"}\\
 | 
| 489 | \<^term>\<open>sum (\<lambda>x. t) {a..<b}\<close> & @{term[source] "sum (\<lambda>x. t) {a..<b}"}\\
 | |
| 490 | \<^term>\<open>sum (\<lambda>x. t) {..b}\<close> & @{term[source] "sum (\<lambda>x. t) {..b}"}\\
 | |
| 491 | \<^term>\<open>sum (\<lambda>x. t) {..<b}\<close> & @{term[source] "sum (\<lambda>x. t) {..<b}"}\\
 | |
| 61996 | 492 | \multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Prod>\<close> instead of \<open>\<Sum>\<close>}\\
 | 
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 493 | \end{tabular}
 | 
| 30321 | 494 | |
| 495 | ||
| 50581 | 496 | \section*{Power}
 | 
| 30293 | 497 | |
| 498 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 499 | \<^const>\<open>Power.power\<close> & \<^typeof>\<open>Power.power\<close> | 
| 30293 | 500 | \end{tabular}
 | 
| 501 | ||
| 502 | ||
| 50581 | 503 | \section*{Option}
 | 
| 30293 | 504 | |
| 69597 | 505 | \<^datatype>\<open>option\<close> | 
| 61996 | 506 | \<^bigskip> | 
| 30293 | 507 | |
| 508 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 509 | \<^const>\<open>Option.the\<close> & \<^typeof>\<open>Option.the\<close>\\ | 
| 510 | \<^const>\<open>map_option\<close> & @{typ[source]"('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option"}\\
 | |
| 511 | \<^const>\<open>set_option\<close> & @{term_type_only set_option "'a option \<Rightarrow> 'a set"}\\
 | |
| 512 | \<^const>\<open>Option.bind\<close> & @{term_type_only Option.bind "'a option \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> 'b option"}
 | |
| 30293 | 513 | \end{tabular}
 | 
| 514 | ||
| 50581 | 515 | \section*{List}
 | 
| 30293 | 516 | |
| 69597 | 517 | \<^datatype>\<open>list\<close> | 
| 61996 | 518 | \<^bigskip> | 
| 30293 | 519 | |
| 520 | \begin{supertabular}{@ {} l @ {~::~} l @ {}}
 | |
| 69597 | 521 | \<^const>\<open>List.append\<close> & \<^typeof>\<open>List.append\<close>\\ | 
| 522 | \<^const>\<open>List.butlast\<close> & \<^typeof>\<open>List.butlast\<close>\\ | |
| 523 | \<^const>\<open>List.concat\<close> & \<^typeof>\<open>List.concat\<close>\\ | |
| 524 | \<^const>\<open>List.distinct\<close> & \<^typeof>\<open>List.distinct\<close>\\ | |
| 525 | \<^const>\<open>List.drop\<close> & \<^typeof>\<open>List.drop\<close>\\ | |
| 526 | \<^const>\<open>List.dropWhile\<close> & \<^typeof>\<open>List.dropWhile\<close>\\ | |
| 527 | \<^const>\<open>List.filter\<close> & \<^typeof>\<open>List.filter\<close>\\ | |
| 528 | \<^const>\<open>List.find\<close> & \<^typeof>\<open>List.find\<close>\\ | |
| 529 | \<^const>\<open>List.fold\<close> & \<^typeof>\<open>List.fold\<close>\\ | |
| 530 | \<^const>\<open>List.foldr\<close> & \<^typeof>\<open>List.foldr\<close>\\ | |
| 531 | \<^const>\<open>List.foldl\<close> & \<^typeof>\<open>List.foldl\<close>\\ | |
| 532 | \<^const>\<open>List.hd\<close> & \<^typeof>\<open>List.hd\<close>\\ | |
| 533 | \<^const>\<open>List.last\<close> & \<^typeof>\<open>List.last\<close>\\ | |
| 534 | \<^const>\<open>List.length\<close> & \<^typeof>\<open>List.length\<close>\\ | |
| 535 | \<^const>\<open>List.lenlex\<close> & @{term_type_only List.lenlex "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 | |
| 536 | \<^const>\<open>List.lex\<close> & @{term_type_only List.lex "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 | |
| 537 | \<^const>\<open>List.lexn\<close> & @{term_type_only List.lexn "('a*'a)set\<Rightarrow>nat\<Rightarrow>('a list * 'a list)set"}\\
 | |
| 538 | \<^const>\<open>List.lexord\<close> & @{term_type_only List.lexord "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 | |
| 539 | \<^const>\<open>List.listrel\<close> & @{term_type_only List.listrel "('a*'b)set\<Rightarrow>('a list * 'b list)set"}\\
 | |
| 540 | \<^const>\<open>List.listrel1\<close> & @{term_type_only List.listrel1 "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 | |
| 541 | \<^const>\<open>List.lists\<close> & @{term_type_only List.lists "'a set\<Rightarrow>'a list set"}\\
 | |
| 542 | \<^const>\<open>List.listset\<close> & @{term_type_only List.listset "'a set list \<Rightarrow> 'a list set"}\\
 | |
| 543 | \<^const>\<open>List.list_all2\<close> & \<^typeof>\<open>List.list_all2\<close>\\ | |
| 544 | \<^const>\<open>List.list_update\<close> & \<^typeof>\<open>List.list_update\<close>\\ | |
| 545 | \<^const>\<open>List.map\<close> & \<^typeof>\<open>List.map\<close>\\ | |
| 546 | \<^const>\<open>List.measures\<close> & @{term_type_only List.measures "('a\<Rightarrow>nat)list\<Rightarrow>('a*'a)set"}\\
 | |
| 547 | \<^const>\<open>List.nth\<close> & \<^typeof>\<open>List.nth\<close>\\ | |
| 548 | \<^const>\<open>List.nths\<close> & \<^typeof>\<open>List.nths\<close>\\ | |
| 79766 | 549 | \<^const>\<open>Groups_List.prod_list\<close> & \<^typeof>\<open>Groups_List.prod_list\<close>\\ | 
| 69597 | 550 | \<^const>\<open>List.remdups\<close> & \<^typeof>\<open>List.remdups\<close>\\ | 
| 551 | \<^const>\<open>List.removeAll\<close> & \<^typeof>\<open>List.removeAll\<close>\\ | |
| 552 | \<^const>\<open>List.remove1\<close> & \<^typeof>\<open>List.remove1\<close>\\ | |
| 553 | \<^const>\<open>List.replicate\<close> & \<^typeof>\<open>List.replicate\<close>\\ | |
| 554 | \<^const>\<open>List.rev\<close> & \<^typeof>\<open>List.rev\<close>\\ | |
| 555 | \<^const>\<open>List.rotate\<close> & \<^typeof>\<open>List.rotate\<close>\\ | |
| 556 | \<^const>\<open>List.rotate1\<close> & \<^typeof>\<open>List.rotate1\<close>\\ | |
| 557 | \<^const>\<open>List.set\<close> & @{term_type_only List.set "'a list \<Rightarrow> 'a set"}\\
 | |
| 558 | \<^const>\<open>List.shuffles\<close> & \<^typeof>\<open>List.shuffles\<close>\\ | |
| 559 | \<^const>\<open>List.sort\<close> & \<^typeof>\<open>List.sort\<close>\\ | |
| 560 | \<^const>\<open>List.sorted\<close> & \<^typeof>\<open>List.sorted\<close>\\ | |
| 561 | \<^const>\<open>List.sorted_wrt\<close> & \<^typeof>\<open>List.sorted_wrt\<close>\\ | |
| 562 | \<^const>\<open>List.splice\<close> & \<^typeof>\<open>List.splice\<close>\\ | |
| 79766 | 563 | \<^const>\<open>Groups_List.sum_list\<close> & \<^typeof>\<open>Groups_List.sum_list\<close>\\ | 
| 69597 | 564 | \<^const>\<open>List.take\<close> & \<^typeof>\<open>List.take\<close>\\ | 
| 565 | \<^const>\<open>List.takeWhile\<close> & \<^typeof>\<open>List.takeWhile\<close>\\ | |
| 566 | \<^const>\<open>List.tl\<close> & \<^typeof>\<open>List.tl\<close>\\ | |
| 567 | \<^const>\<open>List.upt\<close> & \<^typeof>\<open>List.upt\<close>\\ | |
| 568 | \<^const>\<open>List.upto\<close> & \<^typeof>\<open>List.upto\<close>\\ | |
| 569 | \<^const>\<open>List.zip\<close> & \<^typeof>\<open>List.zip\<close>\\ | |
| 30293 | 570 | \end{supertabular}
 | 
| 571 | ||
| 572 | \subsubsection*{Syntax}
 | |
| 573 | ||
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 574 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | 
| 61996 | 575 | \<open>[x\<^sub>1,\<dots>,x\<^sub>n]\<close> & \<open>x\<^sub>1 # \<dots> # x\<^sub>n # []\<close>\\ | 
| 69597 | 576 | \<^term>\<open>[m..<n]\<close> & @{term[source]"upt m n"}\\
 | 
| 577 | \<^term>\<open>[i..j]\<close> & @{term[source]"upto i j"}\\
 | |
| 578 | \<^term>\<open>xs[n := x]\<close> & @{term[source]"list_update xs n x"}\\
 | |
| 579 | \<^term>\<open>\<Sum>x\<leftarrow>xs. e\<close> & @{term[source]"listsum (map (\<lambda>x. e) xs)"}\\
 | |
| 74658 
4c508826fee8
improve pagebreaks by *not* using supertabular too much;
 wenzelm parents: 
74334diff
changeset | 580 | \end{tabular}
 | 
| 61996 | 581 | \<^medskip> | 
| 30293 | 582 | |
| 68364 | 583 | Filter input syntax \<open>[pat \<leftarrow> e. b]\<close>, where | 
| 69597 | 584 | \<open>pat\<close> is a tuple pattern, which stands for \<^term>\<open>filter (\<lambda>pat. b) e\<close>. | 
| 68364 | 585 | |
| 586 | List comprehension input syntax: \<open>[e. q\<^sub>1, \<dots>, q\<^sub>n]\<close> where each | |
| 61996 | 587 | qualifier \<open>q\<^sub>i\<close> is either a generator \mbox{\<open>pat \<leftarrow> e\<close>} or a
 | 
| 30293 | 588 | guard, i.e.\ boolean expression. | 
| 589 | ||
| 77306 | 590 | \section*{Map}
 | 
| 591 | ||
| 592 | Maps model partial functions and are often used as finite tables. However, | |
| 593 | the domain of a map may be infinite. | |
| 594 | ||
| 595 | \begin{tabular}{@ {} l @ {~::~} l @ {}}
 | |
| 596 | \<^const>\<open>Map.empty\<close> & \<^typeof>\<open>Map.empty\<close>\\ | |
| 597 | \<^const>\<open>Map.map_add\<close> & \<^typeof>\<open>Map.map_add\<close>\\ | |
| 598 | \<^const>\<open>Map.map_comp\<close> & \<^typeof>\<open>Map.map_comp\<close>\\ | |
| 599 | \<^const>\<open>Map.restrict_map\<close> & @{term_type_only Map.restrict_map "('a\<Rightarrow>'b option)\<Rightarrow>'a set\<Rightarrow>('a\<Rightarrow>'b option)"}\\
 | |
| 600 | \<^const>\<open>Map.dom\<close> & @{term_type_only Map.dom "('a\<Rightarrow>'b option)\<Rightarrow>'a set"}\\
 | |
| 601 | \<^const>\<open>Map.ran\<close> & @{term_type_only Map.ran "('a\<Rightarrow>'b option)\<Rightarrow>'b set"}\\
 | |
| 602 | \<^const>\<open>Map.map_le\<close> & \<^typeof>\<open>Map.map_le\<close>\\ | |
| 603 | \<^const>\<open>Map.map_of\<close> & \<^typeof>\<open>Map.map_of\<close>\\ | |
| 604 | \<^const>\<open>Map.map_upds\<close> & \<^typeof>\<open>Map.map_upds\<close>\\ | |
| 605 | \end{tabular}
 | |
| 606 | ||
| 607 | \subsubsection*{Syntax}
 | |
| 608 | ||
| 609 | \begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
 | |
| 610 | \<^term>\<open>Map.empty\<close> & @{term[source] "\<lambda>_. None"}\\
 | |
| 611 | \<^term>\<open>m(x:=Some y)\<close> & @{term[source]"m(x:=Some y)"}\\
 | |
| 612 | \<open>m(x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n)\<close> & @{text[source]"m(x\<^sub>1\<mapsto>y\<^sub>1)\<dots>(x\<^sub>n\<mapsto>y\<^sub>n)"}\\
 | |
| 613 | \<open>[x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n]\<close> & @{text[source]"Map.empty(x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n)"}\\
 | |
| 614 | \<^term>\<open>map_upds m xs ys\<close> & @{term[source]"map_upds m xs ys"}\\
 | |
| 615 | \end{tabular}
 | |
| 30293 | 616 | |
| 69597 | 617 | \section*{Infix operators in Main} % \<^theory>\<open>Main\<close>
 | 
| 50581 | 618 | |
| 619 | \begin{center}
 | |
| 50605 | 620 | \begin{tabular}{llll}
 | 
| 621 | & Operator & precedence & associativity \\ | |
| 622 | \hline | |
| 61996 | 623 | Meta-logic & \<open>\<Longrightarrow>\<close> & 1 & right \\ | 
| 624 | & \<open>\<equiv>\<close> & 2 \\ | |
| 50605 | 625 | \hline | 
| 61996 | 626 | Logic & \<open>\<and>\<close> & 35 & right \\ | 
| 627 | &\<open>\<or>\<close> & 30 & right \\ | |
| 628 | &\<open>\<longrightarrow>\<close>, \<open>\<longleftrightarrow>\<close> & 25 & right\\ | |
| 629 | &\<open>=\<close>, \<open>\<noteq>\<close> & 50 & left\\ | |
| 50605 | 630 | \hline | 
| 61996 | 631 | Orderings & \<open>\<le>\<close>, \<open><\<close>, \<open>\<ge>\<close>, \<open>>\<close> & 50 \\ | 
| 50605 | 632 | \hline | 
| 61996 | 633 | Sets & \<open>\<subseteq>\<close>, \<open>\<subset>\<close>, \<open>\<supseteq>\<close>, \<open>\<supset>\<close> & 50 \\ | 
| 634 | &\<open>\<in>\<close>, \<open>\<notin>\<close> & 50 \\ | |
| 635 | &\<open>\<inter>\<close> & 70 & left \\ | |
| 636 | &\<open>\<union>\<close> & 65 & left \\ | |
| 50605 | 637 | \hline | 
| 61996 | 638 | Functions and Relations & \<open>\<circ>\<close> & 55 & left\\ | 
| 639 | &\<open>`\<close> & 90 & right\\ | |
| 640 | &\<open>O\<close> & 75 & right\\ | |
| 641 | &\<open>``\<close> & 90 & right\\ | |
| 642 | &\<open>^^\<close> & 80 & right\\ | |
| 50605 | 643 | \hline | 
| 61996 | 644 | Numbers & \<open>+\<close>, \<open>-\<close> & 65 & left \\ | 
| 645 | &\<open>*\<close>, \<open>/\<close> & 70 & left \\ | |
| 646 | &\<open>div\<close>, \<open>mod\<close> & 70 & left\\ | |
| 647 | &\<open>^\<close> & 80 & right\\ | |
| 648 | &\<open>dvd\<close> & 50 \\ | |
| 50605 | 649 | \hline | 
| 61996 | 650 | Lists & \<open>#\<close>, \<open>@\<close> & 65 & right\\ | 
| 651 | &\<open>!\<close> & 100 & left | |
| 50581 | 652 | \end{tabular}
 | 
| 653 | \end{center}
 | |
| 61996 | 654 | \<close> | 
| 30293 | 655 | (*<*) | 
| 656 | end | |
| 65956 
639eb3617a86
reorganised material on sublists
 eberlm <eberlm@in.tum.de> parents: 
65954diff
changeset | 657 | (*>*) |