src/HOL/Tools/Function/size.ML
changeset 36610 bafd82950e24
parent 35756 cfde251d03a5
child 38348 cf7b2121ad9d
equal deleted inserted replaced
36609:6ed6112f0a95 36610:bafd82950e24
   131       let
   131       let
   132         val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq;
   132         val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq;
   133         val (thm, lthy') = lthy
   133         val (thm, lthy') = lthy
   134           |> Local_Theory.define ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs))
   134           |> Local_Theory.define ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs))
   135           |-> (fn (t, (_, thm)) => Spec_Rules.add Spec_Rules.Equational ([t], [thm]) #> pair thm);
   135           |-> (fn (t, (_, thm)) => Spec_Rules.add Spec_Rules.Equational ([t], [thm]) #> pair thm);
   136         val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy');
   136         val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy');
   137         val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
   137         val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
   138       in (thm', lthy') end;
   138       in (thm', lthy') end;
   139 
   139 
   140     val ((size_def_thms, size_def_thms'), thy') =
   140     val ((size_def_thms, size_def_thms'), thy') =
   141       thy
   141       thy
   150       ||>> fold_map define_overloaded
   150       ||>> fold_map define_overloaded
   151         (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1))
   151         (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1))
   152       ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
   152       ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
   153       ||> Local_Theory.exit_global;
   153       ||> Local_Theory.exit_global;
   154 
   154 
   155     val ctxt = ProofContext.init thy';
   155     val ctxt = ProofContext.init_global thy';
   156 
   156 
   157     val simpset1 = HOL_basic_ss addsimps @{thm Nat.add_0} :: @{thm Nat.add_0_right} ::
   157     val simpset1 = HOL_basic_ss addsimps @{thm Nat.add_0} :: @{thm Nat.add_0_right} ::
   158       size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites;
   158       size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites;
   159     val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2);
   159     val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2);
   160 
   160