| author | wenzelm | 
| Sat, 02 Aug 2014 20:58:15 +0200 | |
| changeset 57845 | a2340800ca1f | 
| parent 56723 | a8f71445c265 | 
| child 59150 | 71b416020f42 | 
| 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: 
38757 
diff
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  | 
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
14  | 
|
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
15  | 
val overloading: (string * (string * typ) * bool) list -> theory -> local_theory  | 
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
16  | 
val overloading_cmd: (string * string * bool) list -> theory -> local_theory  | 
| 25519 | 17  | 
end;  | 
18  | 
||
19  | 
structure Overloading: OVERLOADING =  | 
|
20  | 
struct  | 
|
21  | 
||
| 42404 | 22  | 
(* generic check/uncheck combinators for improvable constants *)  | 
| 26238 | 23  | 
|
| 26249 | 24  | 
type improvable_syntax = ((((string * typ) list * (string * typ) list) *  | 
| 26730 | 25  | 
((((string * typ -> (typ * typ) option) * (string * typ -> (typ * term) option)) * bool) *  | 
| 26249 | 26  | 
(term * term) list)) * bool);  | 
| 25519 | 27  | 
|
| 42404 | 28  | 
structure Improvable_Syntax = Proof_Data  | 
| 33519 | 29  | 
(  | 
| 26249 | 30  | 
  type T = {
 | 
| 26520 | 31  | 
primary_constraints: (string * typ) list,  | 
32  | 
secondary_constraints: (string * typ) list,  | 
|
| 26249 | 33  | 
improve: string * typ -> (typ * typ) option,  | 
34  | 
subst: string * typ -> (typ * term) option,  | 
|
| 26730 | 35  | 
consider_abbrevs: bool,  | 
| 26249 | 36  | 
unchecks: (term * term) list,  | 
37  | 
passed: bool  | 
|
38  | 
};  | 
|
| 26238 | 39  | 
  fun init _ = {
 | 
| 26520 | 40  | 
primary_constraints = [],  | 
41  | 
secondary_constraints = [],  | 
|
| 26238 | 42  | 
improve = K NONE,  | 
43  | 
subst = K NONE,  | 
|
| 26730 | 44  | 
consider_abbrevs = false,  | 
| 26238 | 45  | 
unchecks = [],  | 
46  | 
passed = true  | 
|
47  | 
};  | 
|
48  | 
);  | 
|
| 25536 | 49  | 
|
| 42404 | 50  | 
fun map_improvable_syntax f = Improvable_Syntax.map (fn {primary_constraints,
 | 
51  | 
secondary_constraints, improve, subst, consider_abbrevs, unchecks, passed} =>  | 
|
| 40782 | 52  | 
let  | 
| 26730 | 53  | 
val (((primary_constraints', secondary_constraints'),  | 
54  | 
(((improve', subst'), consider_abbrevs'), unchecks')), passed')  | 
|
55  | 
= f (((primary_constraints, secondary_constraints),  | 
|
56  | 
(((improve, subst), consider_abbrevs), unchecks)), passed)  | 
|
| 42404 | 57  | 
in  | 
58  | 
   {primary_constraints = primary_constraints', secondary_constraints = secondary_constraints',
 | 
|
| 26730 | 59  | 
improve = improve', subst = subst', consider_abbrevs = consider_abbrevs',  | 
| 42404 | 60  | 
unchecks = unchecks', passed = passed'}  | 
| 40782 | 61  | 
end);  | 
| 26238 | 62  | 
|
| 26249 | 63  | 
val mark_passed = (map_improvable_syntax o apsnd) (K true);  | 
| 26238 | 64  | 
|
65  | 
fun improve_term_check ts ctxt =  | 
|
| 25519 | 66  | 
let  | 
| 42388 | 67  | 
val thy = Proof_Context.theory_of ctxt;  | 
68  | 
||
| 42404 | 69  | 
    val {secondary_constraints, improve, subst, consider_abbrevs, passed, ...} =
 | 
70  | 
Improvable_Syntax.get ctxt;  | 
|
| 42360 | 71  | 
val is_abbrev = consider_abbrevs andalso Proof_Context.abbrev_mode ctxt;  | 
| 26730 | 72  | 
val passed_or_abbrev = passed orelse is_abbrev;  | 
| 42388 | 73  | 
fun accumulate_improvements (Const (c, ty)) =  | 
74  | 
(case improve (c, ty) of  | 
|
75  | 
SOME ty_ty' => Sign.typ_match thy ty_ty'  | 
|
| 26238 | 76  | 
| _ => I)  | 
77  | 
| accumulate_improvements _ = I;  | 
|
78  | 
val improvements = (fold o fold_aterms) accumulate_improvements ts Vartab.empty;  | 
|
| 32035 | 79  | 
val ts' = (map o map_types) (Envir.subst_type improvements) ts;  | 
| 42388 | 80  | 
fun apply_subst t =  | 
81  | 
Envir.expand_term  | 
|
82  | 
(fn Const (c, ty) =>  | 
|
83  | 
(case subst (c, ty) of  | 
|
84  | 
SOME (ty', t') =>  | 
|
85  | 
if Sign.typ_instance thy (ty, ty')  | 
|
| 26238 | 86  | 
then SOME (ty', apply_subst t') else NONE  | 
87  | 
| NONE => NONE)  | 
|
| 26259 | 88  | 
| _ => NONE) t;  | 
| 26730 | 89  | 
val ts'' = if is_abbrev then ts' else map apply_subst ts';  | 
| 40782 | 90  | 
in  | 
| 42404 | 91  | 
if eq_list (op aconv) (ts, ts'') andalso passed_or_abbrev then NONE  | 
92  | 
else if passed_or_abbrev then SOME (ts'', ctxt)  | 
|
93  | 
else  | 
|
94  | 
SOME (ts'', ctxt  | 
|
95  | 
|> fold (Proof_Context.add_const_constraint o apsnd SOME) secondary_constraints  | 
|
96  | 
|> mark_passed)  | 
|
| 26238 | 97  | 
end;  | 
| 25519 | 98  | 
|
| 31698 | 99  | 
fun rewrite_liberal thy unchecks t =  | 
| 42404 | 100  | 
(case try (Pattern.rewrite_term thy unchecks []) t of  | 
101  | 
NONE => NONE  | 
|
102  | 
| SOME t' => if t aconv t' then NONE else SOME t');  | 
|
| 31698 | 103  | 
|
| 26238 | 104  | 
fun improve_term_uncheck ts ctxt =  | 
| 25519 | 105  | 
let  | 
| 42360 | 106  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 42404 | 107  | 
    val {unchecks, ...} = Improvable_Syntax.get ctxt;
 | 
| 31698 | 108  | 
val ts' = map (rewrite_liberal thy unchecks) ts;  | 
109  | 
in if exists is_some ts' then SOME (map2 the_default ts ts', ctxt) else NONE end;  | 
|
| 26238 | 110  | 
|
| 26520 | 111  | 
fun set_primary_constraints ctxt =  | 
| 42404 | 112  | 
  let val {primary_constraints, ...} = Improvable_Syntax.get ctxt;
 | 
| 42360 | 113  | 
in fold (Proof_Context.add_const_constraint o apsnd SOME) primary_constraints ctxt end;  | 
| 26259 | 114  | 
|
| 
39378
 
df86b1b4ce10
more precise name for activation of improveable syntax
 
haftmann 
parents: 
38757 
diff
changeset
 | 
115  | 
val activate_improvable_syntax =  | 
| 26259 | 116  | 
Context.proof_map  | 
| 45444 | 117  | 
(Syntax_Phases.term_check' 0 "improvement" improve_term_check  | 
118  | 
#> Syntax_Phases.term_uncheck' 0 "improvement" improve_term_uncheck)  | 
|
| 26520 | 119  | 
#> set_primary_constraints;  | 
| 26259 | 120  | 
|
121  | 
||
| 42404 | 122  | 
(* overloading target *)  | 
| 26259 | 123  | 
|
| 38382 | 124  | 
structure Data = Proof_Data  | 
| 26259 | 125  | 
(  | 
126  | 
type T = ((string * typ) * (string * bool)) list;  | 
|
127  | 
fun init _ = [];  | 
|
128  | 
);  | 
|
129  | 
||
| 38382 | 130  | 
val get_overloading = Data.get o Local_Theory.target_of;  | 
131  | 
val map_overloading = Local_Theory.target o Data.map;  | 
|
| 26259 | 132  | 
|
| 42404 | 133  | 
fun operation lthy b =  | 
134  | 
get_overloading lthy  | 
|
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
29606 
diff
changeset
 | 
135  | 
|> get_first (fn ((c, _), (v, checked)) =>  | 
| 38382 | 136  | 
if Binding.name_of b = v then SOME (c, (v, checked)) else NONE);  | 
| 26259 | 137  | 
|
| 32343 | 138  | 
fun synchronize_syntax ctxt =  | 
139  | 
let  | 
|
| 38382 | 140  | 
val overloading = Data.get ctxt;  | 
| 42404 | 141  | 
fun subst (c, ty) =  | 
142  | 
(case AList.lookup (op =) overloading (c, ty) of  | 
|
143  | 
SOME (v, _) => SOME (ty, Free (v, ty))  | 
|
144  | 
| NONE => NONE);  | 
|
| 32343 | 145  | 
val unchecks =  | 
146  | 
map (fn (c_ty as (_, ty), (v, _)) => (Free (v, ty), Const c_ty)) overloading;  | 
|
147  | 
in  | 
|
148  | 
ctxt  | 
|
149  | 
|> map_improvable_syntax (K ((([], []), (((K NONE, subst), false), unchecks)), false))  | 
|
| 
46916
 
e7ea35b41e2d
Local_Theory.define no longer hard-wires default theorem name -- targets/packages need to take care of it;
 
wenzelm 
parents: 
45444 
diff
changeset
 | 
150  | 
end;  | 
| 26259 | 151  | 
|
| 38382 | 152  | 
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: 
38382 
diff
changeset
 | 
153  | 
Local_Theory.background_theory_result  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
154  | 
(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: 
45444 
diff
changeset
 | 
155  | 
(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: 
45444 
diff
changeset
 | 
156  | 
Logic.mk_equals (Const (c, Term.fastype_of rhs), rhs)))  | 
| 38382 | 157  | 
##> 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: 
47081 
diff
changeset
 | 
158  | 
##> 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: 
45444 
diff
changeset
 | 
159  | 
#-> (fn (_, def) => pair (Const (c, U), def));  | 
| 26259 | 160  | 
|
| 47289 | 161  | 
fun foundation (((b, U), mx), (b_def, rhs)) params lthy =  | 
| 42404 | 162  | 
(case operation lthy b of  | 
163  | 
SOME (c, (v, checked)) =>  | 
|
164  | 
if mx <> NoSyn  | 
|
165  | 
      then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
 | 
|
166  | 
else lthy |> define_overloaded (c, U) (v, checked) (b_def, rhs)  | 
|
| 47289 | 167  | 
| NONE => lthy |> Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) params);  | 
| 25519 | 168  | 
|
| 25606 | 169  | 
fun pretty lthy =  | 
170  | 
let  | 
|
171  | 
val overloading = get_overloading lthy;  | 
|
172  | 
fun pr_operation ((c, ty), (v, _)) =  | 
|
| 42359 | 173  | 
Pretty.block (Pretty.breaks  | 
| 55304 | 174  | 
[Pretty.str v, Pretty.str "==", Proof_Context.pretty_const lthy c,  | 
| 42359 | 175  | 
Pretty.str "::", Syntax.pretty_typ lthy ty]);  | 
| 55763 | 176  | 
in Pretty.keyword1 "overloading" :: map pr_operation overloading end;  | 
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
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: 
36610 
diff
changeset
 | 
187  | 
|
| 38382 | 188  | 
fun gen_overloading prep_const raw_overloading thy =  | 
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
189  | 
let  | 
| 42360 | 190  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
47061
 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 
wenzelm 
parents: 
46923 
diff
changeset
 | 
191  | 
val naming = Sign.naming_of thy;  | 
| 38382 | 192  | 
val _ = if null raw_overloading then error "At least one parameter must be given" else ();  | 
193  | 
val overloading = raw_overloading |> map (fn (v, const, checked) =>  | 
|
194  | 
(Term.dest_Const (prep_const ctxt const), (v, checked)));  | 
|
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
195  | 
in  | 
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
196  | 
thy  | 
| 
56056
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
55763 
diff
changeset
 | 
197  | 
|> Sign.change_begin  | 
| 42360 | 198  | 
|> Proof_Context.init_global  | 
| 38382 | 199  | 
|> Data.put overloading  | 
200  | 
|> fold (fn ((_, ty), (v, _)) => Variable.declare_names (Free (v, ty))) overloading  | 
|
| 
39378
 
df86b1b4ce10
more precise name for activation of improveable syntax
 
haftmann 
parents: 
38757 
diff
changeset
 | 
201  | 
|> activate_improvable_syntax  | 
| 38382 | 202  | 
|> synchronize_syntax  | 
| 
47061
 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 
wenzelm 
parents: 
46923 
diff
changeset
 | 
203  | 
|> Local_Theory.init naming  | 
| 38382 | 204  | 
       {define = Generic_Target.define foundation,
 | 
| 
47250
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47245 
diff
changeset
 | 
205  | 
notes = Generic_Target.notes Generic_Target.theory_notes,  | 
| 
47286
 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 
wenzelm 
parents: 
47279 
diff
changeset
 | 
206  | 
abbrev = Generic_Target.abbrev Generic_Target.theory_abbrev,  | 
| 
47279
 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 
wenzelm 
parents: 
47250 
diff
changeset
 | 
207  | 
declaration = K Generic_Target.theory_declaration,  | 
| 
56723
 
a8f71445c265
subscription as target-specific implementation device
 
haftmann 
parents: 
56056 
diff
changeset
 | 
208  | 
subscription = Generic_Target.theory_registration,  | 
| 38382 | 209  | 
pretty = pretty,  | 
| 
56056
 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 
wenzelm 
parents: 
55763 
diff
changeset
 | 
210  | 
exit = conclude #> Local_Theory.target_of #> Sign.change_end_local}  | 
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
211  | 
end;  | 
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
212  | 
|
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
213  | 
val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);  | 
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
214  | 
val overloading_cmd = gen_overloading Syntax.read_term;  | 
| 
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
36610 
diff
changeset
 | 
215  | 
|
| 25519 | 216  | 
end;  |