--- a/src/CCL/Wfd.thy Wed Jan 22 21:35:05 2025 +0100
+++ b/src/CCL/Wfd.thy Wed Jan 22 22:22:19 2025 +0100
@@ -424,9 +424,9 @@
| get_bno l n (Bound m) = (m-length(l),n)
(* Not a great way of identifying induction hypothesis! *)
-fun could_IH x = Term.could_unify(x,hd (Thm.prems_of @{thm rcallT})) orelse
- Term.could_unify(x,hd (Thm.prems_of @{thm rcall2T})) orelse
- Term.could_unify(x,hd (Thm.prems_of @{thm rcall3T}))
+fun could_IH x = Term.could_unify(x,hd (Thm.take_prems_of 1 @{thm rcallT})) orelse
+ Term.could_unify(x,hd (Thm.take_prems_of 1 @{thm rcall2T})) orelse
+ Term.could_unify(x,hd (Thm.take_prems_of 1 @{thm rcall3T}))
fun IHinst tac rls = SUBGOAL (fn (Bi,i) =>
let val bvs = bvars Bi []