1.1 --- a/src/Pure/goal.ML Sat Apr 12 17:00:38 2008 +0200
1.2 +++ b/src/Pure/goal.ML Sat Apr 12 17:00:40 2008 +0200
1.3 @@ -20,7 +20,6 @@
1.4 val conclude: thm -> thm
1.5 val finish: thm -> thm
1.6 val norm_result: thm -> thm
1.7 - val close_result: thm -> thm
1.8 val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
1.9 val prove_multi: Proof.context -> string list -> term list -> term list ->
1.10 ({prems: thm list, context: Proof.context} -> tactic) -> thm list
1.11 @@ -84,18 +83,12 @@
1.12
1.13 (** results **)
1.14
1.15 -(* normal form *)
1.16 -
1.17 val norm_result =
1.18 Drule.flexflex_unique
1.19 #> MetaSimplifier.norm_hhf_protect
1.20 #> Thm.strip_shyps
1.21 #> Drule.zero_var_indexes;
1.22
1.23 -val close_result =
1.24 - Thm.compress
1.25 - #> Drule.close_derivation;
1.26 -
1.27
1.28
1.29 (** tactical theorem proving **)