src/Pure/Isar/code.ML
changeset 32662 2faf1148c062
parent 32661 f4d179d91af2
child 32738 15bb09ca0378
equal deleted inserted replaced
32661:f4d179d91af2 32662:2faf1148c062
   779 fun change_yield thy = Code.change_yield_data data_op thy;
   779 fun change_yield thy = Code.change_yield_data data_op thy;
   780 
   780 
   781 end;
   781 end;
   782 
   782 
   783 (** datastructure to log definitions for preprocessing of the predicate compiler **)
   783 (** datastructure to log definitions for preprocessing of the predicate compiler **)
   784 (*
   784 (* obviously a clone of Named_Thms *)
   785 structure Predicate_Compile_Preproc_Const_Defs = Named_Thms
   785 
       
   786 signature PREDICATE_COMPILE_PREPROC_CONST_DEFS =
       
   787 sig
       
   788   val get: Proof.context -> thm list
       
   789   val add_thm: thm -> Context.generic -> Context.generic
       
   790   val del_thm: thm -> Context.generic -> Context.generic
       
   791   
       
   792   val add_attribute : attribute
       
   793   val del_attribute : attribute
       
   794   
       
   795   val add_attrib : Attrib.src
       
   796   
       
   797   val setup: theory -> theory
       
   798 end;
       
   799 
       
   800 structure Predicate_Compile_Preproc_Const_Defs : PREDICATE_COMPILE_PREPROC_CONST_DEFS =
       
   801 struct
       
   802 
       
   803 structure Data = GenericDataFun
   786 (
   804 (
   787   val name = "predicate_compile_preproc_const_def"
   805   type T = thm list;
   788   val description =
   806   val empty = [];
   789     "definitions of constants as needed by the preprocessing for the predicate compiler"
   807   val extend = I;
   790 )
   808   fun merge _ = Thm.merge_thms;
   791 *)
   809 );
       
   810 
       
   811 val get = Data.get o Context.Proof;
       
   812 
       
   813 val add_thm = Data.map o Thm.add_thm;
       
   814 val del_thm = Data.map o Thm.del_thm;
       
   815 
       
   816 val add_attribute = Thm.declaration_attribute add_thm;
       
   817 val del_attribute = Thm.declaration_attribute del_thm;
       
   818 
       
   819 val add_attrib = Attrib.internal (K add_attribute)
       
   820 
       
   821 val setup =
       
   822   Attrib.setup (Binding.name "pred_compile_preproc") (Attrib.add_del add_attribute del_attribute)
       
   823     ("declaration of definition for preprocessing of the predicate compiler") #>
       
   824   PureThy.add_thms_dynamic (Binding.name "pred_compile_preproc", Data.get);
       
   825 
       
   826 end;
       
   827 
   792 structure Code : CODE = struct open Code; end;
   828 structure Code : CODE = struct open Code; end;