--- a/src/HOL/Extraction.thy Sat Nov 15 11:25:17 2008 +0100
+++ b/src/HOL/Extraction.thy Sat Nov 15 21:31:13 2008 +0100
@@ -40,8 +40,7 @@
("set", ([realizes_set_proc], SOME mk_realizes_set))] #>
Extraction.set_preprocessor (fn thy =>
Proofterm.rewrite_proof_notypes
- ([], ("HOL/elim_cong", RewriteHOLProof.elim_cong) ::
- ProofRewriteRules.rprocs true) o
+ ([], RewriteHOLProof.elim_cong :: ProofRewriteRules.rprocs true) o
Proofterm.rewrite_proof thy
(RewriteHOLProof.rews, ProofRewriteRules.rprocs true) o
ProofRewriteRules.elim_vars (curry Const @{const_name default}))