--- a/src/HOL/Tools/rewrite_hol_proof.ML Mon Nov 02 20:48:08 2009 +0100
+++ b/src/HOL/Tools/rewrite_hol_proof.ML Mon Nov 02 20:50:48 2009 +0100
@@ -15,8 +15,8 @@
open Proofterm;
-val rews = map (pairself (ProofSyntax.proof_of_term @{theory} true) o
- Logic.dest_equals o Logic.varify o ProofSyntax.read_term @{theory} propT)
+val rews = map (pairself (Proof_Syntax.proof_of_term @{theory} true) o
+ Logic.dest_equals o Logic.varify o Proof_Syntax.read_term @{theory} propT)
(** eliminate meta-equality rules **)
--- a/src/Pure/Isar/isar_cmd.ML Mon Nov 02 20:48:08 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML Mon Nov 02 20:50:48 2009 +0100
@@ -460,11 +460,11 @@
val prop = Thm.full_prop_of thm;
val prf' = Proofterm.rewrite_proof_notypes ([], []) prf;
in
- ProofSyntax.pretty_proof ctxt
+ Proof_Syntax.pretty_proof ctxt
(if full then Reconstruct.reconstruct_proof thy prop prf' else prf')
end
| SOME args => Pretty.chunks
- (map (ProofSyntax.pretty_proof_of (Proof.context_of state) full)
+ (map (Proof_Syntax.pretty_proof_of (Proof.context_of state) full)
(Proof.get_thmss state args)));
fun string_of_prop state s =
--- a/src/Pure/Proof/extraction.ML Mon Nov 02 20:48:08 2009 +0100
+++ b/src/Pure/Proof/extraction.ML Mon Nov 02 20:50:48 2009 +0100
@@ -300,7 +300,7 @@
val rtypes = map fst types;
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
val thy' = add_syntax thy;
- val rd = ProofSyntax.read_proof thy' false
+ val rd = Proof_Syntax.read_proof thy' false;
in fn (thm, (vs, s1, s2)) =>
let
val name = Thm.get_name thm;
--- a/src/Pure/Proof/proof_syntax.ML Mon Nov 02 20:48:08 2009 +0100
+++ b/src/Pure/Proof/proof_syntax.ML Mon Nov 02 20:50:48 2009 +0100
@@ -19,7 +19,7 @@
val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T
end;
-structure ProofSyntax : PROOF_SYNTAX =
+structure Proof_Syntax : PROOF_SYNTAX =
struct
open Proofterm;
--- a/src/Pure/Thy/thy_output.ML Mon Nov 02 20:48:08 2009 +0100
+++ b/src/Pure/Thy/thy_output.ML Mon Nov 02 20:50:48 2009 +0100
@@ -478,7 +478,7 @@
val ctxt' = Variable.auto_fixes eq ctxt;
in ProofContext.pretty_term_abbrev ctxt' eq end;
-fun pretty_prf full ctxt = ProofSyntax.pretty_proof_of ctxt full;
+fun pretty_prf full ctxt = Proof_Syntax.pretty_proof_of ctxt full;
fun pretty_theory ctxt name =
(Theory.requires (ProofContext.theory_of ctxt) name "presentation"; Pretty.str name);