| author | wenzelm | 
| Sun, 09 Jun 2024 22:40:13 +0200 | |
| changeset 80314 | 594356f16810 | 
| parent 74561 | 8e6c973003c8 | 
| child 80636 | 4041e7c8059d | 
| permissions | -rw-r--r-- | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 1 | (* Title: HOL/Tools/Predicate_Compile/core_data.ML | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 3 | |
| 41941 | 4 | Data of the predicate compiler core. | 
| 5 | *) | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 6 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 7 | signature CORE_DATA = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 8 | sig | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 9 | type mode = Predicate_Compile_Aux.mode | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 10 | type compilation = Predicate_Compile_Aux.compilation | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 11 | type compilation_funs = Predicate_Compile_Aux.compilation_funs | 
| 59501 | 12 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 13 |   datatype predfun_data = PredfunData of {
 | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 14 | definition : thm, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 15 | intro : thm, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 16 | elim : thm, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 17 | neg_intro : thm option | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 18 | }; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 19 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 20 |   datatype pred_data = PredData of {
 | 
| 55543 | 21 | pos : Position.T, | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 22 | intros : (string option * thm) list, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 23 | elim : thm option, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 24 | preprocessed : bool, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 25 | function_names : (compilation * (mode * string) list) list, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 26 | predfun_data : (mode * predfun_data) list, | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 27 | needs_random : mode list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 28 | }; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 29 | |
| 58823 | 30 | structure PredData : THEORY_DATA (* FIXME keep data private *) | 
| 59501 | 31 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 32 | (* queries *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 33 | val defined_functions : compilation -> Proof.context -> string -> bool | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 34 | val is_registered : Proof.context -> string -> bool | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 35 | val function_name_of : compilation -> Proof.context -> string -> mode -> string | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 36 | val the_elim_of : Proof.context -> string -> thm | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 37 | val has_elim : Proof.context -> string -> bool | 
| 59501 | 38 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 39 | val needs_random : Proof.context -> string -> mode -> bool | 
| 59501 | 40 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 41 | val predfun_intro_of : Proof.context -> string -> mode -> thm | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 42 | val predfun_neg_intro_of : Proof.context -> string -> mode -> thm option | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 43 | val predfun_elim_of : Proof.context -> string -> mode -> thm | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 44 | val predfun_definition_of : Proof.context -> string -> mode -> thm | 
| 59501 | 45 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 46 | val all_preds_of : Proof.context -> string list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 47 | val modes_of: compilation -> Proof.context -> string -> mode list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 48 | val all_modes_of : compilation -> Proof.context -> (string * mode list) list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 49 | val all_random_modes_of : Proof.context -> (string * mode list) list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 50 | val intros_of : Proof.context -> string -> thm list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 51 | val names_of : Proof.context -> string -> string option list | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 52 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 53 | val intros_graph_of : Proof.context -> thm list Graph.T | 
| 59501 | 54 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 55 | (* updaters *) | 
| 59501 | 56 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 57 | val register_predicate : (string * thm list * thm) -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 58 | val register_intros : string * thm list -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 59 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 60 | (* FIXME: naming of function is strange *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 61 | val defined_function_of : compilation -> string -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 62 | val add_intro : string option * thm -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 63 | val set_elim : thm -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 64 | val set_function_name : compilation -> string -> mode -> string -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 65 | val add_predfun_data : string -> mode -> thm * ((thm * thm) * thm option) -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 66 | val set_needs_random : string -> mode list -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 67 | (* sophisticated updaters *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 68 | val extend_intro_graph : string list -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 69 | val preprocess_intros : string -> theory -> theory | 
| 59501 | 70 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 71 | (* alternative function definitions *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 72 | val register_alternative_function : string -> mode -> string -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 73 | val alternative_compilation_of_global : theory -> string -> mode -> | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 74 | (compilation_funs -> typ -> term) option | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 75 | val alternative_compilation_of : Proof.context -> string -> mode -> | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 76 | (compilation_funs -> typ -> term) option | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 77 | val functional_compilation : string -> mode -> compilation_funs -> typ -> term | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 78 | val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 79 | val force_modes_and_compilations : string -> | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 80 | (mode * ((compilation_funs -> typ -> term) * bool)) list -> theory -> theory | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 81 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 82 | end; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 83 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 84 | structure Core_Data : CORE_DATA = | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 85 | struct | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 86 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 87 | open Predicate_Compile_Aux; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 88 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 89 | (* book-keeping *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 90 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 91 | datatype predfun_data = PredfunData of {
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 92 | definition : thm, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 93 | intro : thm, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 94 | elim : thm, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 95 | neg_intro : thm option | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 96 | }; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 97 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 98 | fun rep_predfun_data (PredfunData data) = data; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 99 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 100 | fun mk_predfun_data (definition, ((intro, elim), neg_intro)) = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 101 |   PredfunData {definition = definition, intro = intro, elim = elim, neg_intro = neg_intro}
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 102 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 103 | datatype pred_data = PredData of {
 | 
| 55543 | 104 | pos: Position.T, | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 105 | intros : (string option * thm) list, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 106 | elim : thm option, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 107 | preprocessed : bool, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 108 | function_names : (compilation * (mode * string) list) list, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 109 | predfun_data : (mode * predfun_data) list, | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 110 | needs_random : mode list | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 111 | }; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 112 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 113 | fun rep_pred_data (PredData data) = data; | 
| 55543 | 114 | val pos_of = #pos o rep_pred_data; | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 115 | |
| 55543 | 116 | fun mk_pred_data | 
| 117 | (pos, (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random)))) = | |
| 118 |   PredData {pos = pos, intros = intros, elim = elim, preprocessed = preprocessed,
 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 119 | function_names = function_names, predfun_data = predfun_data, needs_random = needs_random} | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 120 | |
| 55543 | 121 | fun map_pred_data f | 
| 122 |     (PredData {pos, intros, elim, preprocessed, function_names, predfun_data, needs_random}) =
 | |
| 123 | mk_pred_data | |
| 124 | (f (pos, (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))))) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 125 | |
| 59501 | 126 | fun eq_pred_data (PredData d1, PredData d2) = | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 127 | eq_list (eq_pair (op =) Thm.eq_thm) (#intros d1, #intros d2) andalso | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 128 | eq_option Thm.eq_thm (#elim d1, #elim d2) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 129 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 130 | structure PredData = Theory_Data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 131 | ( | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 132 | type T = pred_data Graph.T; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 133 | val empty = Graph.empty; | 
| 55543 | 134 | val merge = | 
| 135 | Graph.join (fn key => fn (x, y) => | |
| 136 | if eq_pred_data (x, y) | |
| 137 | then raise Graph.SAME | |
| 138 | else | |
| 139 |         error ("Duplicate predicate declarations for " ^ quote key ^
 | |
| 140 | Position.here (pos_of x) ^ Position.here (pos_of y))); | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 141 | ); | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 142 | |
| 55437 | 143 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 144 | (* queries *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 145 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 146 | fun lookup_pred_data ctxt name = | 
| 42361 | 147 | Option.map rep_pred_data (try (Graph.get_node (PredData.get (Proof_Context.theory_of ctxt))) name) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 148 | |
| 55437 | 149 | fun the_pred_data ctxt name = | 
| 150 | (case lookup_pred_data ctxt name of | |
| 59501 | 151 |     NONE => error ("No such predicate: " ^ quote name)
 | 
| 55437 | 152 | | SOME data => data) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 153 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 154 | val is_registered = is_some oo lookup_pred_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 155 | |
| 42361 | 156 | val all_preds_of = Graph.keys o PredData.get o Proof_Context.theory_of | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 157 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 158 | val intros_of = map snd o #intros oo the_pred_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 159 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 160 | val names_of = map fst o #intros oo the_pred_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 161 | |
| 55437 | 162 | fun the_elim_of ctxt name = | 
| 163 | (case #elim (the_pred_data ctxt name) of | |
| 164 |     NONE => error ("No elimination rule for predicate " ^ quote name)
 | |
| 165 | | SOME thm => thm) | |
| 59501 | 166 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 167 | val has_elim = is_some o #elim oo the_pred_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 168 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 169 | fun function_names_of compilation ctxt name = | 
| 55437 | 170 | (case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation of | 
| 171 | NONE => | |
| 172 |       error ("No " ^ string_of_compilation compilation ^
 | |
| 173 | " functions defined for predicate " ^ quote name) | |
| 174 | | SOME fun_names => fun_names) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 175 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 176 | fun function_name_of compilation ctxt name mode = | 
| 55437 | 177 | (case AList.lookup eq_mode (function_names_of compilation ctxt name) mode of | 
| 178 | NONE => | |
| 179 |       error ("No " ^ string_of_compilation compilation ^
 | |
| 180 | " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name) | |
| 181 | | SOME function_name => function_name) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 182 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 183 | fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 184 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 185 | fun all_modes_of compilation ctxt = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 186 | map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 187 | (all_preds_of ctxt) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 188 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 189 | val all_random_modes_of = all_modes_of Random | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 190 | |
| 55437 | 191 | fun defined_functions compilation ctxt name = | 
| 192 | (case lookup_pred_data ctxt name of | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 193 | NONE => false | 
| 55437 | 194 | | SOME data => AList.defined (op =) (#function_names data) compilation) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 195 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 196 | fun needs_random ctxt s m = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 197 | member (op =) (#needs_random (the_pred_data ctxt s)) m | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 198 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 199 | fun lookup_predfun_data ctxt name mode = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 200 | Option.map rep_predfun_data | 
| 40140 
8282b87f957c
using mode_eq instead of op = for lookup in the predicate compiler
 bulwahn parents: 
40101diff
changeset | 201 | (AList.lookup eq_mode (#predfun_data (the_pred_data ctxt name)) mode) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 202 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 203 | fun the_predfun_data ctxt name mode = | 
| 55437 | 204 | (case lookup_predfun_data ctxt name mode of | 
| 205 | NONE => | |
| 206 |       error ("No function defined for mode " ^ string_of_mode mode ^
 | |
| 207 | " of predicate " ^ name) | |
| 208 | | SOME data => data) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 209 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 210 | val predfun_definition_of = #definition ooo the_predfun_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 211 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 212 | val predfun_intro_of = #intro ooo the_predfun_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 213 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 214 | val predfun_elim_of = #elim ooo the_predfun_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 215 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 216 | val predfun_neg_intro_of = #neg_intro ooo the_predfun_data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 217 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 218 | val intros_graph_of = | 
| 42361 | 219 | Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o Proof_Context.theory_of | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 220 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 221 | fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 222 | let | 
| 42361 | 223 | val thy = Proof_Context.theory_of ctxt | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 224 | val nargs = length (binder_types (fastype_of pred)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 225 |     fun meta_eq_of th = th RS @{thm eq_reflection}
 | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60752diff
changeset | 226 |     val tuple_rew_rules = map meta_eq_of [@{thm fst_conv}, @{thm snd_conv}, @{thm prod.inject}]
 | 
| 59501 | 227 | |
| 228 |     fun instantiate i n ({context = ctxt2, prems, ...}: Subgoal.focus) =
 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 229 | let | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59636diff
changeset | 230 | fun inst_pair_of (ix, (ty, t)) = ((ix, ty), t) | 
| 59501 | 231 | fun inst_of_matches tts = | 
| 232 | fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty) | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59636diff
changeset | 233 | |> snd |> Vartab.dest |> map (apsnd (Thm.cterm_of ctxt2) o inst_pair_of) | 
| 59501 | 234 | val (cases, (eqs, prems1)) = apsnd (chop (nargs - nparams)) (chop n prems) | 
| 52230 | 235 | val case_th = | 
| 59501 | 236 |           rewrite_rule ctxt2 (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1))
 | 
| 237 | val prems2 = maps (dest_conjunct_prem o rewrite_rule ctxt2 tuple_rew_rules) prems1 | |
| 55437 | 238 | val pats = | 
| 59501 | 239 | map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop) | 
| 240 | (take nargs (Thm.prems_of case_th)) | |
| 241 | val case_th' = | |
| 74282 | 242 | Thm.instantiate (TVars.empty, Vars.make (inst_of_matches pats)) case_th | 
| 59501 | 243 |             OF replicate nargs @{thm refl}
 | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 244 | val thesis = | 
| 74282 | 245 | Thm.instantiate (TVars.empty, | 
| 246 | Vars.make (inst_of_matches (Thm.prems_of case_th' ~~ map Thm.prop_of prems2))) | |
| 59582 | 247 | case_th' OF prems2 | 
| 60752 | 248 | in resolve_tac ctxt2 [thesis] 1 end | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 249 | in | 
| 59501 | 250 | Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule | 
| 251 |       (fn {context = ctxt1, ...} =>
 | |
| 252 | eresolve_tac ctxt1 [pre_cases_rule] 1 THEN (fn st => | |
| 253 | let val n = Thm.nprems_of st in | |
| 254 | st |> ALLGOALS (fn i => | |
| 255 |               rewrite_goal_tac ctxt1 @{thms split_paired_all} i THEN
 | |
| 256 | SUBPROOF (instantiate i n) ctxt1 i) | |
| 257 | end)) | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 258 | end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 259 | |
| 55437 | 260 | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 261 | (* updaters *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 262 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 263 | (* fetching introduction rules or registering introduction rules *) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 264 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 265 | val no_compilation = ([], ([], [])) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 266 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 267 | fun fetch_pred_data ctxt name = | 
| 65411 
3c628937899d
use Item_Net to store inductive info
 Lars Hupel <lars.hupel@mytum.de> parents: 
62391diff
changeset | 268 | (case try (Inductive.the_inductive_global ctxt) name of | 
| 59501 | 269 | SOME (info as (_, result)) => | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 270 | let | 
| 55543 | 271 | val thy = Proof_Context.theory_of ctxt | 
| 272 | ||
| 273 | val pos = Position.thread_data () | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 274 | fun is_intro_of intro = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 275 | let | 
| 59582 | 276 | val (const, _) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of intro)) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 277 | in (fst (dest_Const const) = name) end; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 278 | val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 279 | val index = find_index (fn s => s = name) (#names (fst info)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 280 | val pre_elim = nth (#elims result) index | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 281 | val pred = nth (#preds result) index | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 282 | val elim_t = mk_casesrule ctxt pred intros | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 283 | val nparams = length (Inductive.params_of (#raw_induct result)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 284 | val elim = prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 285 | in | 
| 55543 | 286 | mk_pred_data (pos, (((map (pair NONE) intros, SOME elim), true), no_compilation)) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 287 | end | 
| 55437 | 288 |   | NONE => error ("No such predicate: " ^ quote name))
 | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 289 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 290 | fun add_predfun_data name mode data = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 291 | let | 
| 55543 | 292 | val add = (apsnd o apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data)) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 293 | in PredData.map (Graph.map_node name (map_pred_data add)) end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 294 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 295 | fun is_inductive_predicate ctxt name = | 
| 65411 
3c628937899d
use Item_Net to store inductive info
 Lars Hupel <lars.hupel@mytum.de> parents: 
62391diff
changeset | 296 | is_some (try (Inductive.the_inductive_global ctxt) name) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 297 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 298 | fun depending_preds_of ctxt (key, value) = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 299 | let | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 300 | val intros = map (Thm.prop_of o snd) ((#intros o rep_pred_data) value) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 301 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 302 | fold Term.add_const_names intros [] | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 303 | |> (fn cs => | 
| 69593 | 304 | if member (op =) cs \<^const_name>\<open>HOL.eq\<close> then | 
| 305 | insert (op =) \<^const_name>\<open>Predicate.eq\<close> cs | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 306 | else cs) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 307 | |> filter (fn c => (not (c = key)) andalso | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 308 | (is_inductive_predicate ctxt c orelse is_registered ctxt c)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 309 | end; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 310 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 311 | fun add_intro (opt_case_name, thm) thy = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 312 | let | 
| 50056 | 313 | val (name, _) = dest_Const (fst (strip_intro_concl thm)) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 314 | fun cons_intro gr = | 
| 55437 | 315 | (case try (Graph.get_node gr) name of | 
| 316 | SOME _ => | |
| 317 | Graph.map_node name (map_pred_data | |
| 55543 | 318 | (apsnd (apfst (apfst (apfst (fn intros => intros @ [(opt_case_name, thm)])))))) gr | 
| 55437 | 319 | | NONE => | 
| 320 | Graph.new_node | |
| 55543 | 321 | (name, | 
| 322 | mk_pred_data (Position.thread_data (), | |
| 323 | (((([(opt_case_name, thm)], NONE), false), no_compilation)))) gr) | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 324 | in PredData.map cons_intro thy end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 325 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 326 | fun set_elim thm = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 327 | let | 
| 55437 | 328 | val (name, _) = | 
| 59582 | 329 | dest_Const (fst (strip_comb (HOLogic.dest_Trueprop (hd (Thm.prems_of thm))))) | 
| 55543 | 330 | in | 
| 331 | PredData.map (Graph.map_node name (map_pred_data (apsnd (apfst (apfst (apsnd (K (SOME thm)))))))) | |
| 332 | end | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 333 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 334 | fun register_predicate (constname, intros, elim) thy = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 335 | let | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 336 | val named_intros = map (pair NONE) intros | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 337 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 338 | if not (member (op =) (Graph.keys (PredData.get thy)) constname) then | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 339 | PredData.map | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 340 | (Graph.new_node (constname, | 
| 55543 | 341 | mk_pred_data (Position.thread_data (), | 
| 342 | (((named_intros, SOME elim), false), no_compilation)))) thy | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 343 | else thy | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 344 | end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 345 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 346 | fun register_intros (constname, pre_intros) thy = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 347 | let | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 348 | val T = Sign.the_const_type thy constname | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 349 | fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl intr))) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 350 | val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 351 |       error ("register_intros: Introduction rules of different constants are used\n" ^
 | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 352 | "expected rules for " ^ constname ^ ", but received rules for " ^ | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 353 | commas (map constname_of_intro pre_intros)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 354 | else () | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 355 | val pred = Const (constname, T) | 
| 59501 | 356 | val pre_elim = | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 357 | (Drule.export_without_context o Skip_Proof.make_thm thy) | 
| 42361 | 358 | (mk_casesrule (Proof_Context.init_global thy) pred pre_intros) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 359 | in register_predicate (constname, pre_intros, pre_elim) thy end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 360 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 361 | fun defined_function_of compilation pred = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 362 | let | 
| 55543 | 363 | val set = (apsnd o apsnd o apfst) (cons (compilation, [])) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 364 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 365 | PredData.map (Graph.map_node pred (map_pred_data set)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 366 | end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 367 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 368 | fun set_function_name compilation pred mode name = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 369 | let | 
| 55543 | 370 | val set = (apsnd o apsnd o apfst) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 371 | (AList.map_default (op =) (compilation, [(mode, name)]) (cons (mode, name))) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 372 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 373 | PredData.map (Graph.map_node pred (map_pred_data set)) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 374 | end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 375 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 376 | fun set_needs_random name modes = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 377 | let | 
| 55543 | 378 | val set = (apsnd o apsnd o apsnd o apsnd) (K modes) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 379 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 380 | PredData.map (Graph.map_node name (map_pred_data set)) | 
| 59501 | 381 | end | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 382 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 383 | fun extend' value_of edges_of key (G, visited) = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 384 | let | 
| 55437 | 385 | val (G', v) = | 
| 386 | (case try (Graph.get_node G) key of | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 387 | SOME v => (G, v) | 
| 55437 | 388 | | NONE => (Graph.new_node (key, value_of key) G, value_of key)) | 
| 389 | val (G'', visited') = | |
| 390 | fold (extend' value_of edges_of) | |
| 391 | (subtract (op =) visited (edges_of (key, v))) | |
| 392 | (G', key :: visited) | |
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 393 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 394 | (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited') | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 395 | end; | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 396 | |
| 59501 | 397 | fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 398 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 399 | fun extend_intro_graph names thy = | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 400 | let | 
| 42361 | 401 | val ctxt = Proof_Context.init_global thy | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 402 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 403 | PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names) thy | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 404 | end | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 405 | |
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 406 | fun preprocess_intros name thy = | 
| 55543 | 407 | PredData.map (Graph.map_node name (map_pred_data (apsnd (apfst (fn (rules, preprocessed) => | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 408 | if preprocessed then (rules, preprocessed) | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 409 | else | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 410 | let | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 411 | val (named_intros, SOME elim) = rules | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 412 | val named_intros' = map (apsnd (preprocess_intro thy)) named_intros | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 413 | val pred = Const (name, Sign.the_const_type thy name) | 
| 42361 | 414 | val ctxt = Proof_Context.init_global thy | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 415 | val elim_t = mk_casesrule ctxt pred (map snd named_intros') | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 416 | val elim' = prove_casesrule ctxt (pred, (elim, 0)) elim_t | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 417 | in | 
| 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 418 | ((named_intros', SOME elim'), true) | 
| 55543 | 419 | end))))) | 
| 59501 | 420 | thy | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
40052diff
changeset | 421 | |
| 55437 | 422 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 423 | (* registration of alternative function names *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 424 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 425 | structure Alt_Compilations_Data = Theory_Data | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 426 | ( | 
| 55437 | 427 | type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table | 
| 428 | val empty = Symtab.empty | |
| 429 | fun merge data : T = Symtab.merge (K true) data | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 430 | ); | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 431 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 432 | fun alternative_compilation_of_global thy pred_name mode = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 433 | AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 434 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 435 | fun alternative_compilation_of ctxt pred_name mode = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 436 | AList.lookup eq_mode | 
| 42361 | 437 | (Symtab.lookup_list (Alt_Compilations_Data.get (Proof_Context.theory_of ctxt)) pred_name) mode | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 438 | |
| 55543 | 439 | fun force_modes_and_compilations pred_name compilations thy = | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 440 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 441 | (* thm refl is a dummy thm *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 442 | val modes = map fst compilations | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58823diff
changeset | 443 | val (needs_random, non_random_modes) = | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58823diff
changeset | 444 | apply2 (map fst) (List.partition (fn (_, (_, random)) => random) compilations) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 445 | val non_random_dummys = map (rpair "dummy") non_random_modes | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 446 | val all_dummys = map (rpair "dummy") modes | 
| 55437 | 447 | val dummy_function_names = | 
| 448 | map (rpair all_dummys) Predicate_Compile_Aux.random_compilations @ | |
| 449 | map (rpair non_random_dummys) Predicate_Compile_Aux.non_random_compilations | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 450 | val alt_compilations = map (apsnd fst) compilations | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 451 | in | 
| 55543 | 452 | thy |> | 
| 55437 | 453 | PredData.map | 
| 454 | (Graph.new_node | |
| 455 | (pred_name, | |
| 55543 | 456 | mk_pred_data | 
| 457 | (Position.thread_data (), | |
| 458 |               ((([], SOME @{thm refl}), true), (dummy_function_names, ([], needs_random))))))
 | |
| 459 | |> Alt_Compilations_Data.map (Symtab.insert (K false) (pred_name, alt_compilations)) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 460 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 461 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 462 | fun functional_compilation fun_name mode compfuns T = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 463 | let | 
| 55437 | 464 | val (inpTs, outpTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE)) mode (binder_types T) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 465 | val bs = map (pair "x") inpTs | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 466 | val bounds = map Bound (rev (0 upto (length bs) - 1)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 467 | val f = Const (fun_name, inpTs ---> HOLogic.mk_tupleT outpTs) | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
42361diff
changeset | 468 | in fold_rev Term.abs bs (mk_single compfuns (list_comb (f, bounds))) end | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 469 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 470 | fun register_alternative_function pred_name mode fun_name = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 471 | Alt_Compilations_Data.map (Symtab.insert_list (eq_pair eq_mode (K false)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 472 | (pred_name, (mode, functional_compilation fun_name mode))) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 473 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 474 | fun force_modes_and_functions pred_name fun_names = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 475 | force_modes_and_compilations pred_name | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 476 | (map (fn (mode, (fun_name, random)) => (mode, (functional_compilation fun_name mode, random))) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 477 | fun_names) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 478 | |
| 62391 | 479 | end |