src/Pure/goal.ML
changeset 26628 63306cb94313
parent 25301 24e027f55f45
child 26713 1c6181def1d0
--- a/src/Pure/goal.ML	Sat Apr 12 17:00:38 2008 +0200
+++ b/src/Pure/goal.ML	Sat Apr 12 17:00:40 2008 +0200
@@ -20,7 +20,6 @@
   val conclude: thm -> thm
   val finish: thm -> thm
   val norm_result: thm -> thm
-  val close_result: thm -> thm
   val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
   val prove_multi: Proof.context -> string list -> term list -> term list ->
     ({prems: thm list, context: Proof.context} -> tactic) -> thm list
@@ -84,18 +83,12 @@
 
 (** results **)
 
-(* normal form *)
-
 val norm_result =
   Drule.flexflex_unique
   #> MetaSimplifier.norm_hhf_protect
   #> Thm.strip_shyps
   #> Drule.zero_var_indexes;
 
-val close_result =
-  Thm.compress
-  #> Drule.close_derivation;
-
 
 
 (** tactical theorem proving **)