| author | wenzelm | 
| Sat, 27 Jan 2018 17:27:06 +0100 | |
| changeset 67516 | 656720e8f443 | 
| parent 67147 | dea94b1aabc3 | 
| child 69829 | 3bfa28b3a5b2 | 
| permissions | -rw-r--r-- | 
| 25519 | 1 | (* Title: Pure/Isar/overloading.ML | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | ||
| 4 | Overloaded definitions without any discipline. | |
| 5 | *) | |
| 6 | ||
| 7 | signature OVERLOADING = | |
| 8 | sig | |
| 26238 | 9 | type improvable_syntax | 
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
38757diff
changeset | 10 | val activate_improvable_syntax: Proof.context -> Proof.context | 
| 26238 | 11 | val map_improvable_syntax: (improvable_syntax -> improvable_syntax) | 
| 12 | -> Proof.context -> Proof.context | |
| 26520 | 13 | val set_primary_constraints: Proof.context -> Proof.context | 
| 60339 
0e6742f89c03
dedicated config options to deactivate uncheck phase for improvable syntax
 haftmann parents: 
60338diff
changeset | 14 | val show_reverted_improvements: bool Config.T; | 
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 15 | |
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 16 | val overloading: (string * (string * typ) * bool) list -> theory -> local_theory | 
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 17 | val overloading_cmd: (string * string * bool) list -> theory -> local_theory | 
| 25519 | 18 | end; | 
| 19 | ||
| 20 | structure Overloading: OVERLOADING = | |
| 21 | struct | |
| 22 | ||
| 42404 | 23 | (* generic check/uncheck combinators for improvable constants *) | 
| 26238 | 24 | |
| 60338 | 25 | type improvable_syntax = {
 | 
| 26 | primary_constraints: (string * typ) list, | |
| 27 | secondary_constraints: (string * typ) list, | |
| 28 | improve: string * typ -> (typ * typ) option, | |
| 29 | subst: string * typ -> (typ * term) option, | |
| 30 | no_subst_in_abbrev_mode: bool, | |
| 31 | unchecks: (term * term) list | |
| 32 | } | |
| 25519 | 33 | |
| 42404 | 34 | structure Improvable_Syntax = Proof_Data | 
| 33519 | 35 | ( | 
| 60338 | 36 |   type T = {syntax: improvable_syntax, secondary_pass: bool}
 | 
| 37 |   fun init _ = {syntax = {
 | |
| 38 | primary_constraints = [], | |
| 39 | secondary_constraints = [], | |
| 40 | improve = K NONE, | |
| 41 | subst = K NONE, | |
| 42 | no_subst_in_abbrev_mode = false, | |
| 43 | unchecks = [] | |
| 44 | }, secondary_pass = false}: T; | |
| 26238 | 45 | ); | 
| 25536 | 46 | |
| 60338 | 47 | fun map_improvable_syntax f = | 
| 48 |   Improvable_Syntax.map (fn {syntax, secondary_pass} => {syntax = f syntax, secondary_pass = secondary_pass});
 | |
| 26238 | 49 | |
| 60338 | 50 | val mark_passed = | 
| 51 |   Improvable_Syntax.map (fn {syntax, secondary_pass} => {syntax = syntax, secondary_pass = true});
 | |
| 26238 | 52 | |
| 53 | fun improve_term_check ts ctxt = | |
| 25519 | 54 | let | 
| 42388 | 55 | val thy = Proof_Context.theory_of ctxt; | 
| 56 | ||
| 60338 | 57 |     val {syntax = {secondary_constraints, improve, subst, no_subst_in_abbrev_mode, ...}, secondary_pass} =
 | 
| 42404 | 58 | Improvable_Syntax.get ctxt; | 
| 60338 | 59 | val no_subst = Proof_Context.abbrev_mode ctxt andalso no_subst_in_abbrev_mode; | 
| 42388 | 60 | fun accumulate_improvements (Const (c, ty)) = | 
| 61 | (case improve (c, ty) of | |
| 62 | SOME ty_ty' => Sign.typ_match thy ty_ty' | |
| 26238 | 63 | | _ => I) | 
| 64 | | accumulate_improvements _ = I; | |
| 42388 | 65 | fun apply_subst t = | 
| 66 | Envir.expand_term | |
| 67 | (fn Const (c, ty) => | |
| 68 | (case subst (c, ty) of | |
| 69 | SOME (ty', t') => | |
| 70 | if Sign.typ_instance thy (ty, ty') | |
| 26238 | 71 | then SOME (ty', apply_subst t') else NONE | 
| 72 | | NONE => NONE) | |
| 26259 | 73 | | _ => NONE) t; | 
| 60338 | 74 | val improvements = (fold o fold_aterms) accumulate_improvements ts Vartab.empty; | 
| 75 | val ts' = | |
| 76 | ts | |
| 77 | |> (map o map_types) (Envir.subst_type improvements) | |
| 78 | |> not no_subst ? map apply_subst; | |
| 40782 | 79 | in | 
| 60338 | 80 | if secondary_pass orelse no_subst then | 
| 81 | if eq_list (op aconv) (ts, ts') then NONE | |
| 82 | else SOME (ts', ctxt) | |
| 42404 | 83 | else | 
| 60338 | 84 | SOME (ts', ctxt | 
| 42404 | 85 | |> fold (Proof_Context.add_const_constraint o apsnd SOME) secondary_constraints | 
| 86 | |> mark_passed) | |
| 26238 | 87 | end; | 
| 25519 | 88 | |
| 31698 | 89 | fun rewrite_liberal thy unchecks t = | 
| 60347 | 90 | (case try (Pattern.rewrite_term_top thy unchecks []) t of | 
| 42404 | 91 | NONE => NONE | 
| 92 | | SOME t' => if t aconv t' then NONE else SOME t'); | |
| 31698 | 93 | |
| 60339 
0e6742f89c03
dedicated config options to deactivate uncheck phase for improvable syntax
 haftmann parents: 
60338diff
changeset | 94 | val show_reverted_improvements = | 
| 67147 | 95 | Attrib.setup_config_bool \<^binding>\<open>show_reverted_improvements\<close> (K true); | 
| 60339 
0e6742f89c03
dedicated config options to deactivate uncheck phase for improvable syntax
 haftmann parents: 
60338diff
changeset | 96 | |
| 26238 | 97 | fun improve_term_uncheck ts ctxt = | 
| 25519 | 98 | let | 
| 42360 | 99 | val thy = Proof_Context.theory_of ctxt; | 
| 60338 | 100 |     val {syntax = {unchecks, ...}, ...} = Improvable_Syntax.get ctxt;
 | 
| 60339 
0e6742f89c03
dedicated config options to deactivate uncheck phase for improvable syntax
 haftmann parents: 
60338diff
changeset | 101 | val revert = Config.get ctxt show_reverted_improvements; | 
| 31698 | 102 | val ts' = map (rewrite_liberal thy unchecks) ts; | 
| 60339 
0e6742f89c03
dedicated config options to deactivate uncheck phase for improvable syntax
 haftmann parents: 
60338diff
changeset | 103 | in if revert andalso exists is_some ts' then SOME (map2 the_default ts ts', ctxt) else NONE end; | 
| 26238 | 104 | |
| 26520 | 105 | fun set_primary_constraints ctxt = | 
| 60338 | 106 | let | 
| 107 |     val {syntax = {primary_constraints, ...}, ...} = Improvable_Syntax.get ctxt;
 | |
| 42360 | 108 | in fold (Proof_Context.add_const_constraint o apsnd SOME) primary_constraints ctxt end; | 
| 26259 | 109 | |
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
38757diff
changeset | 110 | val activate_improvable_syntax = | 
| 26259 | 111 | Context.proof_map | 
| 45444 | 112 | (Syntax_Phases.term_check' 0 "improvement" improve_term_check | 
| 113 | #> Syntax_Phases.term_uncheck' 0 "improvement" improve_term_uncheck) | |
| 26520 | 114 | #> set_primary_constraints; | 
| 26259 | 115 | |
| 116 | ||
| 42404 | 117 | (* overloading target *) | 
| 26259 | 118 | |
| 38382 | 119 | structure Data = Proof_Data | 
| 26259 | 120 | ( | 
| 121 | type T = ((string * typ) * (string * bool)) list; | |
| 122 | fun init _ = []; | |
| 123 | ); | |
| 124 | ||
| 38382 | 125 | val get_overloading = Data.get o Local_Theory.target_of; | 
| 126 | val map_overloading = Local_Theory.target o Data.map; | |
| 26259 | 127 | |
| 42404 | 128 | fun operation lthy b = | 
| 129 | get_overloading lthy | |
| 30519 
c05c0199826f
coherent binding policy with primitive target operations
 haftmann parents: 
29606diff
changeset | 130 | |> get_first (fn ((c, _), (v, checked)) => | 
| 38382 | 131 | if Binding.name_of b = v then SOME (c, (v, checked)) else NONE); | 
| 26259 | 132 | |
| 32343 | 133 | fun synchronize_syntax ctxt = | 
| 134 | let | |
| 38382 | 135 | val overloading = Data.get ctxt; | 
| 42404 | 136 | fun subst (c, ty) = | 
| 137 | (case AList.lookup (op =) overloading (c, ty) of | |
| 138 | SOME (v, _) => SOME (ty, Free (v, ty)) | |
| 139 | | NONE => NONE); | |
| 32343 | 140 | val unchecks = | 
| 141 | map (fn (c_ty as (_, ty), (v, _)) => (Free (v, ty), Const c_ty)) overloading; | |
| 64596 | 142 | in | 
| 32343 | 143 | ctxt | 
| 60338 | 144 |     |> map_improvable_syntax (K {primary_constraints = [],
 | 
| 145 | secondary_constraints = [], improve = K NONE, subst = subst, | |
| 146 | no_subst_in_abbrev_mode = false, unchecks = unchecks}) | |
| 46916 
e7ea35b41e2d
Local_Theory.define no longer hard-wires default theorem name -- targets/packages need to take care of it;
 wenzelm parents: 
45444diff
changeset | 147 | end; | 
| 26259 | 148 | |
| 38382 | 149 | fun define_overloaded (c, U) (v, checked) (b_def, rhs) = | 
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38382diff
changeset | 150 | Local_Theory.background_theory_result | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 151 | (Thm.add_def_global (not checked) true | 
| 46916 
e7ea35b41e2d
Local_Theory.define no longer hard-wires default theorem name -- targets/packages need to take care of it;
 wenzelm parents: 
45444diff
changeset | 152 | (Thm.def_binding_optional (Binding.name v) b_def, | 
| 
e7ea35b41e2d
Local_Theory.define no longer hard-wires default theorem name -- targets/packages need to take care of it;
 wenzelm parents: 
45444diff
changeset | 153 | Logic.mk_equals (Const (c, Term.fastype_of rhs), rhs))) | 
| 38382 | 154 | ##> map_overloading (filter_out (fn (_, (v', _)) => v' = v)) | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 155 | ##> Local_Theory.map_contexts (K synchronize_syntax) | 
| 46916 
e7ea35b41e2d
Local_Theory.define no longer hard-wires default theorem name -- targets/packages need to take care of it;
 wenzelm parents: 
45444diff
changeset | 156 | #-> (fn (_, def) => pair (Const (c, U), def)); | 
| 26259 | 157 | |
| 47289 | 158 | fun foundation (((b, U), mx), (b_def, rhs)) params lthy = | 
| 42404 | 159 | (case operation lthy b of | 
| 160 | SOME (c, (v, checked)) => | |
| 62752 | 161 | if Mixfix.is_empty mx then | 
| 162 | lthy |> define_overloaded (c, U) (v, checked) (b_def, rhs) | |
| 62765 
5b95a12b7b19
tuned messages -- position is usually missing here;
 wenzelm parents: 
62752diff
changeset | 163 |       else error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
 | 
| 60341 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60339diff
changeset | 164 | | NONE => lthy |> Generic_Target.theory_target_foundation (((b, U), mx), (b_def, rhs)) params); | 
| 25519 | 165 | |
| 25606 | 166 | fun pretty lthy = | 
| 167 | let | |
| 168 | val overloading = get_overloading lthy; | |
| 169 | fun pr_operation ((c, ty), (v, _)) = | |
| 42359 | 170 | Pretty.block (Pretty.breaks | 
| 64596 | 171 | [Pretty.str v, Pretty.str "\<equiv>", Proof_Context.pretty_const lthy c, | 
| 42359 | 172 | Pretty.str "::", Syntax.pretty_typ lthy ty]); | 
| 60246 | 173 | in | 
| 174 | [Pretty.block | |
| 175 | (Pretty.fbreaks (Pretty.keyword1 "overloading" :: map pr_operation overloading))] | |
| 176 | end; | |
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 177 | |
| 38382 | 178 | fun conclude lthy = | 
| 179 | let | |
| 180 | val overloading = get_overloading lthy; | |
| 40782 | 181 | val _ = | 
| 182 | if null overloading then () | |
| 183 | else | |
| 42404 | 184 |         error ("Missing definition(s) for parameter(s) " ^
 | 
| 185 | commas_quote (map (Syntax.string_of_term lthy o Const o fst) overloading)); | |
| 38382 | 186 | in lthy end; | 
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 187 | |
| 38382 | 188 | fun gen_overloading prep_const raw_overloading thy = | 
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 189 | let | 
| 42360 | 190 | val ctxt = Proof_Context.init_global thy; | 
| 38382 | 191 | val _ = if null raw_overloading then error "At least one parameter must be given" else (); | 
| 192 | val overloading = raw_overloading |> map (fn (v, const, checked) => | |
| 193 | (Term.dest_Const (prep_const ctxt const), (v, checked))); | |
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 194 | in | 
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 195 | thy | 
| 66337 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 196 | |> Generic_Target.init | 
| 66335 
a849ce33923d
treat exit separate from regular local theory operations
 haftmann parents: 
64596diff
changeset | 197 |        {background_naming = Sign.naming_of thy,
 | 
| 66337 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 198 | setup = Proof_Context.init_global | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 199 | #> Data.put overloading | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 200 | #> fold (fn ((_, ty), (v, _)) => Variable.declare_names (Free (v, ty))) overloading | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 201 | #> activate_improvable_syntax | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 202 | #> synchronize_syntax, | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66335diff
changeset | 203 | conclude = conclude} | 
| 38382 | 204 |        {define = Generic_Target.define foundation,
 | 
| 60341 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60339diff
changeset | 205 | notes = Generic_Target.notes Generic_Target.theory_target_notes, | 
| 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60339diff
changeset | 206 | abbrev = Generic_Target.abbrev Generic_Target.theory_target_abbrev, | 
| 47279 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 wenzelm parents: 
47250diff
changeset | 207 | declaration = K Generic_Target.theory_declaration, | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
60347diff
changeset | 208 | theory_registration = Generic_Target.theory_registration, | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
60347diff
changeset | 209 | locale_dependency = fn _ => error "Not possible in overloading target", | 
| 66335 
a849ce33923d
treat exit separate from regular local theory operations
 haftmann parents: 
64596diff
changeset | 210 | pretty = pretty} | 
| 38342 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 211 | end; | 
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 212 | |
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 213 | val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const); | 
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 214 | val overloading_cmd = gen_overloading Syntax.read_term; | 
| 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 haftmann parents: 
36610diff
changeset | 215 | |
| 25519 | 216 | end; |