| author | wenzelm | 
| Sat, 11 Jan 2025 22:18:47 +0100 | |
| changeset 81767 | 3f90880136fe | 
| parent 80590 | 505f97165f52 | 
| child 82256 | f65ac4962b66 | 
| permissions | -rw-r--r-- | 
| 61656 | 1  | 
(*:maxLineLen=78:*)  | 
2  | 
||
| 29755 | 3  | 
theory Logic  | 
4  | 
imports Base  | 
|
5  | 
begin  | 
|
| 18537 | 6  | 
|
| 58618 | 7  | 
chapter \<open>Primitive logic \label{ch:logic}\<close>
 | 
| 18537 | 8  | 
|
| 58618 | 9  | 
text \<open>  | 
| 61854 | 10  | 
The logical foundations of Isabelle/Isar are that of the Pure logic, which  | 
| 76987 | 11  | 
has been introduced as a Natural Deduction framework in \<^cite>\<open>paulson700\<close>.  | 
| 61854 | 12  | 
This is essentially the same logic as ``\<open>\<lambda>HOL\<close>'' in the more abstract  | 
| 76987 | 13  | 
setting of Pure Type Systems (PTS) \<^cite>\<open>"Barendregt-Geuvers:2001"\<close>,  | 
| 61854 | 14  | 
although there are some key differences in the specific treatment of simple  | 
15  | 
types in Isabelle/Pure.  | 
|
| 20480 | 16  | 
|
| 61854 | 17  | 
Following type-theoretic parlance, the Pure logic consists of three levels  | 
18  | 
of \<open>\<lambda>\<close>-calculus with corresponding arrows, \<open>\<Rightarrow>\<close> for syntactic function space  | 
|
19  | 
(terms depending on terms), \<open>\<And>\<close> for universal quantification (proofs  | 
|
20  | 
depending on terms), and \<open>\<Longrightarrow>\<close> for implication (proofs depending on proofs).  | 
|
| 20480 | 21  | 
|
| 20537 | 22  | 
Derivations are relative to a logical theory, which declares type  | 
| 61854 | 23  | 
constructors, constants, and axioms. Theory declarations support schematic  | 
24  | 
polymorphism, which is strictly speaking outside the logic.\<^footnote>\<open>This is the  | 
|
25  | 
deeper logical reason, why the theory context \<open>\<Theta>\<close> is separate from the proof  | 
|
26  | 
context \<open>\<Gamma>\<close> of the core calculus: type constructors, term constants, and  | 
|
27  | 
facts (proof constants) may involve arbitrary type schemes, but the type of  | 
|
28  | 
a locally fixed term parameter is also fixed!\<close>  | 
|
| 58618 | 29  | 
\<close>  | 
| 20480 | 30  | 
|
31  | 
||
| 58618 | 32  | 
section \<open>Types \label{sec:types}\<close>
 | 
| 20437 | 33  | 
|
| 58618 | 34  | 
text \<open>  | 
| 61854 | 35  | 
The language of types is an uninterpreted order-sorted first-order algebra;  | 
36  | 
types are qualified by ordered type classes.  | 
|
| 20480 | 37  | 
|
| 61416 | 38  | 
\<^medskip>  | 
| 61854 | 39  | 
A \<^emph>\<open>type class\<close> is an abstract syntactic entity declared in the theory  | 
40  | 
context. The \<^emph>\<open>subclass relation\<close> \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close> is specified by stating an  | 
|
41  | 
acyclic generating relation; the transitive closure is maintained  | 
|
42  | 
internally. The resulting relation is an ordering: reflexive, transitive,  | 
|
43  | 
and antisymmetric.  | 
|
| 20451 | 44  | 
|
| 61854 | 45  | 
  A \<^emph>\<open>sort\<close> is a list of type classes written as \<open>s = {c\<^sub>1, \<dots>, c\<^sub>m}\<close>, it
 | 
46  | 
represents symbolic intersection. Notationally, the curly braces are omitted  | 
|
47  | 
for singleton intersections, i.e.\ any class \<open>c\<close> may be read as a sort  | 
|
48  | 
  \<open>{c}\<close>. The ordering on type classes is extended to sorts according to the
 | 
|
49  | 
  meaning of intersections: \<open>{c\<^sub>1, \<dots> c\<^sub>m} \<subseteq> {d\<^sub>1, \<dots>, d\<^sub>n}\<close> iff \<open>\<forall>j. \<exists>i. c\<^sub>i \<subseteq>
 | 
|
50  | 
  d\<^sub>j\<close>. The empty intersection \<open>{}\<close> refers to the universal sort, which is the
 | 
|
51  | 
  largest element wrt.\ the sort order. Thus \<open>{}\<close> represents the ``full
 | 
|
52  | 
sort'', not the empty one! The intersection of all (finitely many) classes  | 
|
53  | 
declared in the current theory is the least element wrt.\ the sort ordering.  | 
|
| 20480 | 54  | 
|
| 61416 | 55  | 
\<^medskip>  | 
| 61854 | 56  | 
A \<^emph>\<open>fixed type variable\<close> is a pair of a basic name (starting with a \<open>'\<close>  | 
57  | 
  character) and a sort constraint, e.g.\ \<open>('a, s)\<close> which is usually printed
 | 
|
58  | 
as \<open>\<alpha>\<^sub>s\<close>. A \<^emph>\<open>schematic type variable\<close> is a pair of an indexname and a sort  | 
|
59  | 
  constraint, e.g.\ \<open>(('a, 0), s)\<close> which is usually printed as \<open>?\<alpha>\<^sub>s\<close>.
 | 
|
| 20451 | 60  | 
|
| 61854 | 61  | 
Note that \<^emph>\<open>all\<close> syntactic components contribute to the identity of type  | 
62  | 
variables: basic name, index, and sort constraint. The core logic handles  | 
|
63  | 
type variables with the same name but different sorts as different, although  | 
|
64  | 
the type-inference layer (which is outside the core) rejects anything like  | 
|
65  | 
that.  | 
|
| 20451 | 66  | 
|
| 61854 | 67  | 
A \<^emph>\<open>type constructor\<close> \<open>\<kappa>\<close> is a \<open>k\<close>-ary operator on types declared in the  | 
68  | 
theory. Type constructor application is written postfix as \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>k)\<kappa>\<close>.  | 
|
69  | 
For \<open>k = 0\<close> the argument tuple is omitted, e.g.\ \<open>prop\<close> instead of \<open>()prop\<close>.  | 
|
70  | 
For \<open>k = 1\<close> the parentheses are omitted, e.g.\ \<open>\<alpha> list\<close> instead of  | 
|
71  | 
\<open>(\<alpha>)list\<close>. Further notation is provided for specific constructors, notably  | 
|
72  | 
the right-associative infix \<open>\<alpha> \<Rightarrow> \<beta>\<close> instead of \<open>(\<alpha>, \<beta>)fun\<close>.  | 
|
| 65446 | 73  | 
|
| 61854 | 74  | 
The logical category \<^emph>\<open>type\<close> is defined inductively over type variables and  | 
75  | 
type constructors as follows: \<open>\<tau> = \<alpha>\<^sub>s | ?\<alpha>\<^sub>s | (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close>.  | 
|
| 20451 | 76  | 
|
| 61854 | 77  | 
A \<^emph>\<open>type abbreviation\<close> is a syntactic definition \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close> of an  | 
78  | 
arbitrary type expression \<open>\<tau>\<close> over variables \<open>\<^vec>\<alpha>\<close>. Type abbreviations  | 
|
79  | 
appear as type constructors in the syntax, but are expanded before entering  | 
|
80  | 
the logical core.  | 
|
| 20480 | 81  | 
|
| 61854 | 82  | 
A \<^emph>\<open>type arity\<close> declares the image behavior of a type constructor wrt.\ the  | 
83  | 
algebra of sorts: \<open>\<kappa> :: (s\<^sub>1, \<dots>, s\<^sub>k)s\<close> means that \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close> is of  | 
|
84  | 
sort \<open>s\<close> if every argument type \<open>\<tau>\<^sub>i\<close> is of sort \<open>s\<^sub>i\<close>. Arity declarations  | 
|
85  | 
are implicitly completed, i.e.\ \<open>\<kappa> :: (\<^vec>s)c\<close> entails \<open>\<kappa> ::  | 
|
| 61493 | 86  | 
(\<^vec>s)c'\<close> for any \<open>c' \<supseteq> c\<close>.  | 
| 20491 | 87  | 
|
| 61416 | 88  | 
\<^medskip>  | 
| 61854 | 89  | 
The sort algebra is always maintained as \<^emph>\<open>coregular\<close>, which means that type  | 
90  | 
arities are consistent with the subclass relation: for any type constructor  | 
|
91  | 
\<open>\<kappa>\<close>, and classes \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close>, and arities \<open>\<kappa> :: (\<^vec>s\<^sub>1)c\<^sub>1\<close> and \<open>\<kappa> ::  | 
|
92  | 
(\<^vec>s\<^sub>2)c\<^sub>2\<close> holds \<open>\<^vec>s\<^sub>1 \<subseteq> \<^vec>s\<^sub>2\<close> component-wise.  | 
|
| 20451 | 93  | 
|
| 20491 | 94  | 
The key property of a coregular order-sorted algebra is that sort  | 
| 20537 | 95  | 
constraints can be solved in a most general fashion: for each type  | 
| 61854 | 96  | 
constructor \<open>\<kappa>\<close> and sort \<open>s\<close> there is a most general vector of argument  | 
97  | 
sorts \<open>(s\<^sub>1, \<dots>, s\<^sub>k)\<close> such that a type scheme \<open>(\<alpha>\<^bsub>s\<^sub>1\<^esub>, \<dots>, \<alpha>\<^bsub>s\<^sub>k\<^esub>)\<kappa>\<close> is of  | 
|
98  | 
sort \<open>s\<close>. Consequently, type unification has most general solutions (modulo  | 
|
99  | 
equivalence of sorts), so type-inference produces primary types as expected  | 
|
| 76987 | 100  | 
\<^cite>\<open>"nipkow-prehofer"\<close>.  | 
| 58618 | 101  | 
\<close>  | 
| 20451 | 102  | 
|
| 58618 | 103  | 
text %mlref \<open>  | 
| 20480 | 104  | 
  \begin{mldecls}
 | 
| 73764 | 105  | 
  @{define_ML_type class = string} \\
 | 
106  | 
  @{define_ML_type sort = "class list"} \\
 | 
|
107  | 
  @{define_ML_type arity = "string * sort list * sort"} \\
 | 
|
108  | 
  @{define_ML_type typ} \\
 | 
|
109  | 
  @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\
 | 
|
110  | 
  @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
 | 
|
| 20547 | 111  | 
  \end{mldecls}
 | 
112  | 
  \begin{mldecls}
 | 
|
| 73764 | 113  | 
  @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
 | 
114  | 
  @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
 | 
|
115  | 
  @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\
 | 
|
116  | 
  @{define_ML Sign.add_type_abbrev: "Proof.context ->
 | 
|
| 
42401
 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 
wenzelm 
parents: 
40255 
diff
changeset
 | 
117  | 
binding * string list * typ -> theory -> theory"} \\  | 
| 73764 | 118  | 
  @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
 | 
119  | 
  @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
 | 
|
120  | 
  @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
 | 
|
| 20480 | 121  | 
  \end{mldecls}
 | 
122  | 
||
| 69597 | 123  | 
\<^descr> Type \<^ML_type>\<open>class\<close> represents type classes.  | 
| 20480 | 124  | 
|
| 69597 | 125  | 
\<^descr> Type \<^ML_type>\<open>sort\<close> represents sorts, i.e.\ finite intersections of  | 
126  | 
classes. The empty list \<^ML>\<open>[]: sort\<close> refers to the empty class  | 
|
| 61854 | 127  | 
intersection, i.e.\ the ``full sort''.  | 
| 20451 | 128  | 
|
| 69597 | 129  | 
\<^descr> Type \<^ML_type>\<open>arity\<close> represents type arities. A triple \<open>(\<kappa>, \<^vec>s, s)  | 
| 61854 | 130  | 
: arity\<close> represents \<open>\<kappa> :: (\<^vec>s)s\<close> as described above.  | 
| 20480 | 131  | 
|
| 69597 | 132  | 
\<^descr> Type \<^ML_type>\<open>typ\<close> represents types; this is a datatype with constructors  | 
133  | 
\<^ML>\<open>TFree\<close>, \<^ML>\<open>TVar\<close>, \<^ML>\<open>Type\<close>.  | 
|
| 20480 | 134  | 
|
| 69597 | 135  | 
\<^descr> \<^ML>\<open>Term.map_atyps\<close>~\<open>f \<tau>\<close> applies the mapping \<open>f\<close> to all atomic types  | 
136  | 
(\<^ML>\<open>TFree\<close>, \<^ML>\<open>TVar\<close>) occurring in \<open>\<tau>\<close>.  | 
|
| 20519 | 137  | 
|
| 69597 | 138  | 
\<^descr> \<^ML>\<open>Term.fold_atyps\<close>~\<open>f \<tau>\<close> iterates the operation \<open>f\<close> over all  | 
139  | 
occurrences of atomic types (\<^ML>\<open>TFree\<close>, \<^ML>\<open>TVar\<close>) in \<open>\<tau>\<close>; the type  | 
|
| 61854 | 140  | 
structure is traversed from left to right.  | 
| 20494 | 141  | 
|
| 69597 | 142  | 
\<^descr> \<^ML>\<open>Sign.subsort\<close>~\<open>thy (s\<^sub>1, s\<^sub>2)\<close> tests the subsort relation \<open>s\<^sub>1 \<subseteq>  | 
| 61854 | 143  | 
s\<^sub>2\<close>.  | 
| 20480 | 144  | 
|
| 69597 | 145  | 
\<^descr> \<^ML>\<open>Sign.of_sort\<close>~\<open>thy (\<tau>, s)\<close> tests whether type \<open>\<tau>\<close> is of sort \<open>s\<close>.  | 
| 20480 | 146  | 
|
| 69597 | 147  | 
\<^descr> \<^ML>\<open>Sign.add_type\<close>~\<open>ctxt (\<kappa>, k, mx)\<close> declares a new type constructors \<open>\<kappa>\<close>  | 
| 61854 | 148  | 
with \<open>k\<close> arguments and optional mixfix syntax.  | 
| 20451 | 149  | 
|
| 69597 | 150  | 
\<^descr> \<^ML>\<open>Sign.add_type_abbrev\<close>~\<open>ctxt (\<kappa>, \<^vec>\<alpha>, \<tau>)\<close> defines a new type  | 
| 61854 | 151  | 
abbreviation \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close>.  | 
| 20480 | 152  | 
|
| 69597 | 153  | 
\<^descr> \<^ML>\<open>Sign.primitive_class\<close>~\<open>(c, [c\<^sub>1, \<dots>, c\<^sub>n])\<close> declares a new class \<open>c\<close>,  | 
| 61854 | 154  | 
together with class relations \<open>c \<subseteq> c\<^sub>i\<close>, for \<open>i = 1, \<dots>, n\<close>.  | 
| 20480 | 155  | 
|
| 69597 | 156  | 
\<^descr> \<^ML>\<open>Sign.primitive_classrel\<close>~\<open>(c\<^sub>1, c\<^sub>2)\<close> declares the class relation  | 
| 61854 | 157  | 
\<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close>.  | 
| 20480 | 158  | 
|
| 69597 | 159  | 
\<^descr> \<^ML>\<open>Sign.primitive_arity\<close>~\<open>(\<kappa>, \<^vec>s, s)\<close> declares the arity \<open>\<kappa> ::  | 
| 61854 | 160  | 
(\<^vec>s)s\<close>.  | 
| 58618 | 161  | 
\<close>  | 
| 20437 | 162  | 
|
| 58618 | 163  | 
text %mlantiq \<open>  | 
| 39832 | 164  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 165  | 
  @{ML_antiquotation_def "class"} & : & \<open>ML_antiquotation\<close> \\
 | 
166  | 
  @{ML_antiquotation_def "sort"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
167  | 
  @{ML_antiquotation_def "type_name"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
168  | 
  @{ML_antiquotation_def "type_abbrev"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
169  | 
  @{ML_antiquotation_def "nonterminal"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
170  | 
  @{ML_antiquotation_def "typ"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
171  | 
  @{ML_antiquotation_def "Type"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
172  | 
  @{ML_antiquotation_def "Type_fn"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 39832 | 173  | 
  \end{matharray}
 | 
174  | 
||
| 69597 | 175  | 
\<^rail>\<open>  | 
| 67146 | 176  | 
  @@{ML_antiquotation class} embedded
 | 
| 39832 | 177  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
178  | 
  @@{ML_antiquotation sort} sort
 | 
| 39832 | 179  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
180  | 
  (@@{ML_antiquotation type_name} |
 | 
| 
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
181  | 
   @@{ML_antiquotation type_abbrev} |
 | 
| 67146 | 182  | 
   @@{ML_antiquotation nonterminal}) embedded
 | 
| 39832 | 183  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
184  | 
  @@{ML_antiquotation typ} type
 | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
185  | 
;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
186  | 
  (@@{ML_antiquotation Type} | @@{ML_antiquotation Type_fn}) embedded
 | 
| 69597 | 187  | 
\<close>  | 
| 39832 | 188  | 
|
| 69597 | 189  | 
  \<^descr> \<open>@{class c}\<close> inlines the internalized class \<open>c\<close> --- as \<^ML_type>\<open>string\<close>
 | 
| 39832 | 190  | 
literal.  | 
191  | 
||
| 69597 | 192  | 
  \<^descr> \<open>@{sort s}\<close> inlines the internalized sort \<open>s\<close> --- as \<^ML_type>\<open>string
 | 
193  | 
list\<close> literal.  | 
|
| 61854 | 194  | 
|
195  | 
  \<^descr> \<open>@{type_name c}\<close> inlines the internalized type constructor \<open>c\<close> --- as
 | 
|
| 69597 | 196  | 
\<^ML_type>\<open>string\<close> literal.  | 
| 61854 | 197  | 
|
198  | 
  \<^descr> \<open>@{type_abbrev c}\<close> inlines the internalized type abbreviation \<open>c\<close> --- as
 | 
|
| 69597 | 199  | 
\<^ML_type>\<open>string\<close> literal.  | 
| 61854 | 200  | 
|
201  | 
  \<^descr> \<open>@{nonterminal c}\<close> inlines the internalized syntactic type~/ grammar
 | 
|
| 69597 | 202  | 
nonterminal \<open>c\<close> --- as \<^ML_type>\<open>string\<close> literal.  | 
| 61854 | 203  | 
|
204  | 
  \<^descr> \<open>@{typ \<tau>}\<close> inlines the internalized type \<open>\<tau>\<close> --- as constructor term for
 | 
|
| 69597 | 205  | 
datatype \<^ML_type>\<open>typ\<close>.  | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
206  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
207  | 
  \<^descr> \<open>@{Type source}\<close> refers to embedded source text to produce a type
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
208  | 
constructor with name (formally checked) and arguments (inlined ML text).  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
209  | 
  The embedded \<open>source\<close> follows the syntax category @{syntax type_const}
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
210  | 
defined below.  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
211  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
212  | 
  \<^descr> \<open>@{Type_fn source}\<close> is similar to \<open>@{Type source}\<close>, but produces a partial
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
213  | 
ML function that matches against a type constructor pattern, following  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
214  | 
  syntax category @{syntax type_const_fn} below.
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
215  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
216  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
217  | 
\<^rail>\<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
218  | 
    @{syntax_def type_const}: @{syntax name} (@{syntax embedded_ml}*)
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
219  | 
;  | 
| 74880 | 220  | 
    @{syntax_def type_const_fn}: @{syntax type_const} @'=>' @{syntax embedded_ml}
 | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
221  | 
;  | 
| 74880 | 222  | 
    @{syntax_def embedded_ml}:
 | 
223  | 
      @'_' | @{syntax embedded} | @{syntax control_symbol} @{syntax embedded}
 | 
|
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
224  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
225  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
226  | 
  The text provided as @{syntax embedded_ml} is treated as regular Isabelle/ML
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
227  | 
source, possibly with nested antiquotations. ML text that only consists of a  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
228  | 
single antiquotation in compact control-cartouche notation, can be written  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
229  | 
without an extra level of nesting embedded text (see the last example  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
230  | 
below).  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
231  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
232  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
233  | 
text %mlex \<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
234  | 
Here are some minimal examples for type constructor antiquotations.  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
235  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
236  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
237  | 
ML \<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
238  | 
\<comment> \<open>type constructor without arguments:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
239  | 
val natT = \<^Type>\<open>nat\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
240  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
241  | 
\<comment> \<open>type constructor with arguments:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
242  | 
fun mk_funT (A, B) = \<^Type>\<open>fun A B\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
243  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
244  | 
\<comment> \<open>type constructor decomposition as partial function:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
245  | 
val dest_funT = \<^Type_fn>\<open>fun A B => \<open>(A, B)\<close>\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
246  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
247  | 
\<comment> \<open>nested type constructors:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
248  | 
fun mk_relT A = \<^Type>\<open>fun A \<^Type>\<open>fun A \<^Type>\<open>bool\<close>\<close>\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
249  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
250  | 
\<^assert> (mk_relT natT = \<^typ>\<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close>);  | 
| 58618 | 251  | 
\<close>  | 
| 39832 | 252  | 
|
| 20437 | 253  | 
|
| 58618 | 254  | 
section \<open>Terms \label{sec:terms}\<close>
 | 
| 18537 | 255  | 
|
| 58618 | 256  | 
text \<open>  | 
| 61854 | 257  | 
The language of terms is that of simply-typed \<open>\<lambda>\<close>-calculus with de-Bruijn  | 
| 76987 | 258  | 
indices for bound variables (cf.\ \<^cite>\<open>debruijn72\<close> or \<^cite>\<open>"paulson-ml2"\<close>), with the types being determined by the corresponding  | 
| 61854 | 259  | 
binders. In contrast, free variables and constants have an explicit name and  | 
260  | 
type in each occurrence.  | 
|
| 20514 | 261  | 
|
| 61416 | 262  | 
\<^medskip>  | 
| 61854 | 263  | 
A \<^emph>\<open>bound variable\<close> is a natural number \<open>b\<close>, which accounts for the number  | 
264  | 
of intermediate binders between the variable occurrence in the body and its  | 
|
265  | 
binding position. For example, the de-Bruijn term \<open>\<lambda>\<^bsub>bool\<^esub>. \<lambda>\<^bsub>bool\<^esub>. 1 \<and> 0\<close>  | 
|
266  | 
would correspond to \<open>\<lambda>x\<^bsub>bool\<^esub>. \<lambda>y\<^bsub>bool\<^esub>. x \<and> y\<close> in a named representation.  | 
|
267  | 
Note that a bound variable may be represented by different de-Bruijn indices  | 
|
268  | 
at different occurrences, depending on the nesting of abstractions.  | 
|
| 20537 | 269  | 
|
| 61854 | 270  | 
A \<^emph>\<open>loose variable\<close> is a bound variable that is outside the scope of local  | 
271  | 
binders. The types (and names) for loose variables can be managed as a  | 
|
272  | 
separate context, that is maintained as a stack of hypothetical binders. The  | 
|
273  | 
core logic operates on closed terms, without any loose variables.  | 
|
| 20514 | 274  | 
|
| 61854 | 275  | 
A \<^emph>\<open>fixed variable\<close> is a pair of a basic name and a type, e.g.\ \<open>(x, \<tau>)\<close>  | 
276  | 
which is usually printed \<open>x\<^sub>\<tau>\<close> here. A \<^emph>\<open>schematic variable\<close> is a pair of an  | 
|
277  | 
indexname and a type, e.g.\ \<open>((x, 0), \<tau>)\<close> which is likewise printed as  | 
|
278  | 
\<open>?x\<^sub>\<tau>\<close>.  | 
|
| 20491 | 279  | 
|
| 61416 | 280  | 
\<^medskip>  | 
| 61854 | 281  | 
A \<^emph>\<open>constant\<close> is a pair of a basic name and a type, e.g.\ \<open>(c, \<tau>)\<close> which is  | 
282  | 
usually printed as \<open>c\<^sub>\<tau>\<close> here. Constants are declared in the context as  | 
|
283  | 
polymorphic families \<open>c :: \<sigma>\<close>, meaning that all substitution instances \<open>c\<^sub>\<tau>\<close>  | 
|
284  | 
for \<open>\<tau> = \<sigma>\<vartheta>\<close> are valid.  | 
|
| 20514 | 285  | 
|
| 61854 | 286  | 
The vector of \<^emph>\<open>type arguments\<close> of constant \<open>c\<^sub>\<tau>\<close> wrt.\ the declaration \<open>c  | 
287  | 
  :: \<sigma>\<close> is defined as the codomain of the matcher \<open>\<vartheta> = {?\<alpha>\<^sub>1 \<mapsto> \<tau>\<^sub>1,
 | 
|
288  | 
\<dots>, ?\<alpha>\<^sub>n \<mapsto> \<tau>\<^sub>n}\<close> presented in canonical order \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)\<close>, corresponding  | 
|
289  | 
to the left-to-right occurrences of the \<open>\<alpha>\<^sub>i\<close> in \<open>\<sigma>\<close>. Within a given theory  | 
|
290  | 
context, there is a one-to-one correspondence between any constant \<open>c\<^sub>\<tau>\<close> and  | 
|
291  | 
the application \<open>c(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)\<close> of its type arguments. For example, with  | 
|
292  | 
\<open>plus :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>\<close>, the instance \<open>plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>\<close> corresponds to  | 
|
| 61493 | 293  | 
\<open>plus(nat)\<close>.  | 
| 20480 | 294  | 
|
| 61854 | 295  | 
Constant declarations \<open>c :: \<sigma>\<close> may contain sort constraints for type  | 
296  | 
variables in \<open>\<sigma>\<close>. These are observed by type-inference as expected, but  | 
|
297  | 
\<^emph>\<open>ignored\<close> by the core logic. This means the primitive logic is able to  | 
|
298  | 
reason with instances of polymorphic constants that the user-level  | 
|
299  | 
type-checker would reject due to violation of type class restrictions.  | 
|
| 20480 | 300  | 
|
| 61416 | 301  | 
\<^medskip>  | 
| 61854 | 302  | 
An \<^emph>\<open>atomic term\<close> is either a variable or constant. The logical category  | 
303  | 
\<^emph>\<open>term\<close> is defined inductively over atomic terms, with abstraction and  | 
|
304  | 
application as follows: \<open>t = b | x\<^sub>\<tau> | ?x\<^sub>\<tau> | c\<^sub>\<tau> | \<lambda>\<^sub>\<tau>. t | t\<^sub>1 t\<^sub>2\<close>.  | 
|
305  | 
Parsing and printing takes care of converting between an external  | 
|
306  | 
representation with named bound variables. Subsequently, we shall use the  | 
|
307  | 
latter notation instead of internal de-Bruijn representation.  | 
|
| 20498 | 308  | 
|
| 61854 | 309  | 
The inductive relation \<open>t :: \<tau>\<close> assigns a (unique) type to a term according  | 
310  | 
to the structure of atomic terms, abstractions, and applications:  | 
|
| 20498 | 311  | 
\[  | 
| 61493 | 312  | 
  \infer{\<open>a\<^sub>\<tau> :: \<tau>\<close>}{}
 | 
| 20498 | 313  | 
\qquad  | 
| 61493 | 314  | 
  \infer{\<open>(\<lambda>x\<^sub>\<tau>. t) :: \<tau> \<Rightarrow> \<sigma>\<close>}{\<open>t :: \<sigma>\<close>}
 | 
| 20501 | 315  | 
\qquad  | 
| 61493 | 316  | 
  \infer{\<open>t u :: \<sigma>\<close>}{\<open>t :: \<tau> \<Rightarrow> \<sigma>\<close> & \<open>u :: \<tau>\<close>}
 | 
| 20498 | 317  | 
\]  | 
| 61477 | 318  | 
A \<^emph>\<open>well-typed term\<close> is a term that can be typed according to these rules.  | 
| 20498 | 319  | 
|
| 61854 | 320  | 
Typing information can be omitted: type-inference is able to reconstruct the  | 
321  | 
most general type of a raw term, while assigning most general types to all  | 
|
322  | 
of its variables and constants. Type-inference depends on a context of type  | 
|
323  | 
constraints for fixed variables, and declarations for polymorphic constants.  | 
|
| 20514 | 324  | 
|
| 20537 | 325  | 
The identity of atomic terms consists both of the name and the type  | 
| 61854 | 326  | 
component. This means that different variables \<open>x\<^bsub>\<tau>\<^sub>1\<^esub>\<close> and \<open>x\<^bsub>\<tau>\<^sub>2\<^esub>\<close> may  | 
327  | 
become the same after type instantiation. Type-inference rejects variables  | 
|
328  | 
of the same name, but different types. In contrast, mixed instances of  | 
|
| 34929 | 329  | 
polymorphic constants occur routinely.  | 
| 20514 | 330  | 
|
| 61416 | 331  | 
\<^medskip>  | 
| 61854 | 332  | 
The \<^emph>\<open>hidden polymorphism\<close> of a term \<open>t :: \<sigma>\<close> is the set of type variables  | 
333  | 
occurring in \<open>t\<close>, but not in its type \<open>\<sigma>\<close>. This means that the term  | 
|
334  | 
implicitly depends on type arguments that are not accounted in the result  | 
|
335  | 
type, i.e.\ there are different type instances \<open>t\<vartheta> :: \<sigma>\<close> and  | 
|
336  | 
\<open>t\<vartheta>' :: \<sigma>\<close> with the same type. This slightly pathological  | 
|
337  | 
situation notoriously demands additional care.  | 
|
| 20514 | 338  | 
|
| 61416 | 339  | 
\<^medskip>  | 
| 61854 | 340  | 
A \<^emph>\<open>term abbreviation\<close> is a syntactic definition \<open>c\<^sub>\<sigma> \<equiv> t\<close> of a closed term  | 
341  | 
\<open>t\<close> of type \<open>\<sigma>\<close>, without any hidden polymorphism. A term abbreviation looks  | 
|
342  | 
like a constant in the syntax, but is expanded before entering the logical  | 
|
343  | 
core. Abbreviations are usually reverted when printing terms, using \<open>t \<rightarrow>  | 
|
344  | 
c\<^sub>\<sigma>\<close> as rules for higher-order rewriting.  | 
|
| 20519 | 345  | 
|
| 61416 | 346  | 
\<^medskip>  | 
| 61854 | 347  | 
Canonical operations on \<open>\<lambda>\<close>-terms include \<open>\<alpha>\<beta>\<eta>\<close>-conversion: \<open>\<alpha>\<close>-conversion  | 
348  | 
refers to capture-free renaming of bound variables; \<open>\<beta>\<close>-conversion contracts  | 
|
349  | 
an abstraction applied to an argument term, substituting the argument in the  | 
|
350  | 
body: \<open>(\<lambda>x. b)a\<close> becomes \<open>b[a/x]\<close>; \<open>\<eta>\<close>-conversion contracts vacuous  | 
|
351  | 
application-abstraction: \<open>\<lambda>x. f x\<close> becomes \<open>f\<close>, provided that the bound  | 
|
352  | 
variable does not occur in \<open>f\<close>.  | 
|
| 20519 | 353  | 
|
| 61854 | 354  | 
Terms are normally treated modulo \<open>\<alpha>\<close>-conversion, which is implicit in the  | 
355  | 
de-Bruijn representation. Names for bound variables in abstractions are  | 
|
356  | 
maintained separately as (meaningless) comments, mostly for parsing and  | 
|
357  | 
printing. Full \<open>\<alpha>\<beta>\<eta>\<close>-conversion is commonplace in various standard  | 
|
358  | 
  operations (\secref{sec:obj-rules}) that are based on higher-order
 | 
|
359  | 
unification and matching.  | 
|
| 58618 | 360  | 
\<close>  | 
| 18537 | 361  | 
|
| 58618 | 362  | 
text %mlref \<open>  | 
| 20514 | 363  | 
  \begin{mldecls}
 | 
| 73764 | 364  | 
  @{define_ML_type term} \\
 | 
365  | 
  @{define_ML_infix "aconv": "term * term -> bool"} \\
 | 
|
366  | 
  @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\
 | 
|
367  | 
  @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
 | 
|
368  | 
  @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\
 | 
|
369  | 
  @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
 | 
|
| 20547 | 370  | 
  \end{mldecls}
 | 
371  | 
  \begin{mldecls}
 | 
|
| 73764 | 372  | 
  @{define_ML fastype_of: "term -> typ"} \\
 | 
373  | 
  @{define_ML lambda: "term -> term -> term"} \\
 | 
|
374  | 
  @{define_ML betapply: "term * term -> term"} \\
 | 
|
375  | 
  @{define_ML incr_boundvars: "int -> term -> term"} \\
 | 
|
376  | 
  @{define_ML Sign.declare_const: "Proof.context ->
 | 
|
| 
42401
 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 
wenzelm 
parents: 
40255 
diff
changeset
 | 
377  | 
(binding * typ) * mixfix -> theory -> term * theory"} \\  | 
| 73764 | 378  | 
  @{define_ML Sign.add_abbrev: "string -> binding * term ->
 | 
| 
24972
 
acafb18a47dc
replaced Sign.add_consts_i by Sign.declare_const;
 
wenzelm 
parents: 
24828 
diff
changeset
 | 
379  | 
theory -> (term * term) * theory"} \\  | 
| 73764 | 380  | 
  @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
 | 
381  | 
  @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
 | 
|
| 20514 | 382  | 
  \end{mldecls}
 | 
| 18537 | 383  | 
|
| 69597 | 384  | 
\<^descr> Type \<^ML_type>\<open>term\<close> represents de-Bruijn terms, with comments in  | 
| 61854 | 385  | 
abstractions, and explicitly named free variables and constants; this is a  | 
| 73764 | 386  | 
  datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML
 | 
387  | 
  Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}.
 | 
|
| 20519 | 388  | 
|
| 69597 | 389  | 
\<^descr> \<open>t\<close>~\<^ML_text>\<open>aconv\<close>~\<open>u\<close> checks \<open>\<alpha>\<close>-equivalence of two terms. This is the  | 
390  | 
basic equality relation on type \<^ML_type>\<open>term\<close>; raw datatype equality  | 
|
| 61854 | 391  | 
should only be used for operations related to parsing or printing!  | 
| 20537 | 392  | 
|
| 69597 | 393  | 
\<^descr> \<^ML>\<open>Term.map_types\<close>~\<open>f t\<close> applies the mapping \<open>f\<close> to all types occurring  | 
| 61854 | 394  | 
in \<open>t\<close>.  | 
395  | 
||
| 69597 | 396  | 
\<^descr> \<^ML>\<open>Term.fold_types\<close>~\<open>f t\<close> iterates the operation \<open>f\<close> over all  | 
| 61854 | 397  | 
occurrences of types in \<open>t\<close>; the term structure is traversed from left to  | 
398  | 
right.  | 
|
| 20519 | 399  | 
|
| 69597 | 400  | 
\<^descr> \<^ML>\<open>Term.map_aterms\<close>~\<open>f t\<close> applies the mapping \<open>f\<close> to all atomic terms  | 
401  | 
(\<^ML>\<open>Bound\<close>, \<^ML>\<open>Free\<close>, \<^ML>\<open>Var\<close>, \<^ML>\<open>Const\<close>) occurring in \<open>t\<close>.  | 
|
| 20537 | 402  | 
|
| 69597 | 403  | 
\<^descr> \<^ML>\<open>Term.fold_aterms\<close>~\<open>f t\<close> iterates the operation \<open>f\<close> over all  | 
404  | 
occurrences of atomic terms (\<^ML>\<open>Bound\<close>, \<^ML>\<open>Free\<close>, \<^ML>\<open>Var\<close>, \<^ML>\<open>Const\<close>) in \<open>t\<close>; the term structure is traversed from left to right.  | 
|
| 20519 | 405  | 
|
| 69597 | 406  | 
\<^descr> \<^ML>\<open>fastype_of\<close>~\<open>t\<close> determines the type of a well-typed term. This  | 
| 61854 | 407  | 
operation is relatively slow, despite the omission of any sanity checks.  | 
| 20519 | 408  | 
|
| 69597 | 409  | 
\<^descr> \<^ML>\<open>lambda\<close>~\<open>a b\<close> produces an abstraction \<open>\<lambda>a. b\<close>, where occurrences of  | 
| 61854 | 410  | 
the atomic term \<open>a\<close> in the body \<open>b\<close> are replaced by bound variables.  | 
| 20519 | 411  | 
|
| 69597 | 412  | 
\<^descr> \<^ML>\<open>betapply\<close>~\<open>(t, u)\<close> produces an application \<open>t u\<close>, with topmost  | 
| 61854 | 413  | 
\<open>\<beta>\<close>-conversion if \<open>t\<close> is an abstraction.  | 
| 20519 | 414  | 
|
| 69597 | 415  | 
\<^descr> \<^ML>\<open>incr_boundvars\<close>~\<open>j\<close> increments a term's dangling bound variables by  | 
| 61854 | 416  | 
the offset \<open>j\<close>. This is required when moving a subterm into a context where  | 
417  | 
it is enclosed by a different number of abstractions. Bound variables with a  | 
|
418  | 
matching abstraction are unaffected.  | 
|
| 42934 | 419  | 
|
| 69597 | 420  | 
\<^descr> \<^ML>\<open>Sign.declare_const\<close>~\<open>ctxt ((c, \<sigma>), mx)\<close> declares a new constant \<open>c ::  | 
| 61854 | 421  | 
\<sigma>\<close> with optional mixfix syntax.  | 
| 20519 | 422  | 
|
| 69597 | 423  | 
\<^descr> \<^ML>\<open>Sign.add_abbrev\<close>~\<open>print_mode (c, t)\<close> introduces a new term  | 
| 61854 | 424  | 
abbreviation \<open>c \<equiv> t\<close>.  | 
| 20519 | 425  | 
|
| 69597 | 426  | 
\<^descr> \<^ML>\<open>Sign.const_typargs\<close>~\<open>thy (c, \<tau>)\<close> and \<^ML>\<open>Sign.const_instance\<close>~\<open>thy  | 
| 61854 | 427  | 
(c, [\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n])\<close> convert between two representations of polymorphic  | 
428  | 
constants: full type instance vs.\ compact type arguments form.  | 
|
| 58618 | 429  | 
\<close>  | 
| 18537 | 430  | 
|
| 58618 | 431  | 
text %mlantiq \<open>  | 
| 39832 | 432  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 433  | 
  @{ML_antiquotation_def "const_name"} & : & \<open>ML_antiquotation\<close> \\
 | 
434  | 
  @{ML_antiquotation_def "const_abbrev"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
435  | 
  @{ML_antiquotation_def "const"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
436  | 
  @{ML_antiquotation_def "term"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
437  | 
  @{ML_antiquotation_def "prop"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
438  | 
  @{ML_antiquotation_def "Const"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
439  | 
  @{ML_antiquotation_def "Const_"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
440  | 
  @{ML_antiquotation_def "Const_fn"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 39832 | 441  | 
  \end{matharray}
 | 
442  | 
||
| 69597 | 443  | 
\<^rail>\<open>  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
444  | 
  (@@{ML_antiquotation const_name} |
 | 
| 67146 | 445  | 
   @@{ML_antiquotation const_abbrev}) embedded
 | 
| 39832 | 446  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
447  | 
  @@{ML_antiquotation const} ('(' (type + ',') ')')?
 | 
| 39832 | 448  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
449  | 
  @@{ML_antiquotation term} term
 | 
| 39832 | 450  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
451  | 
  @@{ML_antiquotation prop} prop
 | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
452  | 
;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
453  | 
  (@@{ML_antiquotation Const} | @@{ML_antiquotation Const_} | @@{ML_antiquotation Const_fn})
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
454  | 
embedded  | 
| 69597 | 455  | 
\<close>  | 
| 39832 | 456  | 
|
| 61854 | 457  | 
  \<^descr> \<open>@{const_name c}\<close> inlines the internalized logical constant name \<open>c\<close> ---
 | 
| 69597 | 458  | 
as \<^ML_type>\<open>string\<close> literal.  | 
| 61854 | 459  | 
|
460  | 
  \<^descr> \<open>@{const_abbrev c}\<close> inlines the internalized abbreviated constant name \<open>c\<close>
 | 
|
| 69597 | 461  | 
--- as \<^ML_type>\<open>string\<close> literal.  | 
| 39832 | 462  | 
|
| 61854 | 463  | 
  \<^descr> \<open>@{const c(\<^vec>\<tau>)}\<close> inlines the internalized constant \<open>c\<close> with precise
 | 
| 69597 | 464  | 
type instantiation in the sense of \<^ML>\<open>Sign.const_instance\<close> --- as \<^ML>\<open>Const\<close> constructor term for datatype \<^ML_type>\<open>term\<close>.  | 
| 39832 | 465  | 
|
| 61854 | 466  | 
  \<^descr> \<open>@{term t}\<close> inlines the internalized term \<open>t\<close> --- as constructor term for
 | 
| 69597 | 467  | 
datatype \<^ML_type>\<open>term\<close>.  | 
| 39832 | 468  | 
|
| 61854 | 469  | 
  \<^descr> \<open>@{prop \<phi>}\<close> inlines the internalized proposition \<open>\<phi>\<close> --- as constructor
 | 
| 69597 | 470  | 
term for datatype \<^ML_type>\<open>term\<close>.  | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
471  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
472  | 
  \<^descr> \<open>@{Const source}\<close> refers to embedded source text to produce a term
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
473  | 
constructor with name (formally checked), type arguments and term arguments  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
474  | 
(inlined ML text). The embedded \<open>source\<close> follows the syntax category  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
475  | 
  @{syntax term_const} defined below, using @{syntax embedded_ml} from
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
476  | 
  antiquotation @{ML_antiquotation Type} (\secref{sec:types}).
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
477  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
478  | 
  \<^descr> \<open>@{Const_ source}\<close> is similar to \<open>@{Const source}\<close> for patterns instead of
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
479  | 
closed values. This distinction is required due to redundant type  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
480  | 
information within term constants: subtrees with duplicate ML pattern  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
481  | 
variable need to be suppressed (replaced by dummy patterns). The embedded \<open>source\<close> follows  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
482  | 
  the syntax category @{syntax term_const} defined below.
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
483  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
484  | 
  \<^descr> \<open>@{Const_fn source}\<close> is similar to \<open>@{Const_ source}\<close>, but produces a
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
485  | 
proper \<^verbatim>\<open>fn\<close> expression with function body. The embedded \<open>source\<close> follows  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
486  | 
  the syntax category @{syntax term_const_fn} defined below.
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
487  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
488  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
489  | 
\<^rail>\<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
490  | 
    @{syntax_def term_const}:
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
491  | 
      @{syntax name} (@{syntax embedded_ml}*) (@{syntax for_args})?
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
492  | 
;  | 
| 74880 | 493  | 
    @{syntax_def term_const_fn}:
 | 
| 74915 | 494  | 
      @{syntax term_const} @'=>' @{syntax embedded_ml}
 | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
495  | 
;  | 
| 74880 | 496  | 
    @{syntax_def for_args}: @'for' (@{syntax embedded_ml}+)
 | 
| 
74874
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
497  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
498  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
499  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
500  | 
text %mlex \<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
501  | 
Here are some minimal examples for term constant antiquotations. Type  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
502  | 
arguments for constants are analogous to type constructors  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
503  | 
  (\secref{sec:types}). Term arguments are different: a flexible number of
 | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
504  | 
arguments is inserted via curried applications (\<^ML>\<open>op $\<close>).  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
505  | 
\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
506  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
507  | 
ML \<open>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
508  | 
\<comment> \<open>constant without type argument:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
509  | 
fun mk_conj (A, B) = \<^Const>\<open>conj for A B\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
510  | 
val dest_conj = \<^Const_fn>\<open>conj for A B => \<open>(A, B)\<close>\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
511  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
512  | 
\<comment> \<open>constant with type argument:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
513  | 
fun mk_eq T (t, u) = \<^Const>\<open>HOL.eq T for t u\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
514  | 
val dest_eq = \<^Const_fn>\<open>HOL.eq T for t u => \<open>(T, (t, u))\<close>\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
515  | 
|
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
516  | 
\<comment> \<open>constant with variable number of term arguments:\<close>  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
517  | 
val c = Const (\<^const_name>\<open>conj\<close>, \<^typ>\<open>bool \<Rightarrow> bool \<Rightarrow> bool\<close>);  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
518  | 
val P = \<^term>\<open>P::bool\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
519  | 
val Q = \<^term>\<open>Q::bool\<close>;  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
520  | 
\<^assert> (\<^Const>\<open>conj\<close> = c);  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
521  | 
\<^assert> (\<^Const>\<open>conj for P\<close> = c $ P);  | 
| 
 
8baf2e8b16e2
more documentation about Type/Const antiquotations;
 
wenzelm 
parents: 
74604 
diff
changeset
 | 
522  | 
\<^assert> (\<^Const>\<open>conj for P Q\<close> = c $ P $ Q);  | 
| 58618 | 523  | 
\<close>  | 
| 39832 | 524  | 
|
| 18537 | 525  | 
|
| 58618 | 526  | 
section \<open>Theorems \label{sec:thms}\<close>
 | 
| 18537 | 527  | 
|
| 58618 | 528  | 
text \<open>  | 
| 61854 | 529  | 
A \<^emph>\<open>proposition\<close> is a well-typed term of type \<open>prop\<close>, a \<^emph>\<open>theorem\<close> is a  | 
530  | 
proven proposition (depending on a context of hypotheses and the background  | 
|
531  | 
theory). Primitive inferences include plain Natural Deduction rules for the  | 
|
532  | 
primary connectives \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> of the framework. There is also a builtin  | 
|
| 61493 | 533  | 
notion of equality/equivalence \<open>\<equiv>\<close>.  | 
| 58618 | 534  | 
\<close>  | 
| 20521 | 535  | 
|
| 29758 | 536  | 
|
| 58618 | 537  | 
subsection \<open>Primitive connectives and rules \label{sec:prim-rules}\<close>
 | 
| 18537 | 538  | 
|
| 58618 | 539  | 
text \<open>  | 
| 61854 | 540  | 
The theory \<open>Pure\<close> contains constant declarations for the primitive  | 
541  | 
connectives \<open>\<And>\<close>, \<open>\<Longrightarrow>\<close>, and \<open>\<equiv>\<close> of the logical framework, see  | 
|
542  | 
  \figref{fig:pure-connectives}. The derivability judgment \<open>A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B\<close>
 | 
|
543  | 
is defined inductively by the primitive inferences given in  | 
|
544  | 
  \figref{fig:prim-rules}, with the global restriction that the hypotheses
 | 
|
545  | 
must \<^emph>\<open>not\<close> contain any schematic variables. The builtin equality is  | 
|
546  | 
  conceptually axiomatized as shown in \figref{fig:pure-equality}, although
 | 
|
547  | 
the implementation works directly with derived inferences.  | 
|
| 20521 | 548  | 
|
549  | 
  \begin{figure}[htb]
 | 
|
550  | 
  \begin{center}
 | 
|
| 20501 | 551  | 
  \begin{tabular}{ll}
 | 
| 61493 | 552  | 
\<open>all :: (\<alpha> \<Rightarrow> prop) \<Rightarrow> prop\<close> & universal quantification (binder \<open>\<And>\<close>) \\  | 
553  | 
\<open>\<Longrightarrow> :: prop \<Rightarrow> prop \<Rightarrow> prop\<close> & implication (right associative infix) \\  | 
|
554  | 
\<open>\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> & equality relation (infix) \\  | 
|
| 20501 | 555  | 
  \end{tabular}
 | 
| 20537 | 556  | 
  \caption{Primitive connectives of Pure}\label{fig:pure-connectives}
 | 
| 20521 | 557  | 
  \end{center}
 | 
558  | 
  \end{figure}
 | 
|
| 18537 | 559  | 
|
| 20501 | 560  | 
  \begin{figure}[htb]
 | 
561  | 
  \begin{center}
 | 
|
| 20498 | 562  | 
\[  | 
| 61493 | 563  | 
  \infer[\<open>(axiom)\<close>]{\<open>\<turnstile> A\<close>}{\<open>A \<in> \<Theta>\<close>}
 | 
| 20498 | 564  | 
\qquad  | 
| 61493 | 565  | 
  \infer[\<open>(assume)\<close>]{\<open>A \<turnstile> A\<close>}{}
 | 
| 20498 | 566  | 
\]  | 
567  | 
\[  | 
|
| 61493 | 568  | 
  \infer[\<open>(\<And>\<hyphen>intro)\<close>]{\<open>\<Gamma> \<turnstile> \<And>x. B[x]\<close>}{\<open>\<Gamma> \<turnstile> B[x]\<close> & \<open>x \<notin> \<Gamma>\<close>}
 | 
| 20498 | 569  | 
\qquad  | 
| 61493 | 570  | 
  \infer[\<open>(\<And>\<hyphen>elim)\<close>]{\<open>\<Gamma> \<turnstile> B[a]\<close>}{\<open>\<Gamma> \<turnstile> \<And>x. B[x]\<close>}
 | 
| 20498 | 571  | 
\]  | 
572  | 
\[  | 
|
| 61493 | 573  | 
  \infer[\<open>(\<Longrightarrow>\<hyphen>intro)\<close>]{\<open>\<Gamma> - A \<turnstile> A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>}
 | 
| 20498 | 574  | 
\qquad  | 
| 61493 | 575  | 
  \infer[\<open>(\<Longrightarrow>\<hyphen>elim)\<close>]{\<open>\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B\<close>}{\<open>\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B\<close> & \<open>\<Gamma>\<^sub>2 \<turnstile> A\<close>}
 | 
| 20498 | 576  | 
\]  | 
| 20521 | 577  | 
  \caption{Primitive inferences of Pure}\label{fig:prim-rules}
 | 
578  | 
  \end{center}
 | 
|
579  | 
  \end{figure}
 | 
|
580  | 
||
581  | 
  \begin{figure}[htb]
 | 
|
582  | 
  \begin{center}
 | 
|
583  | 
  \begin{tabular}{ll}
 | 
|
| 61493 | 584  | 
\<open>\<turnstile> (\<lambda>x. b[x]) a \<equiv> b[a]\<close> & \<open>\<beta>\<close>-conversion \\  | 
585  | 
\<open>\<turnstile> x \<equiv> x\<close> & reflexivity \\  | 
|
586  | 
\<open>\<turnstile> x \<equiv> y \<Longrightarrow> P x \<Longrightarrow> P y\<close> & substitution \\  | 
|
587  | 
\<open>\<turnstile> (\<And>x. f x \<equiv> g x) \<Longrightarrow> f \<equiv> g\<close> & extensionality \\  | 
|
588  | 
\<open>\<turnstile> (A \<Longrightarrow> B) \<Longrightarrow> (B \<Longrightarrow> A) \<Longrightarrow> A \<equiv> B\<close> & logical equivalence \\  | 
|
| 20521 | 589  | 
  \end{tabular}
 | 
| 20542 | 590  | 
  \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality}
 | 
| 20501 | 591  | 
  \end{center}
 | 
592  | 
  \end{figure}
 | 
|
| 18537 | 593  | 
|
| 61854 | 594  | 
The introduction and elimination rules for \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> are analogous to  | 
595  | 
formation of dependently typed \<open>\<lambda>\<close>-terms representing the underlying proof  | 
|
596  | 
objects. Proof terms are irrelevant in the Pure logic, though; they cannot  | 
|
597  | 
occur within propositions. The system provides a runtime option to record  | 
|
| 52408 | 598  | 
explicit proof terms for primitive inferences, see also  | 
| 61854 | 599  | 
  \secref{sec:proof-terms}. Thus all three levels of \<open>\<lambda>\<close>-calculus become
 | 
| 76987 | 600  | 
explicit: \<open>\<Rightarrow>\<close> for terms, and \<open>\<And>/\<Longrightarrow>\<close> for proofs (cf.\ \<^cite>\<open>"Berghofer-Nipkow:2000:TPHOL"\<close>).  | 
| 20491 | 601  | 
|
| 61854 | 602  | 
Observe that locally fixed parameters (as in \<open>\<And>\<hyphen>intro\<close>) need not be recorded  | 
603  | 
in the hypotheses, because the simple syntactic types of Pure are always  | 
|
604  | 
inhabitable. ``Assumptions'' \<open>x :: \<tau>\<close> for type-membership are only present  | 
|
605  | 
as long as some \<open>x\<^sub>\<tau>\<close> occurs in the statement body.\<^footnote>\<open>This is the key  | 
|
| 76987 | 606  | 
difference to ``\<open>\<lambda>HOL\<close>'' in the PTS framework \<^cite>\<open>"Barendregt-Geuvers:2001"\<close>, where hypotheses \<open>x : A\<close> are treated uniformly  | 
| 61854 | 607  | 
for propositions and types.\<close>  | 
| 20501 | 608  | 
|
| 61416 | 609  | 
\<^medskip>  | 
| 61854 | 610  | 
The axiomatization of a theory is implicitly closed by forming all instances  | 
611  | 
of type and term variables: \<open>\<turnstile> A\<vartheta>\<close> holds for any substitution  | 
|
612  | 
instance of an axiom \<open>\<turnstile> A\<close>. By pushing substitutions through derivations  | 
|
613  | 
inductively, we also get admissible \<open>generalize\<close> and \<open>instantiate\<close> rules as  | 
|
614  | 
  shown in \figref{fig:subst-rules}.
 | 
|
| 20501 | 615  | 
|
616  | 
  \begin{figure}[htb]
 | 
|
617  | 
  \begin{center}
 | 
|
| 20498 | 618  | 
\[  | 
| 61493 | 619  | 
  \infer{\<open>\<Gamma> \<turnstile> B[?\<alpha>]\<close>}{\<open>\<Gamma> \<turnstile> B[\<alpha>]\<close> & \<open>\<alpha> \<notin> \<Gamma>\<close>}
 | 
| 20501 | 620  | 
\quad  | 
| 61493 | 621  | 
  \infer[\quad\<open>(generalize)\<close>]{\<open>\<Gamma> \<turnstile> B[?x]\<close>}{\<open>\<Gamma> \<turnstile> B[x]\<close> & \<open>x \<notin> \<Gamma>\<close>}
 | 
| 20498 | 622  | 
\]  | 
623  | 
\[  | 
|
| 61493 | 624  | 
  \infer{\<open>\<Gamma> \<turnstile> B[\<tau>]\<close>}{\<open>\<Gamma> \<turnstile> B[?\<alpha>]\<close>}
 | 
| 20501 | 625  | 
\quad  | 
| 61493 | 626  | 
  \infer[\quad\<open>(instantiate)\<close>]{\<open>\<Gamma> \<turnstile> B[t]\<close>}{\<open>\<Gamma> \<turnstile> B[?x]\<close>}
 | 
| 20498 | 627  | 
\]  | 
| 20501 | 628  | 
  \caption{Admissible substitution rules}\label{fig:subst-rules}
 | 
629  | 
  \end{center}
 | 
|
630  | 
  \end{figure}
 | 
|
| 18537 | 631  | 
|
| 61854 | 632  | 
Note that \<open>instantiate\<close> does not require an explicit side-condition, because  | 
633  | 
\<open>\<Gamma>\<close> may never contain schematic variables.  | 
|
| 20537 | 634  | 
|
| 61854 | 635  | 
In principle, variables could be substituted in hypotheses as well, but this  | 
636  | 
would disrupt the monotonicity of reasoning: deriving \<open>\<Gamma>\<vartheta> \<turnstile>  | 
|
637  | 
B\<vartheta>\<close> from \<open>\<Gamma> \<turnstile> B\<close> is correct, but \<open>\<Gamma>\<vartheta> \<supseteq> \<Gamma>\<close> does not  | 
|
638  | 
necessarily hold: the result belongs to a different proof context.  | 
|
| 20542 | 639  | 
|
| 70568 | 640  | 
\<^medskip> An \<^emph>\<open>oracle\<close> is a function that produces axioms on the fly. Logically,  | 
641  | 
  this is an instance of the \<open>axiom\<close> rule (\figref{fig:prim-rules}), but there
 | 
|
642  | 
is an operational difference. The inference kernel records oracle  | 
|
643  | 
invocations within derivations of theorems by a unique tag. This also  | 
|
644  | 
includes implicit type-class reasoning via the order-sorted algebra of class  | 
|
645  | 
  relations and type arities (see also @{command_ref instantiation} and
 | 
|
646  | 
  @{command_ref instance}).
 | 
|
| 20542 | 647  | 
|
| 61854 | 648  | 
Axiomatizations should be limited to the bare minimum, typically as part of  | 
649  | 
the initial logical basis of an object-logic formalization. Later on,  | 
|
650  | 
theories are usually developed in a strictly definitional fashion, by  | 
|
651  | 
stating only certain equalities over new constants.  | 
|
| 20542 | 652  | 
|
| 61854 | 653  | 
A \<^emph>\<open>simple definition\<close> consists of a constant declaration \<open>c :: \<sigma>\<close> together  | 
654  | 
with an axiom \<open>\<turnstile> c \<equiv> t\<close>, where \<open>t :: \<sigma>\<close> is a closed term without any hidden  | 
|
655  | 
polymorphism. The RHS may depend on further defined constants, but not \<open>c\<close>  | 
|
656  | 
itself. Definitions of functions may be presented as \<open>c \<^vec>x \<equiv> t\<close>  | 
|
657  | 
instead of the puristic \<open>c \<equiv> \<lambda>\<^vec>x. t\<close>.  | 
|
| 20542 | 658  | 
|
| 61854 | 659  | 
An \<^emph>\<open>overloaded definition\<close> consists of a collection of axioms for the same  | 
660  | 
constant, with zero or one equations \<open>c((\<^vec>\<alpha>)\<kappa>) \<equiv> t\<close> for each type  | 
|
661  | 
constructor \<open>\<kappa>\<close> (for distinct variables \<open>\<^vec>\<alpha>\<close>). The RHS may mention  | 
|
662  | 
previously defined constants as above, or arbitrary constants \<open>d(\<alpha>\<^sub>i)\<close> for  | 
|
663  | 
some \<open>\<alpha>\<^sub>i\<close> projected from \<open>\<^vec>\<alpha>\<close>. Thus overloaded definitions  | 
|
664  | 
essentially work by primitive recursion over the syntactic structure of a  | 
|
| 76987 | 665  | 
single type argument. See also \<^cite>\<open>\<open>\S4.3\<close> in  | 
666  | 
"Haftmann-Wenzel:2006:classes"\<close>.  | 
|
| 58618 | 667  | 
\<close>  | 
| 20498 | 668  | 
|
| 58618 | 669  | 
text %mlref \<open>  | 
| 20521 | 670  | 
  \begin{mldecls}
 | 
| 73764 | 671  | 
  @{define_ML Logic.all: "term -> term -> term"} \\
 | 
672  | 
  @{define_ML Logic.mk_implies: "term * term -> term"} \\
 | 
|
| 46253 | 673  | 
  \end{mldecls}
 | 
674  | 
  \begin{mldecls}
 | 
|
| 73764 | 675  | 
  @{define_ML_type ctyp} \\
 | 
676  | 
  @{define_ML_type cterm} \\
 | 
|
677  | 
  @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\
 | 
|
678  | 
  @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\
 | 
|
679  | 
  @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\
 | 
|
680  | 
  @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\
 | 
|
681  | 
  @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\
 | 
|
682  | 
  @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\
 | 
|
| 20547 | 683  | 
  \end{mldecls}
 | 
684  | 
  \begin{mldecls}
 | 
|
| 73764 | 685  | 
  @{define_ML_type thm} \\
 | 
686  | 
  @{define_ML Thm.transfer: "theory -> thm -> thm"} \\
 | 
|
687  | 
  @{define_ML Thm.assume: "cterm -> thm"} \\
 | 
|
688  | 
  @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
 | 
|
689  | 
  @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
 | 
|
690  | 
  @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\
 | 
|
691  | 
  @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\
 | 
|
| 74266 | 692  | 
  @{define_ML Thm.generalize: "Names.set * Names.set -> int -> thm -> thm"} \\
 | 
| 74282 | 693  | 
  @{define_ML Thm.instantiate: "ctyp TVars.table * cterm Vars.table -> thm -> thm"} \\
 | 
| 73764 | 694  | 
  @{define_ML Thm.add_axiom: "Proof.context ->
 | 
| 
42401
 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 
wenzelm 
parents: 
40255 
diff
changeset
 | 
695  | 
binding * term -> theory -> (string * thm) * theory"} \\  | 
| 73764 | 696  | 
  @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory ->
 | 
| 39821 | 697  | 
  (string * ('a -> thm)) * theory"} \\
 | 
| 73764 | 698  | 
  @{define_ML Thm.add_def: "Defs.context -> bool -> bool ->
 | 
| 
42401
 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 
wenzelm 
parents: 
40255 
diff
changeset
 | 
699  | 
binding * term -> theory -> (string * thm) * theory"} \\  | 
| 20547 | 700  | 
  \end{mldecls}
 | 
701  | 
  \begin{mldecls}
 | 
|
| 73764 | 702  | 
  @{define_ML Theory.add_deps: "Defs.context -> string ->
 | 
| 
61255
 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 
wenzelm 
parents: 
61246 
diff
changeset
 | 
703  | 
Defs.entry -> Defs.entry list -> theory -> theory"} \\  | 
| 77867 | 704  | 
  @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\
 | 
| 20521 | 705  | 
  \end{mldecls}
 | 
706  | 
||
| 69597 | 707  | 
\<^descr> \<^ML>\<open>Logic.all\<close>~\<open>a B\<close> produces a Pure quantification \<open>\<And>a. B\<close>, where  | 
| 61854 | 708  | 
occurrences of the atomic term \<open>a\<close> in the body proposition \<open>B\<close> are replaced  | 
| 69597 | 709  | 
by bound variables. (See also \<^ML>\<open>lambda\<close> on terms.)  | 
| 46253 | 710  | 
|
| 69597 | 711  | 
\<^descr> \<^ML>\<open>Logic.mk_implies\<close>~\<open>(A, B)\<close> produces a Pure implication \<open>A \<Longrightarrow> B\<close>.  | 
| 46253 | 712  | 
|
| 69597 | 713  | 
\<^descr> Types \<^ML_type>\<open>ctyp\<close> and \<^ML_type>\<open>cterm\<close> represent certified types and  | 
| 61854 | 714  | 
terms, respectively. These are abstract datatypes that guarantee that its  | 
715  | 
values have passed the full well-formedness (and well-typedness) checks,  | 
|
716  | 
relative to the declarations of type constructors, constants etc.\ in the  | 
|
| 69597 | 717  | 
background theory. The abstract types \<^ML_type>\<open>ctyp\<close> and \<^ML_type>\<open>cterm\<close>  | 
| 61854 | 718  | 
are part of the same inference kernel that is mainly responsible for  | 
| 69597 | 719  | 
\<^ML_type>\<open>thm\<close>. Thus syntactic operations on \<^ML_type>\<open>ctyp\<close> and \<^ML_type>\<open>cterm\<close> are located in the \<^ML_structure>\<open>Thm\<close> module, even though theorems  | 
| 61854 | 720  | 
are not yet involved at that stage.  | 
| 20542 | 721  | 
|
| 69597 | 722  | 
\<^descr> \<^ML>\<open>Thm.ctyp_of\<close>~\<open>ctxt \<tau>\<close> and \<^ML>\<open>Thm.cterm_of\<close>~\<open>ctxt t\<close> explicitly  | 
| 61854 | 723  | 
check types and terms, respectively. This also involves some basic  | 
724  | 
normalizations, such expansion of type and term abbreviations from the  | 
|
725  | 
underlying theory context. Full re-certification is relatively slow and  | 
|
726  | 
should be avoided in tight reasoning loops.  | 
|
| 20547 | 727  | 
|
| 69597 | 728  | 
\<^descr> \<^ML>\<open>Thm.apply\<close>, \<^ML>\<open>Thm.lambda\<close>, \<^ML>\<open>Thm.all\<close>, \<^ML>\<open>Drule.mk_implies\<close>  | 
| 61854 | 729  | 
etc.\ compose certified terms (or propositions) incrementally. This is  | 
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
730  | 
equivalent to \<^ML>\<open>Thm.cterm_of\<close> after unchecked \<^ML_infix>\<open>$\<close>, \<^ML>\<open>lambda\<close>,  | 
| 69597 | 731  | 
\<^ML>\<open>Logic.all\<close>, \<^ML>\<open>Logic.mk_implies\<close> etc., but there can be a big  | 
| 61854 | 732  | 
difference in performance when large existing entities are composed by a few  | 
733  | 
extra constructions on top. There are separate operations to decompose  | 
|
| 46253 | 734  | 
certified terms and theorems to produce certified terms again.  | 
| 20542 | 735  | 
|
| 69597 | 736  | 
\<^descr> Type \<^ML_type>\<open>thm\<close> represents proven propositions. This is an abstract  | 
| 61854 | 737  | 
datatype that guarantees that its values have been constructed by basic  | 
| 69597 | 738  | 
principles of the \<^ML_structure>\<open>Thm\<close> module. Every \<^ML_type>\<open>thm\<close> value  | 
| 61854 | 739  | 
  refers its background theory, cf.\ \secref{sec:context-theory}.
 | 
| 20542 | 740  | 
|
| 69597 | 741  | 
\<^descr> \<^ML>\<open>Thm.transfer\<close>~\<open>thy thm\<close> transfers the given theorem to a \<^emph>\<open>larger\<close>  | 
| 61854 | 742  | 
  theory, see also \secref{sec:context}. This formal adjustment of the
 | 
743  | 
background context has no logical significance, but is occasionally required  | 
|
744  | 
for formal reasons, e.g.\ when theorems that are imported from more basic  | 
|
745  | 
theories are used in the current situation.  | 
|
| 42933 | 746  | 
|
| 69597 | 747  | 
\<^descr> \<^ML>\<open>Thm.assume\<close>, \<^ML>\<open>Thm.forall_intr\<close>, \<^ML>\<open>Thm.forall_elim\<close>, \<^ML>\<open>Thm.implies_intr\<close>, and \<^ML>\<open>Thm.implies_elim\<close> correspond to the primitive  | 
| 61854 | 748  | 
  inferences of \figref{fig:prim-rules}.
 | 
| 20542 | 749  | 
|
| 69597 | 750  | 
\<^descr> \<^ML>\<open>Thm.generalize\<close>~\<open>(\<^vec>\<alpha>, \<^vec>x)\<close> corresponds to the  | 
| 61854 | 751  | 
  \<open>generalize\<close> rules of \figref{fig:subst-rules}. Here collections of type and
 | 
| 
74200
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
73765 
diff
changeset
 | 
752  | 
term variables are generalized simultaneously, specified by the given sets of  | 
| 
 
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
 
wenzelm 
parents: 
73765 
diff
changeset
 | 
753  | 
basic names.  | 
| 20521 | 754  | 
|
| 69597 | 755  | 
\<^descr> \<^ML>\<open>Thm.instantiate\<close>~\<open>(\<^vec>\<alpha>\<^sub>s, \<^vec>x\<^sub>\<tau>)\<close> corresponds to the  | 
| 61854 | 756  | 
  \<open>instantiate\<close> rules of \figref{fig:subst-rules}. Type variables are
 | 
757  | 
substituted before term variables. Note that the types in \<open>\<^vec>x\<^sub>\<tau>\<close> refer  | 
|
758  | 
to the instantiated versions.  | 
|
| 20542 | 759  | 
|
| 69597 | 760  | 
\<^descr> \<^ML>\<open>Thm.add_axiom\<close>~\<open>ctxt (name, A)\<close> declares an arbitrary proposition as  | 
| 61854 | 761  | 
axiom, and retrieves it as a theorem from the resulting theory, cf.\ \<open>axiom\<close>  | 
762  | 
  in \figref{fig:prim-rules}. Note that the low-level representation in the
 | 
|
763  | 
axiom table may differ slightly from the returned theorem.  | 
|
| 20542 | 764  | 
|
| 69597 | 765  | 
\<^descr> \<^ML>\<open>Thm.add_oracle\<close>~\<open>(binding, oracle)\<close> produces a named oracle rule,  | 
| 61854 | 766  | 
essentially generating arbitrary axioms on the fly, cf.\ \<open>axiom\<close> in  | 
767  | 
  \figref{fig:prim-rules}.
 | 
|
| 20521 | 768  | 
|
| 69597 | 769  | 
\<^descr> \<^ML>\<open>Thm.add_def\<close>~\<open>ctxt unchecked overloaded (name, c \<^vec>x \<equiv> t)\<close>  | 
| 61854 | 770  | 
states a definitional axiom for an existing constant \<open>c\<close>. Dependencies are  | 
| 69597 | 771  | 
recorded via \<^ML>\<open>Theory.add_deps\<close>, unless the \<open>unchecked\<close> option is set.  | 
| 61854 | 772  | 
Note that the low-level representation in the axiom table may differ  | 
773  | 
slightly from the returned theorem.  | 
|
| 20542 | 774  | 
|
| 69597 | 775  | 
\<^descr> \<^ML>\<open>Theory.add_deps\<close>~\<open>ctxt name c\<^sub>\<tau> \<^vec>d\<^sub>\<sigma>\<close> declares dependencies of  | 
| 61854 | 776  | 
a named specification for constant \<open>c\<^sub>\<tau>\<close>, relative to existing  | 
777  | 
specifications for constants \<open>\<^vec>d\<^sub>\<sigma>\<close>. This also works for type  | 
|
778  | 
constructors.  | 
|
| 70568 | 779  | 
|
780  | 
\<^descr> \<^ML>\<open>Thm_Deps.all_oracles\<close>~\<open>thms\<close> returns all oracles used in the  | 
|
781  | 
internal derivation of the given theorems; this covers the full graph of  | 
|
| 
71465
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
782  | 
transitive dependencies. The result contains an authentic oracle name; if  | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
783  | 
  @{ML Proofterm.proofs} was at least @{ML 1} during the oracle inference, it
 | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
784  | 
also contains the position of the oracle invocation and its proposition. See  | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70915 
diff
changeset
 | 
785  | 
  also the command @{command_ref "thm_oracles"}.
 | 
| 58618 | 786  | 
\<close>  | 
| 20521 | 787  | 
|
| 58618 | 788  | 
text %mlantiq \<open>  | 
| 39832 | 789  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 790  | 
  @{ML_antiquotation_def "ctyp"} & : & \<open>ML_antiquotation\<close> \\
 | 
791  | 
  @{ML_antiquotation_def "cterm"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
792  | 
  @{ML_antiquotation_def "cprop"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
793  | 
  @{ML_antiquotation_def "thm"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
794  | 
  @{ML_antiquotation_def "thms"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
795  | 
  @{ML_antiquotation_def "lemma"} & : & \<open>ML_antiquotation\<close> \\
 | 
|
| 70568 | 796  | 
  @{ML_antiquotation_def "oracle_name"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 39832 | 797  | 
  \end{matharray}
 | 
798  | 
||
| 69597 | 799  | 
\<^rail>\<open>  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
800  | 
  @@{ML_antiquotation ctyp} typ
 | 
| 39832 | 801  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
802  | 
  @@{ML_antiquotation cterm} term
 | 
| 39832 | 803  | 
;  | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
804  | 
  @@{ML_antiquotation cprop} prop
 | 
| 39832 | 805  | 
;  | 
| 62969 | 806  | 
  @@{ML_antiquotation thm} thm
 | 
| 
42510
 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 
wenzelm 
parents: 
42401 
diff
changeset
 | 
807  | 
;  | 
| 62969 | 808  | 
  @@{ML_antiquotation thms} thms
 | 
| 39832 | 809  | 
;  | 
| 
55029
 
61a6bf7d4b02
clarified @{rail} syntax: prefer explicit \<newline> symbol;
 
wenzelm 
parents: 
54883 
diff
changeset
 | 
810  | 
  @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \<newline>
 | 
| 74604 | 811  | 
    @{syntax for_fixes} @'by' method method?
 | 
| 70568 | 812  | 
;  | 
813  | 
  @@{ML_antiquotation oracle_name} embedded
 | 
|
| 69597 | 814  | 
\<close>  | 
| 39832 | 815  | 
|
| 61854 | 816  | 
  \<^descr> \<open>@{ctyp \<tau>}\<close> produces a certified type wrt.\ the current background theory
 | 
| 69597 | 817  | 
--- as abstract value of type \<^ML_type>\<open>ctyp\<close>.  | 
| 61854 | 818  | 
|
819  | 
  \<^descr> \<open>@{cterm t}\<close> and \<open>@{cprop \<phi>}\<close> produce a certified term wrt.\ the current
 | 
|
| 69597 | 820  | 
background theory --- as abstract value of type \<^ML_type>\<open>cterm\<close>.  | 
| 39832 | 821  | 
|
| 61854 | 822  | 
  \<^descr> \<open>@{thm a}\<close> produces a singleton fact --- as abstract value of type
 | 
| 69597 | 823  | 
\<^ML_type>\<open>thm\<close>.  | 
| 39832 | 824  | 
|
| 61854 | 825  | 
  \<^descr> \<open>@{thms a}\<close> produces a general fact --- as abstract value of type
 | 
| 69597 | 826  | 
\<^ML_type>\<open>thm list\<close>.  | 
| 39832 | 827  | 
|
| 61854 | 828  | 
  \<^descr> \<open>@{lemma \<phi> by meth}\<close> produces a fact that is proven on the spot according
 | 
829  | 
to the minimal proof, which imitates a terminal Isar proof. The result is an  | 
|
| 69597 | 830  | 
abstract value of type \<^ML_type>\<open>thm\<close> or \<^ML_type>\<open>thm list\<close>, depending on  | 
| 61854 | 831  | 
the number of propositions given here.  | 
| 39832 | 832  | 
|
| 61854 | 833  | 
The internal derivation object lacks a proper theorem name, but it is  | 
834  | 
formally closed, unless the \<open>(open)\<close> option is specified (this may impact  | 
|
835  | 
performance of applications with proof terms).  | 
|
| 39832 | 836  | 
|
| 61854 | 837  | 
Since ML antiquotations are always evaluated at compile-time, there is no  | 
838  | 
run-time overhead even for non-trivial proofs. Nonetheless, the  | 
|
839  | 
  justification is syntactically limited to a single @{command "by"} step.
 | 
|
840  | 
More complex Isar proofs should be done in regular theory source, before  | 
|
841  | 
compiling the corresponding ML text that uses the result.  | 
|
| 70568 | 842  | 
|
843  | 
  \<^descr> \<open>@{oracle_name a}\<close> inlines the internalized oracle name \<open>a\<close> --- as
 | 
|
844  | 
\<^ML_type>\<open>string\<close> literal.  | 
|
| 58618 | 845  | 
\<close>  | 
| 39832 | 846  | 
|
847  | 
||
| 58618 | 848  | 
subsection \<open>Auxiliary connectives \label{sec:logic-aux}\<close>
 | 
| 20521 | 849  | 
|
| 61854 | 850  | 
text \<open>  | 
851  | 
Theory \<open>Pure\<close> provides a few auxiliary connectives that are defined on top  | 
|
852  | 
  of the primitive ones, see \figref{fig:pure-aux}. These special constants
 | 
|
853  | 
are useful in certain internal encodings, and are normally not directly  | 
|
854  | 
exposed to the user.  | 
|
| 20501 | 855  | 
|
856  | 
  \begin{figure}[htb]
 | 
|
857  | 
  \begin{center}
 | 
|
| 20498 | 858  | 
  \begin{tabular}{ll}
 | 
| 61493 | 859  | 
\<open>conjunction :: prop \<Rightarrow> prop \<Rightarrow> prop\<close> & (infix \<open>&&&\<close>) \\  | 
860  | 
\<open>\<turnstile> A &&& B \<equiv> (\<And>C. (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C)\<close> \\[1ex]  | 
|
861  | 
\<open>prop :: prop \<Rightarrow> prop\<close> & (prefix \<open>#\<close>, suppressed) \\  | 
|
862  | 
\<open>#A \<equiv> A\<close> \\[1ex]  | 
|
863  | 
\<open>term :: \<alpha> \<Rightarrow> prop\<close> & (prefix \<open>TERM\<close>) \\  | 
|
864  | 
\<open>term x \<equiv> (\<And>A. A \<Longrightarrow> A)\<close> \\[1ex]  | 
|
865  | 
\<open>type :: \<alpha> itself\<close> & (prefix \<open>TYPE\<close>) \\  | 
|
866  | 
\<open>(unspecified)\<close> \\  | 
|
| 20498 | 867  | 
  \end{tabular}
 | 
| 20521 | 868  | 
  \caption{Definitions of auxiliary connectives}\label{fig:pure-aux}
 | 
| 20501 | 869  | 
  \end{center}
 | 
870  | 
  \end{figure}
 | 
|
871  | 
||
| 61854 | 872  | 
The introduction \<open>A \<Longrightarrow> B \<Longrightarrow> A &&& B\<close>, and eliminations (projections) \<open>A &&& B  | 
873  | 
\<Longrightarrow> A\<close> and \<open>A &&& B \<Longrightarrow> B\<close> are available as derived rules. Conjunction allows to  | 
|
874  | 
treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \<open>A  | 
|
875  | 
\<Longrightarrow> B \<Longrightarrow> C &&& D\<close>. In particular, the goal mechanism represents multiple claims  | 
|
876  | 
as explicit conjunction internally, but this is refined (via backwards  | 
|
877  | 
introduction) into separate sub-goals before the user commences the proof;  | 
|
878  | 
the final result is projected into a list of theorems using eliminations  | 
|
879  | 
  (cf.\ \secref{sec:tactical-goals}).
 | 
|
| 20498 | 880  | 
|
| 61854 | 881  | 
The \<open>prop\<close> marker (\<open>#\<close>) makes arbitrarily complex propositions appear as  | 
882  | 
atomic, without changing the meaning: \<open>\<Gamma> \<turnstile> A\<close> and \<open>\<Gamma> \<turnstile> #A\<close> are  | 
|
883  | 
  interchangeable. See \secref{sec:tactical-goals} for specific operations.
 | 
|
| 20521 | 884  | 
|
| 61854 | 885  | 
The \<open>term\<close> marker turns any well-typed term into a derivable proposition: \<open>\<turnstile>  | 
886  | 
TERM t\<close> holds unconditionally. Although this is logically vacuous, it allows  | 
|
887  | 
to treat terms and proofs uniformly, similar to a type-theoretic framework.  | 
|
| 20498 | 888  | 
|
| 61854 | 889  | 
The \<open>TYPE\<close> constructor is the canonical representative of the unspecified  | 
890  | 
type \<open>\<alpha> itself\<close>; it essentially injects the language of types into that of  | 
|
891  | 
terms. There is specific notation \<open>TYPE(\<tau>)\<close> for \<open>TYPE\<^bsub>\<tau> itself\<^esub>\<close>. Although  | 
|
892  | 
being devoid of any particular meaning, the term \<open>TYPE(\<tau>)\<close> accounts for the  | 
|
893  | 
type \<open>\<tau>\<close> within the term language. In particular, \<open>TYPE(\<alpha>)\<close> may be used as  | 
|
894  | 
formal argument in primitive definitions, in order to circumvent hidden  | 
|
895  | 
  polymorphism (cf.\ \secref{sec:terms}). For example, \<open>c TYPE(\<alpha>) \<equiv> A[\<alpha>]\<close>
 | 
|
896  | 
defines \<open>c :: \<alpha> itself \<Rightarrow> prop\<close> in terms of a proposition \<open>A\<close> that depends on  | 
|
897  | 
an additional type argument, which is essentially a predicate on types.  | 
|
| 58618 | 898  | 
\<close>  | 
| 20501 | 899  | 
|
| 58618 | 900  | 
text %mlref \<open>  | 
| 20521 | 901  | 
  \begin{mldecls}
 | 
| 73764 | 902  | 
  @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\
 | 
903  | 
  @{define_ML Conjunction.elim: "thm -> thm * thm"} \\
 | 
|
904  | 
  @{define_ML Drule.mk_term: "cterm -> thm"} \\
 | 
|
905  | 
  @{define_ML Drule.dest_term: "thm -> cterm"} \\
 | 
|
906  | 
  @{define_ML Logic.mk_type: "typ -> term"} \\
 | 
|
907  | 
  @{define_ML Logic.dest_type: "term -> typ"} \\
 | 
|
| 20521 | 908  | 
  \end{mldecls}
 | 
909  | 
||
| 69597 | 910  | 
\<^descr> \<^ML>\<open>Conjunction.intr\<close> derives \<open>A &&& B\<close> from \<open>A\<close> and \<open>B\<close>.  | 
| 20542 | 911  | 
|
| 69597 | 912  | 
\<^descr> \<^ML>\<open>Conjunction.elim\<close> derives \<open>A\<close> and \<open>B\<close> from \<open>A &&& B\<close>.  | 
| 20542 | 913  | 
|
| 69597 | 914  | 
\<^descr> \<^ML>\<open>Drule.mk_term\<close> derives \<open>TERM t\<close>.  | 
| 20542 | 915  | 
|
| 69597 | 916  | 
\<^descr> \<^ML>\<open>Drule.dest_term\<close> recovers term \<open>t\<close> from \<open>TERM t\<close>.  | 
| 20542 | 917  | 
|
| 69597 | 918  | 
\<^descr> \<^ML>\<open>Logic.mk_type\<close>~\<open>\<tau>\<close> produces the term \<open>TYPE(\<tau>)\<close>.  | 
| 20542 | 919  | 
|
| 69597 | 920  | 
\<^descr> \<^ML>\<open>Logic.dest_type\<close>~\<open>TYPE(\<tau>)\<close> recovers the type \<open>\<tau>\<close>.  | 
| 58618 | 921  | 
\<close>  | 
| 18537 | 922  | 
|
| 20480 | 923  | 
|
| 58618 | 924  | 
subsection \<open>Sort hypotheses\<close>  | 
| 52406 | 925  | 
|
| 61854 | 926  | 
text \<open>  | 
927  | 
  Type variables are decorated with sorts, as explained in \secref{sec:types}.
 | 
|
928  | 
This constrains type instantiation to certain ranges of types: variable  | 
|
929  | 
\<open>\<alpha>\<^sub>s\<close> may only be assigned to types \<open>\<tau>\<close> that belong to sort \<open>s\<close>. Within the  | 
|
930  | 
logic, sort constraints act like implicit preconditions on the result \<open>\<lparr>\<alpha>\<^sub>1  | 
|
931  | 
: s\<^sub>1\<rparr>, \<dots>, \<lparr>\<alpha>\<^sub>n : s\<^sub>n\<rparr>, \<Gamma> \<turnstile> \<phi>\<close> where the type variables \<open>\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n\<close> cover  | 
|
932  | 
the propositions \<open>\<Gamma>\<close>, \<open>\<phi>\<close>, as well as the proof of \<open>\<Gamma> \<turnstile> \<phi>\<close>.  | 
|
| 52406 | 933  | 
|
| 61854 | 934  | 
These \<^emph>\<open>sort hypothesis\<close> of a theorem are passed monotonically through  | 
935  | 
further derivations. They are redundant, as long as the statement of a  | 
|
936  | 
theorem still contains the type variables that are accounted here. The  | 
|
937  | 
logical significance of sort hypotheses is limited to the boundary case  | 
|
938  | 
where type variables disappear from the proposition, e.g.\ \<open>\<lparr>\<alpha>\<^sub>s : s\<rparr> \<turnstile> \<phi>\<close>.  | 
|
939  | 
Since such dangling type variables can be renamed arbitrarily without  | 
|
940  | 
changing the proposition \<open>\<phi>\<close>, the inference kernel maintains sort hypotheses  | 
|
941  | 
in anonymous form \<open>s \<turnstile> \<phi>\<close>.  | 
|
| 52406 | 942  | 
|
| 61854 | 943  | 
In most practical situations, such extra sort hypotheses may be stripped in  | 
944  | 
a final bookkeeping step, e.g.\ at the end of a proof: they are typically  | 
|
945  | 
left over from intermediate reasoning with type classes that can be  | 
|
946  | 
satisfied by some concrete type \<open>\<tau>\<close> of sort \<open>s\<close> to replace the hypothetical  | 
|
947  | 
type variable \<open>\<alpha>\<^sub>s\<close>.  | 
|
948  | 
\<close>  | 
|
| 52406 | 949  | 
|
| 58618 | 950  | 
text %mlref \<open>  | 
| 52406 | 951  | 
  \begin{mldecls}
 | 
| 77869 | 952  | 
  @{define_ML Thm.extra_shyps: "thm -> sort list"} \\
 | 
| 73764 | 953  | 
  @{define_ML Thm.strip_shyps: "thm -> thm"} \\
 | 
| 52406 | 954  | 
  \end{mldecls}
 | 
955  | 
||
| 69597 | 956  | 
\<^descr> \<^ML>\<open>Thm.extra_shyps\<close>~\<open>thm\<close> determines the extraneous sort hypotheses of  | 
| 61854 | 957  | 
the given theorem, i.e.\ the sorts that are not present within type  | 
958  | 
variables of the statement.  | 
|
| 52406 | 959  | 
|
| 69597 | 960  | 
\<^descr> \<^ML>\<open>Thm.strip_shyps\<close>~\<open>thm\<close> removes any extraneous sort hypotheses that  | 
| 61854 | 961  | 
can be witnessed from the type signature.  | 
| 58618 | 962  | 
\<close>  | 
| 52406 | 963  | 
|
| 61854 | 964  | 
text %mlex \<open>  | 
| 69597 | 965  | 
The following artificial example demonstrates the derivation of \<^prop>\<open>False\<close> with a pending sort hypothesis involving a logically empty sort.  | 
| 61854 | 966  | 
\<close>  | 
| 52406 | 967  | 
|
968  | 
class empty =  | 
|
969  | 
assumes bad: "\<And>(x::'a) y. x \<noteq> y"  | 
|
970  | 
||
971  | 
theorem (in empty) false: False  | 
|
972  | 
using bad by blast  | 
|
973  | 
||
| 77869 | 974  | 
ML_val \<open>\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\<open>empty\<close>])\<close>
 | 
| 52406 | 975  | 
|
| 61854 | 976  | 
text \<open>  | 
977  | 
Thanks to the inference kernel managing sort hypothesis according to their  | 
|
978  | 
logical significance, this example is merely an instance of \<^emph>\<open>ex falso  | 
|
979  | 
quodlibet consequitur\<close> --- not a collapse of the logical framework!  | 
|
980  | 
\<close>  | 
|
| 52406 | 981  | 
|
982  | 
||
| 58618 | 983  | 
section \<open>Object-level rules \label{sec:obj-rules}\<close>
 | 
| 18537 | 984  | 
|
| 58618 | 985  | 
text \<open>  | 
| 61854 | 986  | 
The primitive inferences covered so far mostly serve foundational purposes.  | 
987  | 
User-level reasoning usually works via object-level rules that are  | 
|
988  | 
represented as theorems of Pure. Composition of rules involves  | 
|
989  | 
\<^emph>\<open>backchaining\<close>, \<^emph>\<open>higher-order unification\<close> modulo \<open>\<alpha>\<beta>\<eta>\<close>-conversion of  | 
|
990  | 
\<open>\<lambda>\<close>-terms, and so-called \<^emph>\<open>lifting\<close> of rules into a context of \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close>  | 
|
991  | 
connectives. Thus the full power of higher-order Natural Deduction in  | 
|
992  | 
Isabelle/Pure becomes readily available.  | 
|
| 58618 | 993  | 
\<close>  | 
| 20491 | 994  | 
|
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
995  | 
|
| 58618 | 996  | 
subsection \<open>Hereditary Harrop Formulae\<close>  | 
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
997  | 
|
| 58618 | 998  | 
text \<open>  | 
| 61854 | 999  | 
The idea of object-level rules is to model Natural Deduction inferences in  | 
| 76987 | 1000  | 
the style of Gentzen \<^cite>\<open>"Gentzen:1935"\<close>, but we allow arbitrary nesting  | 
1001  | 
similar to \<^cite>\<open>"Schroeder-Heister:1984"\<close>. The most basic rule format is  | 
|
| 69307 | 1002  | 
that of a \<^emph>\<open>Horn Clause\<close>:  | 
| 29768 | 1003  | 
\[  | 
| 61493 | 1004  | 
  \infer{\<open>A\<close>}{\<open>A\<^sub>1\<close> & \<open>\<dots>\<close> & \<open>A\<^sub>n\<close>}
 | 
| 29768 | 1005  | 
\]  | 
| 61854 | 1006  | 
where \<open>A, A\<^sub>1, \<dots>, A\<^sub>n\<close> are atomic propositions of the framework, usually of  | 
1007  | 
the form \<open>Trueprop B\<close>, where \<open>B\<close> is a (compound) object-level statement.  | 
|
1008  | 
This object-level inference corresponds to an iterated implication in Pure  | 
|
1009  | 
like this:  | 
|
| 29768 | 1010  | 
\[  | 
| 61493 | 1011  | 
\<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A\<close>  | 
| 29768 | 1012  | 
\]  | 
| 61854 | 1013  | 
As an example consider conjunction introduction: \<open>A \<Longrightarrow> B \<Longrightarrow> A \<and> B\<close>. Any  | 
1014  | 
parameters occurring in such rule statements are conceptionally treated as  | 
|
1015  | 
arbitrary:  | 
|
| 29768 | 1016  | 
\[  | 
| 61493 | 1017  | 
\<open>\<And>x\<^sub>1 \<dots> x\<^sub>m. A\<^sub>1 x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> \<dots> A\<^sub>n x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> A x\<^sub>1 \<dots> x\<^sub>m\<close>  | 
| 29768 | 1018  | 
\]  | 
| 20491 | 1019  | 
|
| 61854 | 1020  | 
Nesting of rules means that the positions of \<open>A\<^sub>i\<close> may again hold compound  | 
1021  | 
rules, not just atomic propositions. Propositions of this format are called  | 
|
| 76987 | 1022  | 
\<^emph>\<open>Hereditary Harrop Formulae\<close> in the literature \<^cite>\<open>"Miller:1991"\<close>. Here  | 
| 61854 | 1023  | 
we give an inductive characterization as follows:  | 
| 29768 | 1024  | 
|
| 61416 | 1025  | 
\<^medskip>  | 
| 29768 | 1026  | 
  \begin{tabular}{ll}
 | 
| 61493 | 1027  | 
\<open>\<^bold>x\<close> & set of variables \\  | 
1028  | 
\<open>\<^bold>A\<close> & set of atomic propositions \\  | 
|
1029  | 
\<open>\<^bold>H = \<And>\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \<Longrightarrow> \<^bold>A\<close> & set of Hereditary Harrop Formulas \\  | 
|
| 29768 | 1030  | 
  \end{tabular}
 | 
| 61416 | 1031  | 
\<^medskip>  | 
| 29768 | 1032  | 
|
| 61854 | 1033  | 
Thus we essentially impose nesting levels on propositions formed from \<open>\<And>\<close>  | 
1034  | 
and \<open>\<Longrightarrow>\<close>. At each level there is a prefix of parameters and compound  | 
|
1035  | 
premises, concluding an atomic proposition. Typical examples are  | 
|
1036  | 
\<open>\<longrightarrow>\<close>-introduction \<open>(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B\<close> or mathematical induction \<open>P 0 \<Longrightarrow> (\<And>n. P n  | 
|
1037  | 
\<Longrightarrow> P (Suc n)) \<Longrightarrow> P n\<close>. Even deeper nesting occurs in well-founded induction  | 
|
1038  | 
\<open>(\<And>x. (\<And>y. y \<prec> x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x\<close>, but this already marks the limit of  | 
|
1039  | 
rule complexity that is usually seen in practice.  | 
|
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1040  | 
|
| 61416 | 1041  | 
\<^medskip>  | 
| 61854 | 1042  | 
Regular user-level inferences in Isabelle/Pure always maintain the following  | 
1043  | 
canonical form of results:  | 
|
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1044  | 
|
| 61854 | 1045  | 
\<^item> Normalization by \<open>(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)\<close>, which is a theorem of  | 
1046  | 
Pure, means that quantifiers are pushed in front of implication at each  | 
|
1047  | 
level of nesting. The normal form is a Hereditary Harrop Formula.  | 
|
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1048  | 
|
| 61854 | 1049  | 
\<^item> The outermost prefix of parameters is represented via schematic variables:  | 
1050  | 
instead of \<open>\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x\<close> we have \<open>\<^vec>H  | 
|
1051  | 
?\<^vec>x \<Longrightarrow> A ?\<^vec>x\<close>. Note that this representation looses information  | 
|
1052  | 
about the order of parameters, and vacuous quantifiers vanish automatically.  | 
|
| 58618 | 1053  | 
\<close>  | 
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1054  | 
|
| 58618 | 1055  | 
text %mlref \<open>  | 
| 29771 | 1056  | 
  \begin{mldecls}
 | 
| 73764 | 1057  | 
  @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\
 | 
| 29771 | 1058  | 
  \end{mldecls}
 | 
1059  | 
||
| 69597 | 1060  | 
\<^descr> \<^ML>\<open>Simplifier.norm_hhf\<close>~\<open>ctxt thm\<close> normalizes the given theorem  | 
| 61854 | 1061  | 
according to the canonical form specified above. This is occasionally  | 
1062  | 
helpful to repair some low-level tools that do not handle Hereditary Harrop  | 
|
1063  | 
Formulae properly.  | 
|
| 58618 | 1064  | 
\<close>  | 
| 29771 | 1065  | 
|
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1066  | 
|
| 58618 | 1067  | 
subsection \<open>Rule composition\<close>  | 
| 
29769
 
03634a9e91ae
improved section on "Hereditary Harrop Formulae";
 
wenzelm 
parents: 
29768 
diff
changeset
 | 
1068  | 
|
| 58618 | 1069  | 
text \<open>  | 
| 61854 | 1070  | 
  The rule calculus of Isabelle/Pure provides two main inferences: @{inference
 | 
1071  | 
  resolution} (i.e.\ back-chaining of rules) and @{inference assumption}
 | 
|
1072  | 
(i.e.\ closing a branch), both modulo higher-order unification. There are  | 
|
1073  | 
  also combined variants, notably @{inference elim_resolution} and @{inference
 | 
|
1074  | 
dest_resolution}.  | 
|
| 20491 | 1075  | 
|
| 61854 | 1076  | 
  To understand the all-important @{inference resolution} principle, we first
 | 
1077  | 
  consider raw @{inference_def composition} (modulo higher-order unification
 | 
|
1078  | 
with substitution \<open>\<vartheta>\<close>):  | 
|
| 20498 | 1079  | 
\[  | 
| 61493 | 1080  | 
  \infer[(@{inference_def composition})]{\<open>\<^vec>A\<vartheta> \<Longrightarrow> C\<vartheta>\<close>}
 | 
1081  | 
  {\<open>\<^vec>A \<Longrightarrow> B\<close> & \<open>B' \<Longrightarrow> C\<close> & \<open>B\<vartheta> = B'\<vartheta>\<close>}
 | 
|
| 20498 | 1082  | 
\]  | 
| 61854 | 1083  | 
Here the conclusion of the first rule is unified with the premise of the  | 
1084  | 
second; the resulting rule instance inherits the premises of the first and  | 
|
1085  | 
conclusion of the second. Note that \<open>C\<close> can again consist of iterated  | 
|
1086  | 
implications. We can also permute the premises of the second rule  | 
|
1087  | 
back-and-forth in order to compose with \<open>B'\<close> in any position (subsequently  | 
|
1088  | 
we shall always refer to position 1 w.l.o.g.).  | 
|
| 20498 | 1089  | 
|
| 61854 | 1090  | 
  In @{inference composition} the internal structure of the common part \<open>B\<close>
 | 
1091  | 
  and \<open>B'\<close> is not taken into account. For proper @{inference resolution} we
 | 
|
1092  | 
require \<open>B\<close> to be atomic, and explicitly observe the structure \<open>\<And>\<^vec>x.  | 
|
1093  | 
\<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x\<close> of the premise of the second rule. The idea  | 
|
1094  | 
is to adapt the first rule by ``lifting'' it into this context, by means of  | 
|
1095  | 
iterated application of the following inferences:  | 
|
| 20498 | 1096  | 
\[  | 
| 61493 | 1097  | 
  \infer[(@{inference_def imp_lift})]{\<open>(\<^vec>H \<Longrightarrow> \<^vec>A) \<Longrightarrow> (\<^vec>H \<Longrightarrow> B)\<close>}{\<open>\<^vec>A \<Longrightarrow> B\<close>}
 | 
| 20498 | 1098  | 
\]  | 
1099  | 
\[  | 
|
| 61493 | 1100  | 
  \infer[(@{inference_def all_lift})]{\<open>(\<And>\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \<Longrightarrow> (\<And>\<^vec>x. B (?\<^vec>a \<^vec>x))\<close>}{\<open>\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a\<close>}
 | 
| 20498 | 1101  | 
\]  | 
| 29771 | 1102  | 
  By combining raw composition with lifting, we get full @{inference
 | 
1103  | 
resolution} as follows:  | 
|
| 20498 | 1104  | 
\[  | 
| 29771 | 1105  | 
  \infer[(@{inference_def resolution})]
 | 
| 61493 | 1106  | 
  {\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (?\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>\<close>}
 | 
| 20498 | 1107  | 
  {\begin{tabular}{l}
 | 
| 61493 | 1108  | 
\<open>\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a\<close> \\  | 
1109  | 
\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C\<close> \\  | 
|
1110  | 
\<open>(\<lambda>\<^vec>x. B (?\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>\<close> \\  | 
|
| 20498 | 1111  | 
   \end{tabular}}
 | 
1112  | 
\]  | 
|
1113  | 
||
| 61854 | 1114  | 
Continued resolution of rules allows to back-chain a problem towards more  | 
1115  | 
and sub-problems. Branches are closed either by resolving with a rule of 0  | 
|
1116  | 
premises, or by producing a ``short-circuit'' within a solved situation  | 
|
1117  | 
(again modulo unification):  | 
|
| 29771 | 1118  | 
\[  | 
| 61493 | 1119  | 
  \infer[(@{inference_def assumption})]{\<open>C\<vartheta>\<close>}
 | 
| 
61962
 
9c8fc56032e3
eliminated obscure macro that is in conflict with amsmath.sty;
 
wenzelm 
parents: 
61854 
diff
changeset
 | 
1120  | 
  {\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C\<close> & \<open>A\<vartheta> = H\<^sub>i\<vartheta>\<close>~~\mbox{(for some~\<open>i\<close>)}}
 | 
| 29771 | 1121  | 
\]  | 
| 20498 | 1122  | 
|
| 52422 | 1123  | 
  %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution}
 | 
| 58618 | 1124  | 
\<close>  | 
| 18537 | 1125  | 
|
| 58618 | 1126  | 
text %mlref \<open>  | 
| 29768 | 1127  | 
  \begin{mldecls}
 | 
| 73764 | 1128  | 
  @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\
 | 
1129  | 
  @{define_ML_infix "RS": "thm * thm -> thm"} \\
 | 
|
| 46256 | 1130  | 
|
| 73764 | 1131  | 
  @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\
 | 
1132  | 
  @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\
 | 
|
| 46256 | 1133  | 
|
| 73764 | 1134  | 
  @{define_ML_infix "MRS": "thm list * thm -> thm"} \\
 | 
1135  | 
  @{define_ML_infix "OF": "thm * thm list -> thm"} \\
 | 
|
| 29768 | 1136  | 
  \end{mldecls}
 | 
1137  | 
||
| 61854 | 1138  | 
\<^descr> \<open>rule\<^sub>1 RSN (i, rule\<^sub>2)\<close> resolves the conclusion of \<open>rule\<^sub>1\<close> with the  | 
1139  | 
  \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>, according to the @{inference resolution}
 | 
|
1140  | 
principle explained above. Unless there is precisely one resolvent it raises  | 
|
| 69597 | 1141  | 
exception \<^ML>\<open>THM\<close>.  | 
| 46256 | 1142  | 
|
| 61854 | 1143  | 
  This corresponds to the rule attribute @{attribute THEN} in Isar source
 | 
1144  | 
language.  | 
|
| 46256 | 1145  | 
|
| 61854 | 1146  | 
\<^descr> \<open>rule\<^sub>1 RS rule\<^sub>2\<close> abbreviates \<open>rule\<^sub>1 RSN (1, rule\<^sub>2)\<close>.  | 
| 29768 | 1147  | 
|
| 61854 | 1148  | 
\<^descr> \<open>rules\<^sub>1 RLN (i, rules\<^sub>2)\<close> joins lists of rules. For every \<open>rule\<^sub>1\<close> in  | 
1149  | 
\<open>rules\<^sub>1\<close> and \<open>rule\<^sub>2\<close> in \<open>rules\<^sub>2\<close>, it resolves the conclusion of \<open>rule\<^sub>1\<close>  | 
|
1150  | 
with the \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>, accumulating multiple results in one  | 
|
1151  | 
big list. Note that such strict enumerations of higher-order unifications  | 
|
1152  | 
can be inefficient compared to the lazy variant seen in elementary tactics  | 
|
| 69597 | 1153  | 
like \<^ML>\<open>resolve_tac\<close>.  | 
| 46256 | 1154  | 
|
| 61854 | 1155  | 
\<^descr> \<open>rules\<^sub>1 RL rules\<^sub>2\<close> abbreviates \<open>rules\<^sub>1 RLN (1, rules\<^sub>2)\<close>.  | 
| 46256 | 1156  | 
|
| 61854 | 1157  | 
\<^descr> \<open>[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule\<close> resolves \<open>rule\<^sub>i\<close> against premise \<open>i\<close> of  | 
1158  | 
\<open>rule\<close>, for \<open>i = n, \<dots>, 1\<close>. By working from right to left, newly emerging  | 
|
1159  | 
premises are concatenated in the result, without interfering.  | 
|
| 46256 | 1160  | 
|
| 61854 | 1161  | 
\<^descr> \<open>rule OF rules\<close> is an alternative notation for \<open>rules MRS rule\<close>, which  | 
1162  | 
makes rule composition look more like function application. Note that the  | 
|
1163  | 
argument \<open>rules\<close> need not be atomic.  | 
|
| 46256 | 1164  | 
|
| 61854 | 1165  | 
  This corresponds to the rule attribute @{attribute OF} in Isar source
 | 
1166  | 
language.  | 
|
| 58618 | 1167  | 
\<close>  | 
| 30272 | 1168  | 
|
| 52407 | 1169  | 
|
| 58618 | 1170  | 
section \<open>Proof terms \label{sec:proof-terms}\<close>
 | 
| 52407 | 1171  | 
|
| 61854 | 1172  | 
text \<open>  | 
1173  | 
The Isabelle/Pure inference kernel can record the proof of each theorem as a  | 
|
1174  | 
proof term that contains all logical inferences in detail. Rule composition  | 
|
1175  | 
  by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken
 | 
|
1176  | 
down to primitive rules of the logical framework. The proof term can be  | 
|
1177  | 
inspected by a separate proof-checker, for example.  | 
|
| 52407 | 1178  | 
|
| 61854 | 1179  | 
According to the well-known \<^emph>\<open>Curry-Howard isomorphism\<close>, a proof can be  | 
1180  | 
viewed as a \<open>\<lambda>\<close>-term. Following this idea, proofs in Isabelle are internally  | 
|
1181  | 
represented by a datatype similar to the one for terms described in  | 
|
1182  | 
  \secref{sec:terms}. On top of these syntactic terms, two more layers of
 | 
|
1183  | 
\<open>\<lambda>\<close>-calculus are added, which correspond to \<open>\<And>x :: \<alpha>. B x\<close> and \<open>A \<Longrightarrow> B\<close>  | 
|
1184  | 
according to the propositions-as-types principle. The resulting 3-level  | 
|
1185  | 
\<open>\<lambda>\<close>-calculus resembles ``\<open>\<lambda>HOL\<close>'' in the more abstract setting of Pure Type  | 
|
| 76987 | 1186  | 
Systems (PTS) \<^cite>\<open>"Barendregt-Geuvers:2001"\<close>, if some fine points like  | 
| 61854 | 1187  | 
schematic polymorphism and type classes are ignored.  | 
| 52407 | 1188  | 
|
| 61416 | 1189  | 
\<^medskip>  | 
| 61854 | 1190  | 
\<^emph>\<open>Proof abstractions\<close> of the form \<open>\<^bold>\<lambda>x :: \<alpha>. prf\<close> or \<open>\<^bold>\<lambda>p : A. prf\<close>  | 
1191  | 
correspond to introduction of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>, and \<^emph>\<open>proof applications\<close> of the form  | 
|
1192  | 
\<open>p \<cdot> t\<close> or \<open>p \<bullet> q\<close> correspond to elimination of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>. Actual types \<open>\<alpha>\<close>,  | 
|
1193  | 
propositions \<open>A\<close>, and terms \<open>t\<close> might be suppressed and reconstructed from  | 
|
1194  | 
the overall proof term.  | 
|
| 52407 | 1195  | 
|
| 61416 | 1196  | 
\<^medskip>  | 
| 61854 | 1197  | 
Various atomic proofs indicate special situations within the proof  | 
1198  | 
construction as follows.  | 
|
| 52407 | 1199  | 
|
| 61854 | 1200  | 
A \<^emph>\<open>bound proof variable\<close> is a natural number \<open>b\<close> that acts as de-Bruijn  | 
1201  | 
index for proof term abstractions.  | 
|
| 52407 | 1202  | 
|
| 61854 | 1203  | 
A \<^emph>\<open>minimal proof\<close> ``\<open>?\<close>'' is a dummy proof term. This indicates some  | 
1204  | 
unrecorded part of the proof.  | 
|
| 52407 | 1205  | 
|
| 61854 | 1206  | 
\<open>Hyp A\<close> refers to some pending hypothesis by giving its proposition. This  | 
1207  | 
indicates an open context of implicit hypotheses, similar to loose bound  | 
|
1208  | 
  variables or free variables within a term (\secref{sec:terms}).
 | 
|
| 52407 | 1209  | 
|
| 61854 | 1210  | 
An \<^emph>\<open>axiom\<close> or \<^emph>\<open>oracle\<close> \<open>a : A[\<^vec>\<tau>]\<close> refers some postulated \<open>proof  | 
1211  | 
constant\<close>, which is subject to schematic polymorphism of theory content, and  | 
|
1212  | 
the particular type instantiation may be given explicitly. The vector of  | 
|
1213  | 
types \<open>\<^vec>\<tau>\<close> refers to the schematic type variables in the generic  | 
|
| 61493 | 1214  | 
proposition \<open>A\<close> in canonical order.  | 
| 52407 | 1215  | 
|
| 61854 | 1216  | 
A \<^emph>\<open>proof promise\<close> \<open>a : A[\<^vec>\<tau>]\<close> is a placeholder for some proof of  | 
1217  | 
polymorphic proposition \<open>A\<close>, with explicit type instantiation as given by  | 
|
1218  | 
the vector \<open>\<^vec>\<tau>\<close>, as above. Unlike axioms or oracles, proof promises  | 
|
1219  | 
may be \<^emph>\<open>fulfilled\<close> eventually, by substituting \<open>a\<close> by some particular proof  | 
|
1220  | 
\<open>q\<close> at the corresponding type instance. This acts like Hindley-Milner  | 
|
1221  | 
\<open>let\<close>-polymorphism: a generic local proof definition may get used at  | 
|
1222  | 
different type instances, and is replaced by the concrete instance  | 
|
1223  | 
eventually.  | 
|
| 52407 | 1224  | 
|
| 61854 | 1225  | 
A \<^emph>\<open>named theorem\<close> wraps up some concrete proof as a closed formal entity,  | 
1226  | 
in the manner of constant definitions for proof terms. The \<^emph>\<open>proof body\<close> of  | 
|
1227  | 
such boxed theorems involves some digest about oracles and promises  | 
|
1228  | 
occurring in the original proof. This allows the inference kernel to manage  | 
|
1229  | 
this critical information without the full overhead of explicit proof terms.  | 
|
| 58618 | 1230  | 
\<close>  | 
| 52407 | 1231  | 
|
| 52411 | 1232  | 
|
| 58618 | 1233  | 
subsection \<open>Reconstructing and checking proof terms\<close>  | 
| 52411 | 1234  | 
|
| 61854 | 1235  | 
text \<open>  | 
1236  | 
Fully explicit proof terms can be large, but most of this information is  | 
|
1237  | 
redundant and can be reconstructed from the context. Therefore, the  | 
|
1238  | 
Isabelle/Pure inference kernel records only \<^emph>\<open>implicit\<close> proof terms, by  | 
|
1239  | 
omitting all typing information in terms, all term and type labels of proof  | 
|
1240  | 
abstractions, and some argument terms of applications \<open>p \<cdot> t\<close> (if possible).  | 
|
| 52411 | 1241  | 
|
| 61854 | 1242  | 
There are separate operations to reconstruct the full proof term later on,  | 
| 76987 | 1243  | 
using \<^emph>\<open>higher-order pattern unification\<close> \<^cite>\<open>"nipkow-patterns" and  | 
1244  | 
"Berghofer-Nipkow:2000:TPHOL"\<close>.  | 
|
| 52411 | 1245  | 
|
| 61854 | 1246  | 
The \<^emph>\<open>proof checker\<close> expects a fully reconstructed proof term, and can turn  | 
1247  | 
it into a theorem by replaying its primitive inferences within the kernel.  | 
|
1248  | 
\<close>  | 
|
| 52411 | 1249  | 
|
| 52412 | 1250  | 
|
| 58618 | 1251  | 
subsection \<open>Concrete syntax of proof terms\<close>  | 
| 52412 | 1252  | 
|
| 61854 | 1253  | 
text \<open>  | 
1254  | 
The concrete syntax of proof terms is a slight extension of the regular  | 
|
| 76987 | 1255  | 
inner syntax of Isabelle/Pure \<^cite>\<open>"isabelle-isar-ref"\<close>. Its main  | 
| 61854 | 1256  | 
  syntactic category @{syntax (inner) proof} is defined as follows:
 | 
| 52412 | 1257  | 
|
1258  | 
  \begin{center}
 | 
|
1259  | 
  \begin{supertabular}{rclr}
 | 
|
1260  | 
||
| 70389 | 1261  | 
  @{syntax_def (inner) proof} & = & \<open>\<^bold>\<lambda>\<close> \<open>params\<close> \<^verbatim>\<open>.\<close> \<open>proof\<close> \\
 | 
| 61493 | 1262  | 
& \<open>|\<close> & \<open>proof\<close> \<open>\<cdot>\<close> \<open>any\<close> \\  | 
1263  | 
& \<open>|\<close> & \<open>proof\<close> \<open>\<bullet>\<close> \<open>proof\<close> \\  | 
|
1264  | 
& \<open>|\<close> & \<open>id | longid\<close> \\  | 
|
| 52412 | 1265  | 
\\  | 
1266  | 
||
| 61493 | 1267  | 
\<open>param\<close> & = & \<open>idt\<close> \\  | 
| 61503 | 1268  | 
& \<open>|\<close> & \<open>idt\<close> \<^verbatim>\<open>:\<close> \<open>prop\<close> \\  | 
1269  | 
& \<open>|\<close> & \<^verbatim>\<open>(\<close> \<open>param\<close> \<^verbatim>\<open>)\<close> \\  | 
|
| 52412 | 1270  | 
\\  | 
1271  | 
||
| 61493 | 1272  | 
\<open>params\<close> & = & \<open>param\<close> \\  | 
1273  | 
& \<open>|\<close> & \<open>param\<close> \<open>params\<close> \\  | 
|
| 52412 | 1274  | 
|
1275  | 
  \end{supertabular}
 | 
|
1276  | 
  \end{center}
 | 
|
1277  | 
||
| 61854 | 1278  | 
Implicit term arguments in partial proofs are indicated by ``\<open>_\<close>''. Type  | 
1279  | 
arguments for theorems and axioms may be specified using \<open>p \<cdot> TYPE(type)\<close>  | 
|
1280  | 
(they must appear before any other term argument of a theorem or axiom, but  | 
|
1281  | 
may be omitted altogether).  | 
|
| 52412 | 1282  | 
|
| 61416 | 1283  | 
\<^medskip>  | 
| 61854 | 1284  | 
There are separate read and print operations for proof terms, in order to  | 
1285  | 
avoid conflicts with the regular term language.  | 
|
| 58618 | 1286  | 
\<close>  | 
| 52412 | 1287  | 
|
| 58618 | 1288  | 
text %mlref \<open>  | 
| 52408 | 1289  | 
  \begin{mldecls}
 | 
| 73764 | 1290  | 
  @{define_ML_type proof} \\
 | 
1291  | 
  @{define_ML_type proof_body} \\
 | 
|
1292  | 
  @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\
 | 
|
1293  | 
  @{define_ML Proofterm.reconstruct_proof:
 | 
|
| 70449 | 1294  | 
"theory -> term -> proof -> proof"} \\  | 
| 73764 | 1295  | 
  @{define_ML Proofterm.expand_proof: "theory ->
 | 
| 80590 | 1296  | 
(Proofterm.thm_header -> Thm_Name.P option) -> proof -> proof"} \\  | 
| 73764 | 1297  | 
  @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\
 | 
1298  | 
  @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\
 | 
|
1299  | 
  @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\
 | 
|
| 52408 | 1300  | 
  \end{mldecls}
 | 
1301  | 
||
| 69597 | 1302  | 
\<^descr> Type \<^ML_type>\<open>proof\<close> represents proof terms; this is a datatype with  | 
| 73764 | 1303  | 
  constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"},
 | 
1304  | 
  @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML
 | 
|
1305  | 
  Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained
 | 
|
| 
71777
 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 
wenzelm 
parents: 
71465 
diff
changeset
 | 
1306  | 
above. %FIXME PClass (!?)  | 
| 61854 | 1307  | 
|
| 69597 | 1308  | 
\<^descr> Type \<^ML_type>\<open>proof_body\<close> represents the nested proof information of a  | 
| 61854 | 1309  | 
named theorem, consisting of a digest of oracles and named theorem over some  | 
1310  | 
proof term. The digest only covers the directly visible part of the proof:  | 
|
1311  | 
in order to get the full information, the implicit graph of nested theorems  | 
|
| 69597 | 1312  | 
needs to be traversed (e.g.\ using \<^ML>\<open>Proofterm.fold_body_thms\<close>).  | 
| 52408 | 1313  | 
|
| 69597 | 1314  | 
\<^descr> \<^ML>\<open>Thm.proof_of\<close>~\<open>thm\<close> and \<^ML>\<open>Thm.proof_body_of\<close>~\<open>thm\<close> produce the  | 
| 61854 | 1315  | 
proof term or proof body (with digest of oracles and theorems) from a given  | 
1316  | 
theorem. Note that this involves a full join of internal futures that  | 
|
1317  | 
fulfill pending proof promises, and thus disrupts the natural bottom-up  | 
|
1318  | 
construction of proofs by introducing dynamic ad-hoc dependencies. Parallel  | 
|
1319  | 
performance may suffer by inspecting proof terms at run-time.  | 
|
| 52408 | 1320  | 
|
| 69597 | 1321  | 
\<^descr> \<^ML>\<open>Proofterm.proofs\<close> specifies the detail of proof recording within  | 
1322  | 
\<^ML_type>\<open>thm\<close> values produced by the inference kernel: \<^ML>\<open>0\<close> records only  | 
|
1323  | 
the names of oracles, \<^ML>\<open>1\<close> records oracle names and propositions, \<^ML>\<open>2\<close>  | 
|
| 65446 | 1324  | 
additionally records full proof terms. Officially named theorems that  | 
1325  | 
contribute to a result are recorded in any case.  | 
|
| 52408 | 1326  | 
|
| 70449 | 1327  | 
\<^descr> \<^ML>\<open>Proofterm.reconstruct_proof\<close>~\<open>thy prop prf\<close> turns the implicit  | 
| 61854 | 1328  | 
proof term \<open>prf\<close> into a full proof of the given proposition.  | 
| 
52487
 
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
 
wenzelm 
parents: 
52486 
diff
changeset
 | 
1329  | 
|
| 61854 | 1330  | 
Reconstruction may fail if \<open>prf\<close> is not a proof of \<open>prop\<close>, or if it does not  | 
1331  | 
contain sufficient information for reconstruction. Failure may only happen  | 
|
1332  | 
for proofs that are constructed manually, but not for those produced  | 
|
1333  | 
automatically by the inference kernel.  | 
|
| 52411 | 1334  | 
|
| 
70915
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70569 
diff
changeset
 | 
1335  | 
\<^descr> \<^ML>\<open>Proofterm.expand_proof\<close>~\<open>thy expand prf\<close> reconstructs and expands  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70569 
diff
changeset
 | 
1336  | 
the proofs of nested theorems according to the given \<open>expand\<close> function: a  | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70569 
diff
changeset
 | 
1337  | 
  result of @{ML \<open>SOME ""\<close>} means full expansion, @{ML \<open>SOME\<close>}~\<open>name\<close> means to
 | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70569 
diff
changeset
 | 
1338  | 
  keep the theorem node but replace its internal name, @{ML NONE} means no
 | 
| 
 
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
 
wenzelm 
parents: 
70569 
diff
changeset
 | 
1339  | 
change.  | 
| 52411 | 1340  | 
|
| 69597 | 1341  | 
\<^descr> \<^ML>\<open>Proof_Checker.thm_of_proof\<close>~\<open>thy prf\<close> turns the given (full) proof  | 
| 61854 | 1342  | 
into a theorem, by replaying it using only primitive rules of the inference  | 
1343  | 
kernel.  | 
|
| 52411 | 1344  | 
|
| 69597 | 1345  | 
\<^descr> \<^ML>\<open>Proof_Syntax.read_proof\<close>~\<open>thy b\<^sub>1 b\<^sub>2 s\<close> reads in a proof term. The  | 
| 61854 | 1346  | 
Boolean flags indicate the use of sort and type information. Usually, typing  | 
1347  | 
information is left implicit and is inferred during proof reconstruction.  | 
|
1348  | 
%FIXME eliminate flags!?  | 
|
| 52412 | 1349  | 
|
| 69597 | 1350  | 
\<^descr> \<^ML>\<open>Proof_Syntax.pretty_proof\<close>~\<open>ctxt prf\<close> pretty-prints the given proof  | 
| 61854 | 1351  | 
term.  | 
| 58618 | 1352  | 
\<close>  | 
| 52408 | 1353  | 
|
| 61854 | 1354  | 
text %mlex \<open>  | 
| 63680 | 1355  | 
\<^item> \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close> provides basic examples involving  | 
1356  | 
proof terms.  | 
|
| 52410 | 1357  | 
|
| 63680 | 1358  | 
\<^item> \<^file>\<open>~~/src/HOL/Proofs/ex/XML_Data.thy\<close> demonstrates export and import of  | 
1359  | 
proof terms via XML/ML data representation.  | 
|
| 58618 | 1360  | 
\<close>  | 
| 52410 | 1361  | 
|
| 18537 | 1362  | 
end  |