src/HOL/Inductive.thy
changeset 81505 01f2936ec85e
parent 80934 8e72f55295fd
child 81506 f76a5849b570
--- a/src/HOL/Inductive.thy	Thu Nov 28 19:35:30 2024 +0100
+++ b/src/HOL/Inductive.thy	Fri Nov 29 00:25:03 2024 +0100
@@ -532,7 +532,7 @@
   let
     fun fun_tr ctxt [cs] =
       let
-        val x = Syntax.free (fst (Name.variant "x" (Term.declare_term_frees cs Name.context)));
+        val x = Syntax.free (fst (Name.variant "x" (Term.declare_free_names cs Name.context)));
         val ft = Case_Translation.case_tr true ctxt [x, cs];
       in lambda x ft end
   in [(\<^syntax_const>\<open>_lam_pats_syntax\<close>, fun_tr)] end