diff -r 3dd63180cdbf -r 1e57c3c4e05c src/Doc/IsarImplementation/Logic.thy --- a/src/Doc/IsarImplementation/Logic.thy Fri Jun 21 16:21:33 2013 -0700 +++ b/src/Doc/IsarImplementation/Logic.thy Thu Jun 13 17:40:58 2013 +0200 @@ -932,6 +932,72 @@ *} +subsection {* Sort hypotheses *} + +text {* Type variables are decorated with sorts, as explained in + \secref{sec:types}. This constrains type instantiation to certain + ranges of types: variable @{text "\\<^sub>s"} may only be assigned to types + @{text "\"} that belong to sort @{text "s"}. Within the logic, sort + constraints act like implicit preconditions on the result @{text + "\\\<^sub>1 : s\<^sub>1\, \, \\\<^sub>n : s\<^sub>n\, \ \ \"} where the type variables @{text + "\\<^sub>1, \, \\<^sub>n"} cover the propositions @{text "\"}, @{text "\"}, as + well as the proof of @{text "\ \ \"}. + + These \emph{sort hypothesis} of a theorem are passed monotonically + through further derivations. They are redundant, as long as the + statement of a theorem still contains the type variables that are + accounted here. The logical significance of sort hypotheses is + limited to the boundary case where type variables disappear from the + proposition, e.g.\ @{text "\\\<^sub>s : s\ \ \"}. Since such dangling type + variables can be renamed arbitrarily without changing the + proposition @{text "\"}, the inference kernel maintains sort + hypotheses in anonymous form @{text "s \ \"}. + + In most practical situations, such extra sort hypotheses may be + stripped in a final bookkeeping step, e.g.\ at the end of a proof: + they are typically left over from intermediate reasoning with type + classes that can be satisfied by some concrete type @{text "\"} of + sort @{text "s"} to replace the hypothetical type variable @{text + "\\<^sub>s"}. *} + +text %mlref {* + \begin{mldecls} + @{index_ML Thm.extra_shyps: "thm -> sort list"} \\ + @{index_ML Thm.strip_shyps: "thm -> thm"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML Thm.extra_shyps}~@{text "thm"} determines the extraneous + sort hypotheses of the given theorem, i.e.\ the sorts that are not + present within type variables of the statement. + + \item @{ML Thm.strip_shyps}~@{text "thm"} removes any extraneous + sort hypotheses that can be witnessed from the type signature. + + \end{description} +*} + +text %mlex {* The following artificial example demonstrates the + derivation of @{prop False} with a pending sort hypothesis involving + a logically empty sort. *} + +class empty = + assumes bad: "\(x::'a) y. x \ y" + +theorem (in empty) false: False + using bad by blast + +ML {* + @{assert} (Thm.extra_shyps @{thm false} = [@{sort empty}]) +*} + +text {* Thanks to the inference kernel managing sort hypothesis + according to their logical significance, this example is merely an + instance of \emph{ex falso quodlibet consequitur} --- not a collapse + of the logical framework! *} + + section {* Object-level rules \label{sec:obj-rules} *} text {*