# HG changeset patch # User wenzelm # Date 1491757435 -7200 # Node ID ed18feb34c071f0626a106473925b94f697b62c6 # Parent e9e7f5f5794ca1f84a9e84d7d15588efc667dfce tuned signature -- prefer qualified names; diff -r e9e7f5f5794c -r ed18feb34c07 src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy Sat Apr 08 22:36:32 2017 +0200 +++ b/src/Doc/Implementation/Logic.thy Sun Apr 09 19:03:55 2017 +0200 @@ -70,7 +70,7 @@ For \k = 1\ the parentheses are omitted, e.g.\ \\ list\ instead of \(\)list\. Further notation is provided for specific constructors, notably the right-associative infix \\ \ \\ instead of \(\, \)fun\. - + The logical category \<^emph>\type\ is defined inductively over type variables and type constructors as follows: \\ = \\<^sub>s | ?\\<^sub>s | (\\<^sub>1, \, \\<^sub>k)\\. @@ -1186,7 +1186,7 @@ \begin{mldecls} @{index_ML_type proof} \\ @{index_ML_type proof_body} \\ - @{index_ML proofs: "int Unsynchronized.ref"} \\ + @{index_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{index_ML Reconstruct.reconstruct_proof: "Proof.context -> term -> proof -> proof"} \\ @{index_ML Reconstruct.expand_proof: "Proof.context -> @@ -1215,11 +1215,11 @@ construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. - \<^descr> @{ML proofs} specifies the detail of proof recording within @{ML_type thm} - values produced by the inference kernel: @{ML 0} records only the names of - oracles, @{ML 1} records oracle names and propositions, @{ML 2} additionally - records full proof terms. Officially named theorems that contribute to a - result are recorded in any case. + \<^descr> @{ML Proofterm.proofs} specifies the detail of proof recording within + @{ML_type thm} values produced by the inference kernel: @{ML 0} records only + the names of oracles, @{ML 1} records oracle names and propositions, @{ML 2} + additionally records full proof terms. Officially named theorems that + contribute to a result are recorded in any case. \<^descr> @{ML Reconstruct.reconstruct_proof}~\ctxt prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. diff -r e9e7f5f5794c -r ed18feb34c07 src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Sat Apr 08 22:36:32 2017 +0200 +++ b/src/Pure/proofterm.ML Sun Apr 09 19:03:55 2017 +0200 @@ -8,8 +8,6 @@ signature BASIC_PROOFTERM = sig - val proofs: int Unsynchronized.ref - type thm_node datatype proof = MinProof @@ -28,14 +26,13 @@ {oracles: (string * term) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} - val %> : proof * term -> proof end; signature PROOFTERM = sig include BASIC_PROOFTERM - + val proofs: int Unsynchronized.ref type pthm = serial * thm_node type oracle = string * term val proof_of: proof_body -> proof