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