diff -r 0e6adce08fb0 -r d8f5d3391766 src/Provers/Arith/cancel_numeral_factor.ML --- a/src/Provers/Arith/cancel_numeral_factor.ML Thu Aug 08 23:48:31 2002 +0200 +++ b/src/Provers/Arith/cancel_numeral_factor.ML Thu Aug 08 23:49:44 2002 +0200 @@ -30,7 +30,7 @@ as with < and <= but not = and div*) (*proof tools*) val prove_conv: tactic list -> Sign.sg -> - thm list -> term * term -> thm option + thm list -> string list -> term * term -> thm option val trans_tac: thm option -> tactic (*applies the initial lemma*) val norm_tac: tactic (*proves the initial lemma*) val numeral_simp_tac: tactic (*proves the final theorem*) @@ -53,10 +53,7 @@ (*the simplification procedure*) fun proc sg hyps t = let (*first freeze any Vars in the term to prevent flex-flex problems*) - val rand_s = gensym"_" - fun mk_inst (var as Var((a,i),T)) = - (var, Free((a ^ rand_s ^ string_of_int i), T)) - val t' = subst_atomic (map mk_inst (term_vars t)) t + val (t', xs) = Term.adhoc_freeze_vars t; val (t1,t2) = Data.dest_bal t' val (m1, t1') = Data.dest_coeff t1 and (m2, t2') = Data.dest_coeff t2 @@ -73,13 +70,13 @@ else Data.mk_bal (Data.mk_coeff(n1,t1'), Data.mk_coeff(n2,t2')) val reshape = (*Move d to the front and put the rest into standard form i * #m * j == #d * (#n * (j * k)) *) - Data.prove_conv [Data.norm_tac] sg hyps + Data.prove_conv [Data.norm_tac] sg hyps xs (t', Data.mk_bal (newshape(n1,t1'), newshape(n2,t2'))) in apsome Data.simplify_meta_eq (Data.prove_conv [Data.trans_tac reshape, rtac Data.cancel 1, - Data.numeral_simp_tac] sg hyps (t', rhs)) + Data.numeral_simp_tac] sg hyps xs (t', rhs)) end handle TERM _ => None | TYPE _ => None; (*Typically (if thy doesn't include Numeral)