author | wenzelm |
Thu, 19 Jun 2008 21:14:30 +0200 | |
changeset 27283 | ebd0291ea79c |
parent 26939 | 1035c89b4c02 |
child 27285 | def40a211768 |
permissions | -rw-r--r-- |
25519 | 1 |
(* Title: Pure/Isar/overloading.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Overloaded definitions without any discipline. |
|
6 |
*) |
|
7 |
||
8 |
signature OVERLOADING = |
|
9 |
sig |
|
25861 | 10 |
val init: (string * (string * typ) * bool) list -> theory -> local_theory |
25519 | 11 |
val conclude: local_theory -> local_theory |
12 |
val declare: string * typ -> theory -> term * theory |
|
13 |
val confirm: string -> local_theory -> local_theory |
|
14 |
val define: bool -> string -> string * term -> theory -> thm * theory |
|
15 |
val operation: Proof.context -> string -> (string * bool) option |
|
25606 | 16 |
val pretty: Proof.context -> Pretty.T |
26238 | 17 |
|
18 |
type improvable_syntax |
|
19 |
val add_improvable_syntax: Proof.context -> Proof.context |
|
20 |
val map_improvable_syntax: (improvable_syntax -> improvable_syntax) |
|
21 |
-> Proof.context -> Proof.context |
|
26520 | 22 |
val set_primary_constraints: Proof.context -> Proof.context |
25519 | 23 |
end; |
24 |
||
25 |
structure Overloading: OVERLOADING = |
|
26 |
struct |
|
27 |
||
26259 | 28 |
(** generic check/uncheck combinators for improvable constants **) |
26238 | 29 |
|
26249 | 30 |
type improvable_syntax = ((((string * typ) list * (string * typ) list) * |
26730 | 31 |
((((string * typ -> (typ * typ) option) * (string * typ -> (typ * term) option)) * bool) * |
26249 | 32 |
(term * term) list)) * bool); |
25519 | 33 |
|
26238 | 34 |
structure ImprovableSyntax = ProofDataFun( |
26249 | 35 |
type T = { |
26520 | 36 |
primary_constraints: (string * typ) list, |
37 |
secondary_constraints: (string * typ) list, |
|
26249 | 38 |
improve: string * typ -> (typ * typ) option, |
39 |
subst: string * typ -> (typ * term) option, |
|
26730 | 40 |
consider_abbrevs: bool, |
26249 | 41 |
unchecks: (term * term) list, |
42 |
passed: bool |
|
43 |
}; |
|
26238 | 44 |
fun init _ = { |
26520 | 45 |
primary_constraints = [], |
46 |
secondary_constraints = [], |
|
26238 | 47 |
improve = K NONE, |
48 |
subst = K NONE, |
|
26730 | 49 |
consider_abbrevs = false, |
26238 | 50 |
unchecks = [], |
51 |
passed = true |
|
52 |
}; |
|
53 |
); |
|
25536 | 54 |
|
26520 | 55 |
fun map_improvable_syntax f = ImprovableSyntax.map (fn { primary_constraints, |
26730 | 56 |
secondary_constraints, improve, subst, consider_abbrevs, unchecks, passed } => let |
57 |
val (((primary_constraints', secondary_constraints'), |
|
58 |
(((improve', subst'), consider_abbrevs'), unchecks')), passed') |
|
59 |
= f (((primary_constraints, secondary_constraints), |
|
60 |
(((improve, subst), consider_abbrevs), unchecks)), passed) |
|
26520 | 61 |
in { primary_constraints = primary_constraints', secondary_constraints = secondary_constraints', |
26730 | 62 |
improve = improve', subst = subst', consider_abbrevs = consider_abbrevs', |
63 |
unchecks = unchecks', passed = passed' |
|
26249 | 64 |
} end); |
26238 | 65 |
|
26249 | 66 |
val mark_passed = (map_improvable_syntax o apsnd) (K true); |
26238 | 67 |
|
68 |
fun improve_term_check ts ctxt = |
|
25519 | 69 |
let |
26730 | 70 |
val { primary_constraints, secondary_constraints, improve, subst, |
71 |
consider_abbrevs, passed, ... } = ImprovableSyntax.get ctxt; |
|
26238 | 72 |
val tsig = (Sign.tsig_of o ProofContext.theory_of) ctxt; |
26730 | 73 |
val is_abbrev = consider_abbrevs andalso ProofContext.is_abbrev_mode ctxt; |
74 |
val passed_or_abbrev = passed orelse is_abbrev; |
|
26238 | 75 |
fun accumulate_improvements (Const (c, ty)) = (case improve (c, ty) |
26597 | 76 |
of SOME ty_ty' => Type.typ_match tsig ty_ty' |
26238 | 77 |
| _ => I) |
78 |
| accumulate_improvements _ = I; |
|
79 |
val improvements = (fold o fold_aterms) accumulate_improvements ts Vartab.empty; |
|
80 |
val ts' = (map o map_types) (Envir.typ_subst_TVars improvements) ts; |
|
81 |
fun apply_subst t = Envir.expand_term (fn Const (c, ty) => (case subst (c, ty) |
|
26259 | 82 |
of SOME (ty', t') => |
26238 | 83 |
if Type.typ_instance tsig (ty, ty') |
84 |
then SOME (ty', apply_subst t') else NONE |
|
85 |
| NONE => NONE) |
|
26259 | 86 |
| _ => NONE) t; |
26730 | 87 |
val ts'' = if is_abbrev then ts' else map apply_subst ts'; |
88 |
in if eq_list (op aconv) (ts, ts'') andalso passed_or_abbrev then NONE else |
|
89 |
if passed_or_abbrev then SOME (ts'', ctxt) |
|
26238 | 90 |
else SOME (ts'', ctxt |
26520 | 91 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) secondary_constraints |
26238 | 92 |
|> mark_passed) |
93 |
end; |
|
25519 | 94 |
|
26238 | 95 |
fun improve_term_uncheck ts ctxt = |
25519 | 96 |
let |
26238 | 97 |
val thy = ProofContext.theory_of ctxt; |
98 |
val unchecks = (#unchecks o ImprovableSyntax.get) ctxt; |
|
99 |
val ts' = map (Pattern.rewrite_term thy unchecks []) ts; |
|
100 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end; |
|
101 |
||
26520 | 102 |
fun set_primary_constraints ctxt = |
26259 | 103 |
let |
26520 | 104 |
val { primary_constraints, ... } = ImprovableSyntax.get ctxt; |
105 |
in fold (ProofContext.add_const_constraint o apsnd SOME) primary_constraints ctxt end; |
|
26259 | 106 |
|
107 |
val add_improvable_syntax = |
|
108 |
Context.proof_map |
|
26238 | 109 |
(Syntax.add_term_check 0 "improvement" improve_term_check |
110 |
#> Syntax.add_term_uncheck 0 "improvement" improve_term_uncheck) |
|
26520 | 111 |
#> set_primary_constraints; |
26259 | 112 |
|
113 |
||
114 |
(** overloading target **) |
|
115 |
||
116 |
(* bookkeeping *) |
|
117 |
||
118 |
structure OverloadingData = ProofDataFun |
|
119 |
( |
|
120 |
type T = ((string * typ) * (string * bool)) list; |
|
121 |
fun init _ = []; |
|
122 |
); |
|
123 |
||
124 |
val get_overloading = OverloadingData.get o LocalTheory.target_of; |
|
125 |
val map_overloading = LocalTheory.target o OverloadingData.map; |
|
126 |
||
127 |
fun operation lthy v = get_overloading lthy |
|
128 |
|> get_first (fn ((c, _), (v', checked)) => if v = v' then SOME (c, checked) else NONE); |
|
129 |
||
130 |
fun confirm c = map_overloading (filter_out (fn (_, (c', _)) => c' = c)); |
|
131 |
||
132 |
||
133 |
(* overloaded declarations and definitions *) |
|
134 |
||
135 |
fun declare c_ty = pair (Const c_ty); |
|
136 |
||
137 |
fun define checked name (c, t) = |
|
138 |
Thm.add_def (not checked) true (name, Logic.mk_equals (Const (c, Term.fastype_of t), t)); |
|
25519 | 139 |
|
140 |
||
141 |
(* target *) |
|
142 |
||
26238 | 143 |
fun init raw_overloading thy = |
25519 | 144 |
let |
26238 | 145 |
val _ = if null raw_overloading then error "At least one parameter must be given" else (); |
146 |
val overloading = map (fn (v, c_ty, checked) => (c_ty, (v, checked))) raw_overloading; |
|
147 |
fun subst (c, ty) = case AList.lookup (op =) overloading (c, ty) |
|
148 |
of SOME (v, _) => SOME (ty, Free (v, ty)) |
|
149 |
| NONE => NONE; |
|
150 |
val unchecks = |
|
151 |
map (fn (c_ty as (_, ty), (v, _)) => (Free (v, ty), Const c_ty)) overloading; |
|
25519 | 152 |
in |
153 |
thy |
|
154 |
|> ProofContext.init |
|
26238 | 155 |
|> OverloadingData.put overloading |
26259 | 156 |
|> fold (fn ((_, ty), (v, _)) => Variable.declare_names (Free (v, ty))) overloading |
26730 | 157 |
|> map_improvable_syntax (K ((([], []), (((K NONE, subst), false), unchecks)), false)) |
26238 | 158 |
|> add_improvable_syntax |
25519 | 159 |
end; |
160 |
||
161 |
fun conclude lthy = |
|
162 |
let |
|
163 |
val overloading = get_overloading lthy; |
|
164 |
val _ = if null overloading then () else |
|
26259 | 165 |
error ("Missing definition(s) for parameter(s) " ^ commas (map (quote |
25519 | 166 |
o Syntax.string_of_term lthy o Const o fst) overloading)); |
167 |
in |
|
168 |
lthy |
|
169 |
end; |
|
170 |
||
25606 | 171 |
fun pretty lthy = |
172 |
let |
|
173 |
val thy = ProofContext.theory_of lthy; |
|
174 |
val overloading = get_overloading lthy; |
|
175 |
fun pr_operation ((c, ty), (v, _)) = |
|
25861 | 176 |
(Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==", |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26730
diff
changeset
|
177 |
Pretty.str (Sign.extern_const thy c), Pretty.str "::", Syntax.pretty_typ lthy ty]; |
25606 | 178 |
in |
179 |
(Pretty.block o Pretty.fbreaks) |
|
180 |
(Pretty.str "overloading" :: map pr_operation overloading) |
|
181 |
end; |
|
182 |
||
25519 | 183 |
end; |