| author | haftmann | 
| Wed, 21 Jan 2009 16:51:45 +0100 | |
| changeset 29584 | 88ba5e5ac6d8 | 
| parent 29581 | b3b33e0298eb | 
| child 29615 | 24a58ae5dc0e | 
| permissions | -rw-r--r-- | 
| 18537 | 1 | theory logic imports base begin | 
| 2 | ||
| 20470 | 3 | chapter {* Primitive logic \label{ch:logic} *}
 | 
| 18537 | 4 | |
| 20480 | 5 | text {*
 | 
| 6 | The logical foundations of Isabelle/Isar are that of the Pure logic, | |
| 7 | which has been introduced as a natural-deduction framework in | |
| 8 |   \cite{paulson700}.  This is essentially the same logic as ``@{text
 | |
| 20493 | 9 | "\<lambda>HOL"}'' in the more abstract setting of Pure Type Systems (PTS) | 
| 20480 | 10 |   \cite{Barendregt-Geuvers:2001}, although there are some key
 | 
| 20491 | 11 | differences in the specific treatment of simple types in | 
| 12 | Isabelle/Pure. | |
| 20480 | 13 | |
| 14 | Following type-theoretic parlance, the Pure logic consists of three | |
| 20543 | 15 |   levels of @{text "\<lambda>"}-calculus with corresponding arrows, @{text
 | 
| 20480 | 16 |   "\<Rightarrow>"} for syntactic function space (terms depending on terms), @{text
 | 
| 17 | "\<And>"} for universal quantification (proofs depending on terms), and | |
| 18 |   @{text "\<Longrightarrow>"} for implication (proofs depending on proofs).
 | |
| 19 | ||
| 20537 | 20 | Derivations are relative to a logical theory, which declares type | 
| 21 | constructors, constants, and axioms. Theory declarations support | |
| 22 | schematic polymorphism, which is strictly speaking outside the | |
| 23 |   logic.\footnote{This is the deeper logical reason, why the theory
 | |
| 24 |   context @{text "\<Theta>"} is separate from the proof context @{text "\<Gamma>"}
 | |
| 25 | of the core calculus.} | |
| 20480 | 26 | *} | 
| 27 | ||
| 28 | ||
| 20451 | 29 | section {* Types \label{sec:types} *}
 | 
| 20437 | 30 | |
| 31 | text {*
 | |
| 20480 | 32 | The language of types is an uninterpreted order-sorted first-order | 
| 33 | algebra; types are qualified by ordered type classes. | |
| 34 | ||
| 35 |   \medskip A \emph{type class} is an abstract syntactic entity
 | |
| 36 |   declared in the theory context.  The \emph{subclass relation} @{text
 | |
| 37 | "c\<^isub>1 \<subseteq> c\<^isub>2"} is specified by stating an acyclic | |
| 20491 | 38 | generating relation; the transitive closure is maintained | 
| 39 | internally. The resulting relation is an ordering: reflexive, | |
| 40 | transitive, and antisymmetric. | |
| 20451 | 41 | |
| 20537 | 42 |   A \emph{sort} is a list of type classes written as @{text "s =
 | 
| 43 |   {c\<^isub>1, \<dots>, c\<^isub>m}"}, which represents symbolic
 | |
| 20480 | 44 | intersection. Notationally, the curly braces are omitted for | 
| 45 |   singleton intersections, i.e.\ any class @{text "c"} may be read as
 | |
| 46 |   a sort @{text "{c}"}.  The ordering on type classes is extended to
 | |
| 20491 | 47 |   sorts according to the meaning of intersections: @{text
 | 
| 48 |   "{c\<^isub>1, \<dots> c\<^isub>m} \<subseteq> {d\<^isub>1, \<dots>, d\<^isub>n}"} iff
 | |
| 49 |   @{text "\<forall>j. \<exists>i. c\<^isub>i \<subseteq> d\<^isub>j"}.  The empty intersection
 | |
| 50 |   @{text "{}"} refers to the universal sort, which is the largest
 | |
| 51 | element wrt.\ the sort order. The intersections of all (finitely | |
| 52 | many) classes declared in the current theory are the minimal | |
| 53 | elements wrt.\ the sort order. | |
| 20480 | 54 | |
| 20491 | 55 |   \medskip A \emph{fixed type variable} is a pair of a basic name
 | 
| 20537 | 56 |   (starting with a @{text "'"} character) and a sort constraint, e.g.\
 | 
| 57 |   @{text "('a, s)"} which is usually printed as @{text "\<alpha>\<^isub>s"}.
 | |
| 58 |   A \emph{schematic type variable} is a pair of an indexname and a
 | |
| 59 |   sort constraint, e.g.\ @{text "(('a, 0), s)"} which is usually
 | |
| 60 |   printed as @{text "?\<alpha>\<^isub>s"}.
 | |
| 20451 | 61 | |
| 20480 | 62 |   Note that \emph{all} syntactic components contribute to the identity
 | 
| 20493 | 63 | of type variables, including the sort constraint. The core logic | 
| 64 | handles type variables with the same name but different sorts as | |
| 65 | different, although some outer layers of the system make it hard to | |
| 66 | produce anything like this. | |
| 20451 | 67 | |
| 20493 | 68 |   A \emph{type constructor} @{text "\<kappa>"} is a @{text "k"}-ary operator
 | 
| 69 | on types declared in the theory. Type constructor application is | |
| 20537 | 70 |   written postfix as @{text "(\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>k)\<kappa>"}.  For
 | 
| 71 |   @{text "k = 0"} the argument tuple is omitted, e.g.\ @{text "prop"}
 | |
| 72 |   instead of @{text "()prop"}.  For @{text "k = 1"} the parentheses
 | |
| 73 |   are omitted, e.g.\ @{text "\<alpha> list"} instead of @{text "(\<alpha>)list"}.
 | |
| 74 | Further notation is provided for specific constructors, notably the | |
| 75 |   right-associative infix @{text "\<alpha> \<Rightarrow> \<beta>"} instead of @{text "(\<alpha>,
 | |
| 76 | \<beta>)fun"}. | |
| 20480 | 77 | |
| 20537 | 78 |   A \emph{type} is defined inductively over type variables and type
 | 
| 79 |   constructors as follows: @{text "\<tau> = \<alpha>\<^isub>s | ?\<alpha>\<^isub>s |
 | |
| 20543 | 80 | (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"}. | 
| 20451 | 81 | |
| 20514 | 82 |   A \emph{type abbreviation} is a syntactic definition @{text
 | 
| 20494 | 83 |   "(\<^vec>\<alpha>)\<kappa> = \<tau>"} of an arbitrary type expression @{text "\<tau>"} over
 | 
| 20537 | 84 |   variables @{text "\<^vec>\<alpha>"}.  Type abbreviations appear as type
 | 
| 85 | constructors in the syntax, but are expanded before entering the | |
| 86 | logical core. | |
| 20480 | 87 | |
| 88 |   A \emph{type arity} declares the image behavior of a type
 | |
| 20494 | 89 |   constructor wrt.\ the algebra of sorts: @{text "\<kappa> :: (s\<^isub>1, \<dots>,
 | 
| 90 |   s\<^isub>k)s"} means that @{text "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>k)\<kappa>"} is
 | |
| 91 |   of sort @{text "s"} if every argument type @{text "\<tau>\<^isub>i"} is
 | |
| 92 |   of sort @{text "s\<^isub>i"}.  Arity declarations are implicitly
 | |
| 93 |   completed, i.e.\ @{text "\<kappa> :: (\<^vec>s)c"} entails @{text "\<kappa> ::
 | |
| 20491 | 94 |   (\<^vec>s)c'"} for any @{text "c' \<supseteq> c"}.
 | 
| 95 | ||
| 96 |   \medskip The sort algebra is always maintained as \emph{coregular},
 | |
| 97 | which means that type arities are consistent with the subclass | |
| 20537 | 98 |   relation: for any type constructor @{text "\<kappa>"}, and classes @{text
 | 
| 99 |   "c\<^isub>1 \<subseteq> c\<^isub>2"}, and arities @{text "\<kappa> ::
 | |
| 100 |   (\<^vec>s\<^isub>1)c\<^isub>1"} and @{text "\<kappa> ::
 | |
| 101 |   (\<^vec>s\<^isub>2)c\<^isub>2"} holds @{text "\<^vec>s\<^isub>1 \<subseteq>
 | |
| 102 | \<^vec>s\<^isub>2"} component-wise. | |
| 20451 | 103 | |
| 20491 | 104 | The key property of a coregular order-sorted algebra is that sort | 
| 20537 | 105 | constraints can be solved in a most general fashion: for each type | 
| 106 |   constructor @{text "\<kappa>"} and sort @{text "s"} there is a most general
 | |
| 107 |   vector of argument sorts @{text "(s\<^isub>1, \<dots>, s\<^isub>k)"} such
 | |
| 108 |   that a type scheme @{text "(\<alpha>\<^bsub>s\<^isub>1\<^esub>, \<dots>,
 | |
| 109 |   \<alpha>\<^bsub>s\<^isub>k\<^esub>)\<kappa>"} is of sort @{text "s"}.
 | |
| 20543 | 110 | Consequently, type unification has most general solutions (modulo | 
| 111 | equivalence of sorts), so type-inference produces primary types as | |
| 112 |   expected \cite{nipkow-prehofer}.
 | |
| 20480 | 113 | *} | 
| 20451 | 114 | |
| 20480 | 115 | text %mlref {*
 | 
| 116 |   \begin{mldecls}
 | |
| 117 |   @{index_ML_type class} \\
 | |
| 118 |   @{index_ML_type sort} \\
 | |
| 20494 | 119 |   @{index_ML_type arity} \\
 | 
| 20480 | 120 |   @{index_ML_type typ} \\
 | 
| 20519 | 121 |   @{index_ML map_atyps: "(typ -> typ) -> typ -> typ"} \\
 | 
| 20494 | 122 |   @{index_ML fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
 | 
| 20547 | 123 |   \end{mldecls}
 | 
| 124 |   \begin{mldecls}
 | |
| 20480 | 125 |   @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
 | 
| 126 |   @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
 | |
| 20520 | 127 |   @{index_ML Sign.add_types: "(string * int * mixfix) list -> theory -> theory"} \\
 | 
| 20480 | 128 |   @{index_ML Sign.add_tyabbrs_i: "
 | 
| 20520 | 129 | (string * string list * typ * mixfix) list -> theory -> theory"} \\ | 
| 20480 | 130 |   @{index_ML Sign.primitive_class: "string * class list -> theory -> theory"} \\
 | 
| 131 |   @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
 | |
| 132 |   @{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
 | |
| 133 |   \end{mldecls}
 | |
| 134 | ||
| 135 |   \begin{description}
 | |
| 136 | ||
| 137 |   \item @{ML_type class} represents type classes; this is an alias for
 | |
| 138 |   @{ML_type string}.
 | |
| 139 | ||
| 140 |   \item @{ML_type sort} represents sorts; this is an alias for
 | |
| 141 |   @{ML_type "class list"}.
 | |
| 20451 | 142 | |
| 20480 | 143 |   \item @{ML_type arity} represents type arities; this is an alias for
 | 
| 20494 | 144 |   triples of the form @{text "(\<kappa>, \<^vec>s, s)"} for @{text "\<kappa> ::
 | 
| 20480 | 145 | (\<^vec>s)s"} described above. | 
| 146 | ||
| 147 |   \item @{ML_type typ} represents types; this is a datatype with
 | |
| 148 |   constructors @{ML TFree}, @{ML TVar}, @{ML Type}.
 | |
| 149 | ||
| 20537 | 150 |   \item @{ML map_atyps}~@{text "f \<tau>"} applies the mapping @{text "f"}
 | 
| 151 |   to all atomic types (@{ML TFree}, @{ML TVar}) occurring in @{text
 | |
| 152 | "\<tau>"}. | |
| 20519 | 153 | |
| 20537 | 154 |   \item @{ML fold_atyps}~@{text "f \<tau>"} iterates the operation @{text
 | 
| 155 |   "f"} over all occurrences of atomic types (@{ML TFree}, @{ML TVar})
 | |
| 156 |   in @{text "\<tau>"}; the type structure is traversed from left to right.
 | |
| 20494 | 157 | |
| 20480 | 158 |   \item @{ML Sign.subsort}~@{text "thy (s\<^isub>1, s\<^isub>2)"}
 | 
| 159 |   tests the subsort relation @{text "s\<^isub>1 \<subseteq> s\<^isub>2"}.
 | |
| 160 | ||
| 20537 | 161 |   \item @{ML Sign.of_sort}~@{text "thy (\<tau>, s)"} tests whether type
 | 
| 162 |   @{text "\<tau>"} is of sort @{text "s"}.
 | |
| 20480 | 163 | |
| 20537 | 164 |   \item @{ML Sign.add_types}~@{text "[(\<kappa>, k, mx), \<dots>]"} declares a new
 | 
| 20494 | 165 |   type constructors @{text "\<kappa>"} with @{text "k"} arguments and
 | 
| 20480 | 166 | optional mixfix syntax. | 
| 20451 | 167 | |
| 20494 | 168 |   \item @{ML Sign.add_tyabbrs_i}~@{text "[(\<kappa>, \<^vec>\<alpha>, \<tau>, mx), \<dots>]"}
 | 
| 169 |   defines a new type abbreviation @{text "(\<^vec>\<alpha>)\<kappa> = \<tau>"} with
 | |
| 20491 | 170 | optional mixfix syntax. | 
| 20480 | 171 | |
| 172 |   \item @{ML Sign.primitive_class}~@{text "(c, [c\<^isub>1, \<dots>,
 | |
| 20537 | 173 |   c\<^isub>n])"} declares a new class @{text "c"}, together with class
 | 
| 20494 | 174 |   relations @{text "c \<subseteq> c\<^isub>i"}, for @{text "i = 1, \<dots>, n"}.
 | 
| 20480 | 175 | |
| 176 |   \item @{ML Sign.primitive_classrel}~@{text "(c\<^isub>1,
 | |
| 20543 | 177 |   c\<^isub>2)"} declares the class relation @{text "c\<^isub>1 \<subseteq>
 | 
| 20480 | 178 | c\<^isub>2"}. | 
| 179 | ||
| 20494 | 180 |   \item @{ML Sign.primitive_arity}~@{text "(\<kappa>, \<^vec>s, s)"} declares
 | 
| 20537 | 181 |   the arity @{text "\<kappa> :: (\<^vec>s)s"}.
 | 
| 20480 | 182 | |
| 183 |   \end{description}
 | |
| 20437 | 184 | *} | 
| 185 | ||
| 186 | ||
| 20480 | 187 | |
| 20451 | 188 | section {* Terms \label{sec:terms} *}
 | 
| 18537 | 189 | |
| 190 | text {*
 | |
| 20451 | 191 |   \glossary{Term}{FIXME}
 | 
| 18537 | 192 | |
| 20491 | 193 |   The language of terms is that of simply-typed @{text "\<lambda>"}-calculus
 | 
| 20520 | 194 |   with de-Bruijn indices for bound variables (cf.\ \cite{debruijn72}
 | 
| 20537 | 195 |   or \cite{paulson-ml2}), with the types being determined determined
 | 
| 196 | by the corresponding binders. In contrast, free variables and | |
| 197 | constants are have an explicit name and type in each occurrence. | |
| 20514 | 198 | |
| 199 |   \medskip A \emph{bound variable} is a natural number @{text "b"},
 | |
| 20537 | 200 | which accounts for the number of intermediate binders between the | 
| 201 | variable occurrence in the body and its binding position. For | |
| 20543 | 202 |   example, the de-Bruijn term @{text
 | 
| 203 | "\<lambda>\<^bsub>nat\<^esub>. \<lambda>\<^bsub>nat\<^esub>. 1 + 0"} would | |
| 204 |   correspond to @{text
 | |
| 205 | "\<lambda>x\<^bsub>nat\<^esub>. \<lambda>y\<^bsub>nat\<^esub>. x + y"} in a named | |
| 206 | representation. Note that a bound variable may be represented by | |
| 207 | different de-Bruijn indices at different occurrences, depending on | |
| 208 | the nesting of abstractions. | |
| 20537 | 209 | |
| 20543 | 210 |   A \emph{loose variable} is a bound variable that is outside the
 | 
| 20537 | 211 | scope of local binders. The types (and names) for loose variables | 
| 20543 | 212 | can be managed as a separate context, that is maintained as a stack | 
| 213 | of hypothetical binders. The core logic operates on closed terms, | |
| 214 | without any loose variables. | |
| 20514 | 215 | |
| 20537 | 216 |   A \emph{fixed variable} is a pair of a basic name and a type, e.g.\
 | 
| 217 |   @{text "(x, \<tau>)"} which is usually printed @{text "x\<^isub>\<tau>"}.  A
 | |
| 218 |   \emph{schematic variable} is a pair of an indexname and a type,
 | |
| 219 |   e.g.\ @{text "((x, 0), \<tau>)"} which is usually printed as @{text
 | |
| 220 | "?x\<^isub>\<tau>"}. | |
| 20491 | 221 | |
| 20537 | 222 |   \medskip A \emph{constant} is a pair of a basic name and a type,
 | 
| 223 |   e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text
 | |
| 224 | "c\<^isub>\<tau>"}. Constants are declared in the context as polymorphic | |
| 20543 | 225 |   families @{text "c :: \<sigma>"}, meaning that all substitution instances
 | 
| 226 |   @{text "c\<^isub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
 | |
| 20514 | 227 | |
| 20537 | 228 |   The vector of \emph{type arguments} of constant @{text "c\<^isub>\<tau>"}
 | 
| 229 |   wrt.\ the declaration @{text "c :: \<sigma>"} is defined as the codomain of
 | |
| 230 |   the matcher @{text "\<vartheta> = {?\<alpha>\<^isub>1 \<mapsto> \<tau>\<^isub>1, \<dots>,
 | |
| 231 |   ?\<alpha>\<^isub>n \<mapsto> \<tau>\<^isub>n}"} presented in canonical order @{text
 | |
| 232 | "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n)"}. Within a given theory context, | |
| 233 |   there is a one-to-one correspondence between any constant @{text
 | |
| 234 |   "c\<^isub>\<tau>"} and the application @{text "c(\<tau>\<^isub>1, \<dots>,
 | |
| 235 |   \<tau>\<^isub>n)"} of its type arguments.  For example, with @{text "plus
 | |
| 236 |   :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"}, the instance @{text "plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow>
 | |
| 237 |   nat\<^esub>"} corresponds to @{text "plus(nat)"}.
 | |
| 20480 | 238 | |
| 20514 | 239 |   Constant declarations @{text "c :: \<sigma>"} may contain sort constraints
 | 
| 240 |   for type variables in @{text "\<sigma>"}.  These are observed by
 | |
| 241 |   type-inference as expected, but \emph{ignored} by the core logic.
 | |
| 242 | This means the primitive logic is able to reason with instances of | |
| 20537 | 243 | polymorphic constants that the user-level type-checker would reject | 
| 244 | due to violation of type class restrictions. | |
| 20480 | 245 | |
| 20543 | 246 |   \medskip An \emph{atomic} term is either a variable or constant.  A
 | 
| 247 |   \emph{term} is defined inductively over atomic terms, with
 | |
| 248 |   abstraction and application as follows: @{text "t = b | x\<^isub>\<tau> |
 | |
| 249 | ?x\<^isub>\<tau> | c\<^isub>\<tau> | \<lambda>\<^isub>\<tau>. t | t\<^isub>1 t\<^isub>2"}. | |
| 250 | Parsing and printing takes care of converting between an external | |
| 251 | representation with named bound variables. Subsequently, we shall | |
| 252 | use the latter notation instead of internal de-Bruijn | |
| 253 | representation. | |
| 20498 | 254 | |
| 20537 | 255 |   The inductive relation @{text "t :: \<tau>"} assigns a (unique) type to a
 | 
| 256 | term according to the structure of atomic terms, abstractions, and | |
| 257 | applicatins: | |
| 20498 | 258 | \[ | 
| 20501 | 259 |   \infer{@{text "a\<^isub>\<tau> :: \<tau>"}}{}
 | 
| 20498 | 260 | \qquad | 
| 20501 | 261 |   \infer{@{text "(\<lambda>x\<^sub>\<tau>. t) :: \<tau> \<Rightarrow> \<sigma>"}}{@{text "t :: \<sigma>"}}
 | 
| 262 | \qquad | |
| 263 |   \infer{@{text "t u :: \<sigma>"}}{@{text "t :: \<tau> \<Rightarrow> \<sigma>"} & @{text "u :: \<tau>"}}
 | |
| 20498 | 264 | \] | 
| 20514 | 265 |   A \emph{well-typed term} is a term that can be typed according to these rules.
 | 
| 20498 | 266 | |
| 20514 | 267 | Typing information can be omitted: type-inference is able to | 
| 268 | reconstruct the most general type of a raw term, while assigning | |
| 269 | most general types to all of its variables and constants. | |
| 270 | Type-inference depends on a context of type constraints for fixed | |
| 271 | variables, and declarations for polymorphic constants. | |
| 272 | ||
| 20537 | 273 | The identity of atomic terms consists both of the name and the type | 
| 274 |   component.  This means that different variables @{text
 | |
| 275 |   "x\<^bsub>\<tau>\<^isub>1\<^esub>"} and @{text
 | |
| 276 | "x\<^bsub>\<tau>\<^isub>2\<^esub>"} may become the same after type | |
| 277 | instantiation. Some outer layers of the system make it hard to | |
| 278 | produce variables of the same name, but different types. In | |
| 20543 | 279 | contrast, mixed instances of polymorphic constants occur frequently. | 
| 20514 | 280 | |
| 281 |   \medskip The \emph{hidden polymorphism} of a term @{text "t :: \<sigma>"}
 | |
| 282 |   is the set of type variables occurring in @{text "t"}, but not in
 | |
| 20537 | 283 |   @{text "\<sigma>"}.  This means that the term implicitly depends on type
 | 
| 20543 | 284 | arguments that are not accounted in the result type, i.e.\ there are | 
| 20537 | 285 |   different type instances @{text "t\<vartheta> :: \<sigma>"} and @{text
 | 
| 286 | "t\<vartheta>' :: \<sigma>"} with the same type. This slightly | |
| 20543 | 287 | pathological situation notoriously demands additional care. | 
| 20514 | 288 | |
| 289 |   \medskip A \emph{term abbreviation} is a syntactic definition @{text
 | |
| 20537 | 290 |   "c\<^isub>\<sigma> \<equiv> t"} of a closed term @{text "t"} of type @{text "\<sigma>"},
 | 
| 291 | without any hidden polymorphism. A term abbreviation looks like a | |
| 20543 | 292 | constant in the syntax, but is expanded before entering the logical | 
| 293 | core. Abbreviations are usually reverted when printing terms, using | |
| 294 |   @{text "t \<rightarrow> c\<^isub>\<sigma>"} as rules for higher-order rewriting.
 | |
| 20519 | 295 | |
| 296 |   \medskip Canonical operations on @{text "\<lambda>"}-terms include @{text
 | |
| 20537 | 297 |   "\<alpha>\<beta>\<eta>"}-conversion: @{text "\<alpha>"}-conversion refers to capture-free
 | 
| 20519 | 298 |   renaming of bound variables; @{text "\<beta>"}-conversion contracts an
 | 
| 20537 | 299 | abstraction applied to an argument term, substituting the argument | 
| 20519 | 300 |   in the body: @{text "(\<lambda>x. b)a"} becomes @{text "b[a/x]"}; @{text
 | 
| 301 |   "\<eta>"}-conversion contracts vacuous application-abstraction: @{text
 | |
| 302 |   "\<lambda>x. f x"} becomes @{text "f"}, provided that the bound variable
 | |
| 20537 | 303 |   does not occur in @{text "f"}.
 | 
| 20519 | 304 | |
| 20537 | 305 |   Terms are normally treated modulo @{text "\<alpha>"}-conversion, which is
 | 
| 306 | implicit in the de-Bruijn representation. Names for bound variables | |
| 307 | in abstractions are maintained separately as (meaningless) comments, | |
| 308 |   mostly for parsing and printing.  Full @{text "\<alpha>\<beta>\<eta>"}-conversion is
 | |
| 28784 | 309 |   commonplace in various standard operations (\secref{sec:obj-rules})
 | 
| 310 | that are based on higher-order unification and matching. | |
| 18537 | 311 | *} | 
| 312 | ||
| 20514 | 313 | text %mlref {*
 | 
| 314 |   \begin{mldecls}
 | |
| 315 |   @{index_ML_type term} \\
 | |
| 20519 | 316 |   @{index_ML "op aconv": "term * term -> bool"} \\
 | 
| 20547 | 317 |   @{index_ML map_types: "(typ -> typ) -> term -> term"} \\
 | 
| 20519 | 318 |   @{index_ML fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
 | 
| 20514 | 319 |   @{index_ML map_aterms: "(term -> term) -> term -> term"} \\
 | 
| 320 |   @{index_ML fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
 | |
| 20547 | 321 |   \end{mldecls}
 | 
| 322 |   \begin{mldecls}
 | |
| 20514 | 323 |   @{index_ML fastype_of: "term -> typ"} \\
 | 
| 20519 | 324 |   @{index_ML lambda: "term -> term -> term"} \\
 | 
| 325 |   @{index_ML betapply: "term * term -> term"} \\
 | |
| 29581 | 326 |   @{index_ML Sign.declare_const: "Properties.T -> (binding * typ) * mixfix ->
 | 
| 24972 
acafb18a47dc
replaced Sign.add_consts_i by Sign.declare_const;
 wenzelm parents: 
24828diff
changeset | 327 | theory -> term * theory"} \\ | 
| 29581 | 328 |   @{index_ML Sign.add_abbrev: "string -> Properties.T -> binding * term ->
 | 
| 24972 
acafb18a47dc
replaced Sign.add_consts_i by Sign.declare_const;
 wenzelm parents: 
24828diff
changeset | 329 | theory -> (term * term) * theory"} \\ | 
| 20519 | 330 |   @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
 | 
| 331 |   @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
 | |
| 20514 | 332 |   \end{mldecls}
 | 
| 18537 | 333 | |
| 20514 | 334 |   \begin{description}
 | 
| 18537 | 335 | |
| 20537 | 336 |   \item @{ML_type term} represents de-Bruijn terms, with comments in
 | 
| 337 | abstractions, and explicitly named free variables and constants; | |
| 338 |   this is a datatype with constructors @{ML Bound}, @{ML Free}, @{ML
 | |
| 339 |   Var}, @{ML Const}, @{ML Abs}, @{ML "op $"}.
 | |
| 20519 | 340 | |
| 341 |   \item @{text "t"}~@{ML aconv}~@{text "u"} checks @{text
 | |
| 342 | "\<alpha>"}-equivalence of two terms. This is the basic equality relation | |
| 343 |   on type @{ML_type term}; raw datatype equality should only be used
 | |
| 344 | for operations related to parsing or printing! | |
| 345 | ||
| 20547 | 346 |   \item @{ML map_types}~@{text "f t"} applies the mapping @{text
 | 
| 20537 | 347 |   "f"} to all types occurring in @{text "t"}.
 | 
| 348 | ||
| 349 |   \item @{ML fold_types}~@{text "f t"} iterates the operation @{text
 | |
| 350 |   "f"} over all occurrences of types in @{text "t"}; the term
 | |
| 351 | structure is traversed from left to right. | |
| 20519 | 352 | |
| 20537 | 353 |   \item @{ML map_aterms}~@{text "f t"} applies the mapping @{text "f"}
 | 
| 354 |   to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML
 | |
| 355 |   Const}) occurring in @{text "t"}.
 | |
| 356 | ||
| 357 |   \item @{ML fold_aterms}~@{text "f t"} iterates the operation @{text
 | |
| 358 |   "f"} over all occurrences of atomic terms (@{ML Bound}, @{ML Free},
 | |
| 359 |   @{ML Var}, @{ML Const}) in @{text "t"}; the term structure is
 | |
| 20519 | 360 | traversed from left to right. | 
| 361 | ||
| 20537 | 362 |   \item @{ML fastype_of}~@{text "t"} determines the type of a
 | 
| 363 | well-typed term. This operation is relatively slow, despite the | |
| 364 | omission of any sanity checks. | |
| 20519 | 365 | |
| 366 |   \item @{ML lambda}~@{text "a b"} produces an abstraction @{text
 | |
| 20537 | 367 |   "\<lambda>a. b"}, where occurrences of the atomic term @{text "a"} in the
 | 
| 368 |   body @{text "b"} are replaced by bound variables.
 | |
| 20519 | 369 | |
| 20537 | 370 |   \item @{ML betapply}~@{text "(t, u)"} produces an application @{text
 | 
| 371 |   "t u"}, with topmost @{text "\<beta>"}-conversion if @{text "t"} is an
 | |
| 372 | abstraction. | |
| 20519 | 373 | |
| 28110 | 374 |   \item @{ML Sign.declare_const}~@{text "properties ((c, \<sigma>), mx)"}
 | 
| 24972 
acafb18a47dc
replaced Sign.add_consts_i by Sign.declare_const;
 wenzelm parents: 
24828diff
changeset | 375 |   declares a new constant @{text "c :: \<sigma>"} with optional mixfix
 | 
| 
acafb18a47dc
replaced Sign.add_consts_i by Sign.declare_const;
 wenzelm parents: 
24828diff
changeset | 376 | syntax. | 
| 20519 | 377 | |
| 24828 | 378 |   \item @{ML Sign.add_abbrev}~@{text "print_mode properties (c, t)"}
 | 
| 21827 | 379 |   introduces a new term abbreviation @{text "c \<equiv> t"}.
 | 
| 20519 | 380 | |
| 20520 | 381 |   \item @{ML Sign.const_typargs}~@{text "thy (c, \<tau>)"} and @{ML
 | 
| 382 |   Sign.const_instance}~@{text "thy (c, [\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n])"}
 | |
| 20543 | 383 | convert between two representations of polymorphic constants: full | 
| 384 | type instance vs.\ compact type arguments form. | |
| 18537 | 385 | |
| 20514 | 386 |   \end{description}
 | 
| 18537 | 387 | *} | 
| 388 | ||
| 389 | ||
| 20451 | 390 | section {* Theorems \label{sec:thms} *}
 | 
| 18537 | 391 | |
| 392 | text {*
 | |
| 20521 | 393 |   \glossary{Proposition}{FIXME A \seeglossary{term} of
 | 
| 394 |   \seeglossary{type} @{text "prop"}.  Internally, there is nothing
 | |
| 395 | special about propositions apart from their type, but the concrete | |
| 396 | syntax enforces a clear distinction. Propositions are structured | |
| 397 |   via implication @{text "A \<Longrightarrow> B"} or universal quantification @{text
 | |
| 398 | "\<And>x. B x"} --- anything else is considered atomic. The canonical | |
| 399 |   form for propositions is that of a \seeglossary{Hereditary Harrop
 | |
| 400 | Formula}. FIXME} | |
| 20480 | 401 | |
| 20501 | 402 |   \glossary{Theorem}{A proven proposition within a certain theory and
 | 
| 403 |   proof context, formally @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}; both contexts are
 | |
| 404 | rarely spelled out explicitly. Theorems are usually normalized | |
| 405 |   according to the \seeglossary{HHF} format. FIXME}
 | |
| 20480 | 406 | |
| 20519 | 407 |   \glossary{Fact}{Sometimes used interchangeably for
 | 
| 20501 | 408 |   \seeglossary{theorem}.  Strictly speaking, a list of theorems,
 | 
| 409 | essentially an extra-logical conjunction. Facts emerge either as | |
| 410 | local assumptions, or as results of local goal statements --- both | |
| 411 | may be simultaneous, hence the list representation. FIXME} | |
| 18537 | 412 | |
| 20501 | 413 |   \glossary{Schematic variable}{FIXME}
 | 
| 414 | ||
| 415 |   \glossary{Fixed variable}{A variable that is bound within a certain
 | |
| 416 | proof context; an arbitrary-but-fixed entity within a portion of | |
| 417 | proof text. FIXME} | |
| 18537 | 418 | |
| 20501 | 419 |   \glossary{Free variable}{Synonymous for \seeglossary{fixed
 | 
| 420 | variable}. FIXME} | |
| 421 | ||
| 422 |   \glossary{Bound variable}{FIXME}
 | |
| 18537 | 423 | |
| 20501 | 424 |   \glossary{Variable}{See \seeglossary{schematic variable},
 | 
| 425 |   \seeglossary{fixed variable}, \seeglossary{bound variable}, or
 | |
| 426 |   \seeglossary{type variable}.  The distinguishing feature of
 | |
| 427 | different variables is their binding scope. FIXME} | |
| 18537 | 428 | |
| 20543 | 429 |   A \emph{proposition} is a well-typed term of type @{text "prop"}, a
 | 
| 20521 | 430 |   \emph{theorem} is a proven proposition (depending on a context of
 | 
| 431 | hypotheses and the background theory). Primitive inferences include | |
| 432 |   plain natural deduction rules for the primary connectives @{text
 | |
| 20537 | 433 |   "\<And>"} and @{text "\<Longrightarrow>"} of the framework.  There is also a builtin
 | 
| 434 |   notion of equality/equivalence @{text "\<equiv>"}.
 | |
| 20521 | 435 | *} | 
| 436 | ||
| 26872 | 437 | subsection {* Primitive connectives and rules \label{sec:prim-rules} *}
 | 
| 18537 | 438 | |
| 20521 | 439 | text {*
 | 
| 20543 | 440 |   The theory @{text "Pure"} contains constant declarations for the
 | 
| 441 |   primitive connectives @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"} of
 | |
| 442 |   the logical framework, see \figref{fig:pure-connectives}.  The
 | |
| 443 |   derivability judgment @{text "A\<^isub>1, \<dots>, A\<^isub>n \<turnstile> B"} is
 | |
| 444 | defined inductively by the primitive inferences given in | |
| 445 |   \figref{fig:prim-rules}, with the global restriction that the
 | |
| 446 |   hypotheses must \emph{not} contain any schematic variables.  The
 | |
| 447 | builtin equality is conceptually axiomatized as shown in | |
| 20521 | 448 |   \figref{fig:pure-equality}, although the implementation works
 | 
| 20543 | 449 | directly with derived inferences. | 
| 20521 | 450 | |
| 451 |   \begin{figure}[htb]
 | |
| 452 |   \begin{center}
 | |
| 20501 | 453 |   \begin{tabular}{ll}
 | 
| 454 |   @{text "all :: (\<alpha> \<Rightarrow> prop) \<Rightarrow> prop"} & universal quantification (binder @{text "\<And>"}) \\
 | |
| 455 |   @{text "\<Longrightarrow> :: prop \<Rightarrow> prop \<Rightarrow> prop"} & implication (right associative infix) \\
 | |
| 20521 | 456 |   @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop"} & equality relation (infix) \\
 | 
| 20501 | 457 |   \end{tabular}
 | 
| 20537 | 458 |   \caption{Primitive connectives of Pure}\label{fig:pure-connectives}
 | 
| 20521 | 459 |   \end{center}
 | 
| 460 |   \end{figure}
 | |
| 18537 | 461 | |
| 20501 | 462 |   \begin{figure}[htb]
 | 
| 463 |   \begin{center}
 | |
| 20498 | 464 | \[ | 
| 465 |   \infer[@{text "(axiom)"}]{@{text "\<turnstile> A"}}{@{text "A \<in> \<Theta>"}}
 | |
| 466 | \qquad | |
| 467 |   \infer[@{text "(assume)"}]{@{text "A \<turnstile> A"}}{}
 | |
| 468 | \] | |
| 469 | \[ | |
| 20537 | 470 |   \infer[@{text "(\<And>_intro)"}]{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}{@{text "\<Gamma> \<turnstile> b[x]"} & @{text "x \<notin> \<Gamma>"}}
 | 
| 20498 | 471 | \qquad | 
| 20537 | 472 |   \infer[@{text "(\<And>_elim)"}]{@{text "\<Gamma> \<turnstile> b[a]"}}{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}
 | 
| 20498 | 473 | \] | 
| 474 | \[ | |
| 475 |   \infer[@{text "(\<Longrightarrow>_intro)"}]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
 | |
| 476 | \qquad | |
| 477 |   \infer[@{text "(\<Longrightarrow>_elim)"}]{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
 | |
| 478 | \] | |
| 20521 | 479 |   \caption{Primitive inferences of Pure}\label{fig:prim-rules}
 | 
| 480 |   \end{center}
 | |
| 481 |   \end{figure}
 | |
| 482 | ||
| 483 |   \begin{figure}[htb]
 | |
| 484 |   \begin{center}
 | |
| 485 |   \begin{tabular}{ll}
 | |
| 20537 | 486 |   @{text "\<turnstile> (\<lambda>x. b[x]) a \<equiv> b[a]"} & @{text "\<beta>"}-conversion \\
 | 
| 20521 | 487 |   @{text "\<turnstile> x \<equiv> x"} & reflexivity \\
 | 
| 488 |   @{text "\<turnstile> x \<equiv> y \<Longrightarrow> P x \<Longrightarrow> P y"} & substitution \\
 | |
| 489 |   @{text "\<turnstile> (\<And>x. f x \<equiv> g x) \<Longrightarrow> f \<equiv> g"} & extensionality \\
 | |
| 20537 | 490 |   @{text "\<turnstile> (A \<Longrightarrow> B) \<Longrightarrow> (B \<Longrightarrow> A) \<Longrightarrow> A \<equiv> B"} & logical equivalence \\
 | 
| 20521 | 491 |   \end{tabular}
 | 
| 20542 | 492 |   \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality}
 | 
| 20501 | 493 |   \end{center}
 | 
| 494 |   \end{figure}
 | |
| 18537 | 495 | |
| 20501 | 496 |   The introduction and elimination rules for @{text "\<And>"} and @{text
 | 
| 20537 | 497 |   "\<Longrightarrow>"} are analogous to formation of dependently typed @{text
 | 
| 20501 | 498 | "\<lambda>"}-terms representing the underlying proof objects. Proof terms | 
| 20543 | 499 | are irrelevant in the Pure logic, though; they cannot occur within | 
| 500 | propositions. The system provides a runtime option to record | |
| 20537 | 501 | explicit proof terms for primitive inferences. Thus all three | 
| 502 |   levels of @{text "\<lambda>"}-calculus become explicit: @{text "\<Rightarrow>"} for
 | |
| 503 |   terms, and @{text "\<And>/\<Longrightarrow>"} for proofs (cf.\
 | |
| 504 |   \cite{Berghofer-Nipkow:2000:TPHOL}).
 | |
| 20491 | 505 | |
| 20537 | 506 |   Observe that locally fixed parameters (as in @{text "\<And>_intro"}) need
 | 
| 507 | not be recorded in the hypotheses, because the simple syntactic | |
| 20543 | 508 |   types of Pure are always inhabitable.  ``Assumptions'' @{text "x ::
 | 
| 509 |   \<tau>"} for type-membership are only present as long as some @{text
 | |
| 510 |   "x\<^isub>\<tau>"} occurs in the statement body.\footnote{This is the key
 | |
| 511 |   difference to ``@{text "\<lambda>HOL"}'' in the PTS framework
 | |
| 512 |   \cite{Barendregt-Geuvers:2001}, where hypotheses @{text "x : A"} are
 | |
| 513 | treated uniformly for propositions and types.} | |
| 20501 | 514 | |
| 515 | \medskip The axiomatization of a theory is implicitly closed by | |
| 20537 | 516 |   forming all instances of type and term variables: @{text "\<turnstile>
 | 
| 517 | A\<vartheta>"} holds for any substitution instance of an axiom | |
| 20543 | 518 |   @{text "\<turnstile> A"}.  By pushing substitutions through derivations
 | 
| 519 |   inductively, we also get admissible @{text "generalize"} and @{text
 | |
| 520 |   "instance"} rules as shown in \figref{fig:subst-rules}.
 | |
| 20501 | 521 | |
| 522 |   \begin{figure}[htb]
 | |
| 523 |   \begin{center}
 | |
| 20498 | 524 | \[ | 
| 20501 | 525 |   \infer{@{text "\<Gamma> \<turnstile> B[?\<alpha>]"}}{@{text "\<Gamma> \<turnstile> B[\<alpha>]"} & @{text "\<alpha> \<notin> \<Gamma>"}}
 | 
| 526 | \quad | |
| 527 |   \infer[\quad@{text "(generalize)"}]{@{text "\<Gamma> \<turnstile> B[?x]"}}{@{text "\<Gamma> \<turnstile> B[x]"} & @{text "x \<notin> \<Gamma>"}}
 | |
| 20498 | 528 | \] | 
| 529 | \[ | |
| 20501 | 530 |   \infer{@{text "\<Gamma> \<turnstile> B[\<tau>]"}}{@{text "\<Gamma> \<turnstile> B[?\<alpha>]"}}
 | 
| 531 | \quad | |
| 532 |   \infer[\quad@{text "(instantiate)"}]{@{text "\<Gamma> \<turnstile> B[t]"}}{@{text "\<Gamma> \<turnstile> B[?x]"}}
 | |
| 20498 | 533 | \] | 
| 20501 | 534 |   \caption{Admissible substitution rules}\label{fig:subst-rules}
 | 
| 535 |   \end{center}
 | |
| 536 |   \end{figure}
 | |
| 18537 | 537 | |
| 20537 | 538 |   Note that @{text "instantiate"} does not require an explicit
 | 
| 539 |   side-condition, because @{text "\<Gamma>"} may never contain schematic
 | |
| 540 | variables. | |
| 541 | ||
| 542 | In principle, variables could be substituted in hypotheses as well, | |
| 20543 | 543 | but this would disrupt the monotonicity of reasoning: deriving | 
| 544 |   @{text "\<Gamma>\<vartheta> \<turnstile> B\<vartheta>"} from @{text "\<Gamma> \<turnstile> B"} is
 | |
| 545 |   correct, but @{text "\<Gamma>\<vartheta> \<supseteq> \<Gamma>"} does not necessarily hold:
 | |
| 546 | the result belongs to a different proof context. | |
| 20542 | 547 | |
| 20543 | 548 |   \medskip An \emph{oracle} is a function that produces axioms on the
 | 
| 549 |   fly.  Logically, this is an instance of the @{text "axiom"} rule
 | |
| 550 |   (\figref{fig:prim-rules}), but there is an operational difference.
 | |
| 551 | The system always records oracle invocations within derivations of | |
| 552 | theorems. Tracing plain axioms (and named theorems) is optional. | |
| 20542 | 553 | |
| 554 | Axiomatizations should be limited to the bare minimum, typically as | |
| 555 | part of the initial logical basis of an object-logic formalization. | |
| 20543 | 556 | Later on, theories are usually developed in a strictly definitional | 
| 557 | fashion, by stating only certain equalities over new constants. | |
| 20542 | 558 | |
| 559 |   A \emph{simple definition} consists of a constant declaration @{text
 | |
| 20543 | 560 |   "c :: \<sigma>"} together with an axiom @{text "\<turnstile> c \<equiv> t"}, where @{text "t
 | 
| 561 | :: \<sigma>"} is a closed term without any hidden polymorphism. The RHS | |
| 562 |   may depend on further defined constants, but not @{text "c"} itself.
 | |
| 563 |   Definitions of functions may be presented as @{text "c \<^vec>x \<equiv>
 | |
| 564 |   t"} instead of the puristic @{text "c \<equiv> \<lambda>\<^vec>x. t"}.
 | |
| 20542 | 565 | |
| 20543 | 566 |   An \emph{overloaded definition} consists of a collection of axioms
 | 
| 567 |   for the same constant, with zero or one equations @{text
 | |
| 568 |   "c((\<^vec>\<alpha>)\<kappa>) \<equiv> t"} for each type constructor @{text "\<kappa>"} (for
 | |
| 569 |   distinct variables @{text "\<^vec>\<alpha>"}).  The RHS may mention
 | |
| 570 |   previously defined constants as above, or arbitrary constants @{text
 | |
| 571 |   "d(\<alpha>\<^isub>i)"} for some @{text "\<alpha>\<^isub>i"} projected from @{text
 | |
| 572 | "\<^vec>\<alpha>"}. Thus overloaded definitions essentially work by | |
| 573 | primitive recursion over the syntactic structure of a single type | |
| 574 | argument. | |
| 20521 | 575 | *} | 
| 20498 | 576 | |
| 20521 | 577 | text %mlref {*
 | 
| 578 |   \begin{mldecls}
 | |
| 579 |   @{index_ML_type ctyp} \\
 | |
| 580 |   @{index_ML_type cterm} \\
 | |
| 20547 | 581 |   @{index_ML Thm.ctyp_of: "theory -> typ -> ctyp"} \\
 | 
| 582 |   @{index_ML Thm.cterm_of: "theory -> term -> cterm"} \\
 | |
| 583 |   \end{mldecls}
 | |
| 584 |   \begin{mldecls}
 | |
| 20521 | 585 |   @{index_ML_type thm} \\
 | 
| 20542 | 586 |   @{index_ML proofs: "int ref"} \\
 | 
| 587 |   @{index_ML Thm.assume: "cterm -> thm"} \\
 | |
| 588 |   @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
 | |
| 589 |   @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
 | |
| 590 |   @{index_ML Thm.implies_intr: "cterm -> thm -> thm"} \\
 | |
| 591 |   @{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\
 | |
| 592 |   @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\
 | |
| 593 |   @{index_ML Thm.instantiate: "(ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm"} \\
 | |
| 28674 | 594 |   @{index_ML Thm.axiom: "theory -> string -> thm"} \\
 | 
| 28290 | 595 |   @{index_ML Thm.add_oracle: "bstring * ('a -> cterm) -> theory
 | 
| 596 |   -> (string * ('a -> thm)) * theory"} \\
 | |
| 20547 | 597 |   \end{mldecls}
 | 
| 598 |   \begin{mldecls}
 | |
| 20542 | 599 |   @{index_ML Theory.add_axioms_i: "(string * term) list -> theory -> theory"} \\
 | 
| 600 |   @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\
 | |
| 601 |   @{index_ML Theory.add_defs_i: "bool -> bool -> (bstring * term) list -> theory -> theory"} \\
 | |
| 20521 | 602 |   \end{mldecls}
 | 
| 603 | ||
| 604 |   \begin{description}
 | |
| 605 | ||
| 20542 | 606 |   \item @{ML_type ctyp} and @{ML_type cterm} represent certified types
 | 
| 607 | and terms, respectively. These are abstract datatypes that | |
| 608 | guarantee that its values have passed the full well-formedness (and | |
| 609 | well-typedness) checks, relative to the declarations of type | |
| 610 | constructors, constants etc. in the theory. | |
| 611 | ||
| 20547 | 612 |   \item @{ML ctyp_of}~@{text "thy \<tau>"} and @{ML cterm_of}~@{text "thy
 | 
| 613 | t"} explicitly checks types and terms, respectively. This also | |
| 614 | involves some basic normalizations, such expansion of type and term | |
| 615 | abbreviations from the theory context. | |
| 616 | ||
| 617 | Re-certification is relatively slow and should be avoided in tight | |
| 618 | reasoning loops. There are separate operations to decompose | |
| 619 | certified entities (including actual theorems). | |
| 20542 | 620 | |
| 621 |   \item @{ML_type thm} represents proven propositions.  This is an
 | |
| 622 | abstract datatype that guarantees that its values have been | |
| 623 |   constructed by basic principles of the @{ML_struct Thm} module.
 | |
| 20543 | 624 |   Every @{ML thm} value contains a sliding back-reference to the
 | 
| 625 |   enclosing theory, cf.\ \secref{sec:context-theory}.
 | |
| 20542 | 626 | |
| 20543 | 627 |   \item @{ML proofs} determines the detail of proof recording within
 | 
| 628 |   @{ML_type thm} values: @{ML 0} records only oracles, @{ML 1} records
 | |
| 629 |   oracles, axioms and named theorems, @{ML 2} records full proof
 | |
| 630 | terms. | |
| 20542 | 631 | |
| 632 |   \item @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML
 | |
| 633 |   Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim}
 | |
| 634 |   correspond to the primitive inferences of \figref{fig:prim-rules}.
 | |
| 635 | ||
| 636 |   \item @{ML Thm.generalize}~@{text "(\<^vec>\<alpha>, \<^vec>x)"}
 | |
| 637 |   corresponds to the @{text "generalize"} rules of
 | |
| 20543 | 638 |   \figref{fig:subst-rules}.  Here collections of type and term
 | 
| 639 | variables are generalized simultaneously, specified by the given | |
| 640 | basic names. | |
| 20521 | 641 | |
| 20542 | 642 |   \item @{ML Thm.instantiate}~@{text "(\<^vec>\<alpha>\<^isub>s,
 | 
| 643 |   \<^vec>x\<^isub>\<tau>)"} corresponds to the @{text "instantiate"} rules
 | |
| 644 |   of \figref{fig:subst-rules}.  Type variables are substituted before
 | |
| 645 |   term variables.  Note that the types in @{text "\<^vec>x\<^isub>\<tau>"}
 | |
| 646 | refer to the instantiated versions. | |
| 647 | ||
| 28674 | 648 |   \item @{ML Thm.axiom}~@{text "thy name"} retrieves a named
 | 
| 20542 | 649 |   axiom, cf.\ @{text "axiom"} in \figref{fig:prim-rules}.
 | 
| 650 | ||
| 28290 | 651 |   \item @{ML Thm.add_oracle}~@{text "(name, oracle)"} produces a named
 | 
| 652 | oracle rule, essentially generating arbitrary axioms on the fly, | |
| 653 |   cf.\ @{text "axiom"} in \figref{fig:prim-rules}.
 | |
| 20521 | 654 | |
| 20543 | 655 |   \item @{ML Theory.add_axioms_i}~@{text "[(name, A), \<dots>]"} declares
 | 
| 656 | arbitrary propositions as axioms. | |
| 20542 | 657 | |
| 658 |   \item @{ML Theory.add_deps}~@{text "name c\<^isub>\<tau>
 | |
| 20543 | 659 | \<^vec>d\<^isub>\<sigma>"} declares dependencies of a named specification | 
| 660 |   for constant @{text "c\<^isub>\<tau>"}, relative to existing
 | |
| 661 |   specifications for constants @{text "\<^vec>d\<^isub>\<sigma>"}.
 | |
| 20542 | 662 | |
| 663 |   \item @{ML Theory.add_defs_i}~@{text "unchecked overloaded [(name, c
 | |
| 20543 | 664 | \<^vec>x \<equiv> t), \<dots>]"} states a definitional axiom for an existing | 
| 665 |   constant @{text "c"}.  Dependencies are recorded (cf.\ @{ML
 | |
| 666 |   Theory.add_deps}), unless the @{text "unchecked"} option is set.
 | |
| 20521 | 667 | |
| 668 |   \end{description}
 | |
| 669 | *} | |
| 670 | ||
| 671 | ||
| 20543 | 672 | subsection {* Auxiliary definitions *}
 | 
| 20521 | 673 | |
| 674 | text {*
 | |
| 20543 | 675 |   Theory @{text "Pure"} provides a few auxiliary definitions, see
 | 
| 676 |   \figref{fig:pure-aux}.  These special constants are normally not
 | |
| 677 | exposed to the user, but appear in internal encodings. | |
| 20501 | 678 | |
| 679 |   \begin{figure}[htb]
 | |
| 680 |   \begin{center}
 | |
| 20498 | 681 |   \begin{tabular}{ll}
 | 
| 20521 | 682 |   @{text "conjunction :: prop \<Rightarrow> prop \<Rightarrow> prop"} & (infix @{text "&"}) \\
 | 
| 683 |   @{text "\<turnstile> A & B \<equiv> (\<And>C. (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C)"} \\[1ex]
 | |
| 20543 | 684 |   @{text "prop :: prop \<Rightarrow> prop"} & (prefix @{text "#"}, suppressed) \\
 | 
| 20521 | 685 |   @{text "#A \<equiv> A"} \\[1ex]
 | 
| 686 |   @{text "term :: \<alpha> \<Rightarrow> prop"} & (prefix @{text "TERM"}) \\
 | |
| 687 |   @{text "term x \<equiv> (\<And>A. A \<Longrightarrow> A)"} \\[1ex]
 | |
| 688 |   @{text "TYPE :: \<alpha> itself"} & (prefix @{text "TYPE"}) \\
 | |
| 689 |   @{text "(unspecified)"} \\
 | |
| 20498 | 690 |   \end{tabular}
 | 
| 20521 | 691 |   \caption{Definitions of auxiliary connectives}\label{fig:pure-aux}
 | 
| 20501 | 692 |   \end{center}
 | 
| 693 |   \end{figure}
 | |
| 694 | ||
| 20537 | 695 |   Derived conjunction rules include introduction @{text "A \<Longrightarrow> B \<Longrightarrow> A &
 | 
| 696 |   B"}, and destructions @{text "A & B \<Longrightarrow> A"} and @{text "A & B \<Longrightarrow> B"}.
 | |
| 697 | Conjunction allows to treat simultaneous assumptions and conclusions | |
| 698 | uniformly. For example, multiple claims are intermediately | |
| 20543 | 699 | represented as explicit conjunction, but this is refined into | 
| 700 | separate sub-goals before the user continues the proof; the final | |
| 701 | result is projected into a list of theorems (cf.\ | |
| 20537 | 702 |   \secref{sec:tactical-goals}).
 | 
| 20498 | 703 | |
| 20537 | 704 |   The @{text "prop"} marker (@{text "#"}) makes arbitrarily complex
 | 
| 705 |   propositions appear as atomic, without changing the meaning: @{text
 | |
| 706 |   "\<Gamma> \<turnstile> A"} and @{text "\<Gamma> \<turnstile> #A"} are interchangeable.  See
 | |
| 707 |   \secref{sec:tactical-goals} for specific operations.
 | |
| 20521 | 708 | |
| 20543 | 709 |   The @{text "term"} marker turns any well-typed term into a derivable
 | 
| 710 |   proposition: @{text "\<turnstile> TERM t"} holds unconditionally.  Although
 | |
| 711 | this is logically vacuous, it allows to treat terms and proofs | |
| 712 | uniformly, similar to a type-theoretic framework. | |
| 20498 | 713 | |
| 20537 | 714 |   The @{text "TYPE"} constructor is the canonical representative of
 | 
| 715 |   the unspecified type @{text "\<alpha> itself"}; it essentially injects the
 | |
| 716 | language of types into that of terms. There is specific notation | |
| 717 |   @{text "TYPE(\<tau>)"} for @{text "TYPE\<^bsub>\<tau>
 | |
| 20521 | 718 | itself\<^esub>"}. | 
| 20537 | 719 |   Although being devoid of any particular meaning, the @{text
 | 
| 720 |   "TYPE(\<tau>)"} accounts for the type @{text "\<tau>"} within the term
 | |
| 721 |   language.  In particular, @{text "TYPE(\<alpha>)"} may be used as formal
 | |
| 722 | argument in primitive definitions, in order to circumvent hidden | |
| 723 |   polymorphism (cf.\ \secref{sec:terms}).  For example, @{text "c
 | |
| 724 |   TYPE(\<alpha>) \<equiv> A[\<alpha>]"} defines @{text "c :: \<alpha> itself \<Rightarrow> prop"} in terms of
 | |
| 725 |   a proposition @{text "A"} that depends on an additional type
 | |
| 726 | argument, which is essentially a predicate on types. | |
| 20521 | 727 | *} | 
| 20501 | 728 | |
| 20521 | 729 | text %mlref {*
 | 
| 730 |   \begin{mldecls}
 | |
| 731 |   @{index_ML Conjunction.intr: "thm -> thm -> thm"} \\
 | |
| 732 |   @{index_ML Conjunction.elim: "thm -> thm * thm"} \\
 | |
| 733 |   @{index_ML Drule.mk_term: "cterm -> thm"} \\
 | |
| 734 |   @{index_ML Drule.dest_term: "thm -> cterm"} \\
 | |
| 735 |   @{index_ML Logic.mk_type: "typ -> term"} \\
 | |
| 736 |   @{index_ML Logic.dest_type: "term -> typ"} \\
 | |
| 737 |   \end{mldecls}
 | |
| 738 | ||
| 739 |   \begin{description}
 | |
| 740 | ||
| 20542 | 741 |   \item @{ML Conjunction.intr} derives @{text "A & B"} from @{text
 | 
| 742 |   "A"} and @{text "B"}.
 | |
| 743 | ||
| 20543 | 744 |   \item @{ML Conjunction.elim} derives @{text "A"} and @{text "B"}
 | 
| 20542 | 745 |   from @{text "A & B"}.
 | 
| 746 | ||
| 20543 | 747 |   \item @{ML Drule.mk_term} derives @{text "TERM t"}.
 | 
| 20542 | 748 | |
| 20543 | 749 |   \item @{ML Drule.dest_term} recovers term @{text "t"} from @{text
 | 
| 750 | "TERM t"}. | |
| 20542 | 751 | |
| 752 |   \item @{ML Logic.mk_type}~@{text "\<tau>"} produces the term @{text
 | |
| 753 | "TYPE(\<tau>)"}. | |
| 754 | ||
| 755 |   \item @{ML Logic.dest_type}~@{text "TYPE(\<tau>)"} recovers the type
 | |
| 756 |   @{text "\<tau>"}.
 | |
| 20521 | 757 | |
| 758 |   \end{description}
 | |
| 20491 | 759 | *} | 
| 18537 | 760 | |
| 20480 | 761 | |
| 28784 | 762 | section {* Object-level rules \label{sec:obj-rules} *}
 | 
| 18537 | 763 | |
| 20929 | 764 | text %FIXME {*
 | 
| 18537 | 765 | |
| 766 | FIXME | |
| 767 | ||
| 20491 | 768 |   A \emph{rule} is any Pure theorem in HHF normal form; there is a
 | 
| 769 | separate calculus for rule composition, which is modeled after | |
| 770 |   Gentzen's Natural Deduction \cite{Gentzen:1935}, but allows
 | |
| 771 |   rules to be nested arbitrarily, similar to \cite{extensions91}.
 | |
| 772 | ||
| 773 | Normally, all theorems accessible to the user are proper rules. | |
| 774 | Low-level inferences are occasional required internally, but the | |
| 775 | result should be always presented in canonical form. The higher | |
| 776 | interfaces of Isabelle/Isar will always produce proper rules. It is | |
| 777 | important to maintain this invariant in add-on applications! | |
| 778 | ||
| 779 |   There are two main principles of rule composition: @{text
 | |
| 780 |   "resolution"} (i.e.\ backchaining of rules) and @{text
 | |
| 781 | "by-assumption"} (i.e.\ closing a branch); both principles are | |
| 20519 | 782 |   combined in the variants of @{text "elim-resolution"} and @{text
 | 
| 20491 | 783 |   "dest-resolution"}.  Raw @{text "composition"} is occasionally
 | 
| 784 | useful as well, also it is strictly speaking outside of the proper | |
| 785 | rule calculus. | |
| 786 | ||
| 787 | Rules are treated modulo general higher-order unification, which is | |
| 788 |   unification modulo the equational theory of @{text "\<alpha>\<beta>\<eta>"}-conversion
 | |
| 789 |   on @{text "\<lambda>"}-terms.  Moreover, propositions are understood modulo
 | |
| 790 |   the (derived) equivalence @{text "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}.
 | |
| 791 | ||
| 792 | This means that any operations within the rule calculus may be | |
| 793 |   subject to spontaneous @{text "\<alpha>\<beta>\<eta>"}-HHF conversions.  It is common
 | |
| 794 | practice not to contract or expand unnecessarily. Some mechanisms | |
| 795 | prefer an one form, others the opposite, so there is a potential | |
| 796 | danger to produce some oscillation! | |
| 797 | ||
| 798 |   Only few operations really work \emph{modulo} HHF conversion, but
 | |
| 799 |   expect a normal form: quantifiers @{text "\<And>"} before implications
 | |
| 800 |   @{text "\<Longrightarrow>"} at each level of nesting.
 | |
| 801 | ||
| 18537 | 802 | \glossary{Hereditary Harrop Formula}{The set of propositions in HHF
 | 
| 803 | format is defined inductively as @{text "H = (\<And>x\<^sup>*. H\<^sup>* \<Longrightarrow>
 | |
| 804 | A)"}, for variables @{text "x"} and atomic propositions @{text "A"}.
 | |
| 805 | Any proposition may be put into HHF form by normalizing with the rule | |
| 806 | @{text "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}.  In Isabelle, the outermost
 | |
| 807 | quantifier prefix is represented via \seeglossary{schematic
 | |
| 808 | variables}, such that the top-level structure is merely that of a | |
| 809 | \seeglossary{Horn Clause}}.
 | |
| 810 | ||
| 811 | \glossary{HHF}{See \seeglossary{Hereditary Harrop Formula}.}
 | |
| 812 | ||
| 20498 | 813 | |
| 814 | \[ | |
| 815 |   \infer[@{text "(assumption)"}]{@{text "C\<vartheta>"}}
 | |
| 816 |   {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C"} & @{text "A\<vartheta> = H\<^sub>i\<vartheta>"}~~\text{(for some~@{text i})}}
 | |
| 817 | \] | |
| 818 | ||
| 819 | ||
| 820 | \[ | |
| 821 |   \infer[@{text "(compose)"}]{@{text "\<^vec>A\<vartheta> \<Longrightarrow> C\<vartheta>"}}
 | |
| 822 |   {@{text "\<^vec>A \<Longrightarrow> B"} & @{text "B' \<Longrightarrow> C"} & @{text "B\<vartheta> = B'\<vartheta>"}}
 | |
| 823 | \] | |
| 824 | ||
| 825 | ||
| 826 | \[ | |
| 827 |   \infer[@{text "(\<And>_lift)"}]{@{text "(\<And>\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \<Longrightarrow> (\<And>\<^vec>x. B (?\<^vec>a \<^vec>x))"}}{@{text "\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a"}}
 | |
| 828 | \] | |
| 829 | \[ | |
| 830 |   \infer[@{text "(\<Longrightarrow>_lift)"}]{@{text "(\<^vec>H \<Longrightarrow> \<^vec>A) \<Longrightarrow> (\<^vec>H \<Longrightarrow> B)"}}{@{text "\<^vec>A \<Longrightarrow> B"}}
 | |
| 831 | \] | |
| 832 | ||
| 833 |   The @{text resolve} scheme is now acquired from @{text "\<And>_lift"},
 | |
| 834 |   @{text "\<Longrightarrow>_lift"}, and @{text compose}.
 | |
| 835 | ||
| 836 | \[ | |
| 837 |   \infer[@{text "(resolution)"}]
 | |
| 838 |   {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (?\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
 | |
| 839 |   {\begin{tabular}{l}
 | |
| 840 |     @{text "\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a"} \\
 | |
| 841 |     @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
 | |
| 842 |     @{text "(\<lambda>\<^vec>x. B (?\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
 | |
| 843 |    \end{tabular}}
 | |
| 844 | \] | |
| 845 | ||
| 846 | ||
| 847 |   FIXME @{text "elim_resolution"}, @{text "dest_resolution"}
 | |
| 18537 | 848 | *} | 
| 849 | ||
| 20498 | 850 | |
| 18537 | 851 | end |