1 (* Title: Pure/Isar/specification.ML
5 Derived local theory specifications --- with type-inference and
9 signature SPECIFICATION =
11 val quiet_mode: bool ref
12 val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit
13 val read_specification: (string * string option * mixfix) list ->
14 ((string * Attrib.src list) * string list) list -> local_theory ->
15 (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) *
17 val cert_specification: (string * typ option * mixfix) list ->
18 ((string * Attrib.src list) * term list) list -> local_theory ->
19 (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) *
21 val axiomatization: (string * string option * mixfix) list ->
22 ((bstring * Attrib.src list) * string list) list -> local_theory ->
23 (term list * (bstring * thm list) list) * local_theory
24 val axiomatization_i: (string * typ option * mixfix) list ->
25 ((bstring * Attrib.src list) * term list) list -> local_theory ->
26 (term list * (bstring * thm list) list) * local_theory
28 (string * string option * mixfix) option * ((string * Attrib.src list) * string) ->
29 local_theory -> (term * (bstring * thm)) * local_theory
31 (string * typ option * mixfix) option * ((string * Attrib.src list) * term) ->
32 local_theory -> (term * (bstring * thm)) * local_theory
33 val abbreviation: Syntax.mode -> (string * string option * mixfix) option * string ->
34 local_theory -> local_theory
35 val abbreviation_i: Syntax.mode -> (string * typ option * mixfix) option * term ->
36 local_theory -> local_theory
37 val notation: Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
38 val notation_i: Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
39 val theorems: string -> ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list
40 -> local_theory -> (bstring * thm list) list * local_theory
41 val theorems_i: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list
42 -> local_theory -> (bstring * thm list) list * local_theory
43 val theorem: string -> Method.text option -> (thm list list -> local_theory -> local_theory) ->
44 string * Attrib.src list -> Element.context Locale.element list -> Element.statement ->
45 local_theory -> Proof.state
46 val theorem_i: string -> Method.text option -> (thm list list -> local_theory -> local_theory) ->
47 string * Attrib.src list -> Element.context_i Locale.element list -> Element.statement_i ->
48 local_theory -> Proof.state
51 structure Specification: SPECIFICATION =
57 val quiet_mode = ref false;
59 fun print_consts _ _ [] = ()
60 | print_consts ctxt pred cs =
61 if ! quiet_mode then () else Pretty.writeln (ProofDisplay.pretty_consts ctxt pred cs);
63 fun present_results ctxt res =
64 if ! quiet_mode then () else ProofDisplay.present_results ctxt res;
66 fun present_results' ctxt kind res = present_results ctxt ((kind, ""), res);
69 (* prepare specification *)
71 fun prep_specification prep_vars prep_propp prep_att raw_vars raw_specs ctxt =
73 val thy = ProofContext.theory_of ctxt;
75 val (vars, vars_ctxt) = ctxt |> prep_vars raw_vars;
76 val (xs, params_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars;
77 val ((specs, vs), specs_ctxt) =
78 prep_propp (params_ctxt, map (map (rpair []) o snd) raw_specs)
79 |> swap |>> map (map fst)
80 ||>> fold_map ProofContext.inferred_param xs;
82 val params = vs ~~ map #3 vars;
83 val names = map (fst o fst) raw_specs;
84 val atts = map (map (prep_att thy) o snd o fst) raw_specs;
85 in ((params, (names ~~ atts) ~~ specs), specs_ctxt) end;
87 fun read_specification x =
88 prep_specification ProofContext.read_vars ProofContext.read_propp Attrib.intern_src x;
90 fun cert_specification x =
91 prep_specification ProofContext.cert_vars ProofContext.cert_propp (K I) x;
96 fun gen_axioms prep raw_vars raw_specs lthy =
98 val (vars, specs) = fst (prep raw_vars raw_specs lthy);
99 val cs = map fst vars;
100 val spec_frees = member (op =) (fold (fold Term.add_frees o snd) specs []);
102 val ((consts, axioms), lthy') = lthy
103 |> LocalTheory.consts spec_frees vars
104 ||> fold (fold Variable.auto_fixes o snd) specs (* FIXME !? *)
105 ||>> LocalTheory.axioms Thm.axiomK specs;
107 (* FIXME generic target!? *)
108 val hs = map (Term.head_of o #2 o Logic.dest_equals o Thm.prop_of o #2) consts;
109 val lthy'' = lthy' |> LocalTheory.theory (Theory.add_finals_i false hs);
111 val _ = print_consts lthy' spec_frees cs;
112 in ((map #1 consts, axioms), lthy'') end;
114 val axiomatization = gen_axioms read_specification;
115 val axiomatization_i = gen_axioms cert_specification;
120 fun gen_def prep (raw_var, (raw_a, raw_prop)) lthy =
122 val (vars, [((raw_name, atts), [prop])]) =
123 fst (prep (the_list raw_var) [(raw_a, [raw_prop])] lthy);
124 val (((x, T), rhs), prove) = LocalDefs.derived_def lthy true prop;
125 val name = Thm.def_name_optional x raw_name;
126 val mx = (case vars of [] => NoSyn | [((x', _), mx)] =>
128 else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote x'));
129 val ((lhs, (_, th)), lthy2) = lthy
130 (* |> LocalTheory.def ((x, mx), ((name ^ "_raw", []), rhs)); FIXME *)
131 |> LocalTheory.def Thm.definitionK ((x, mx), ((name, []), rhs));
132 val ((b, [th']), lthy3) = lthy2
133 |> LocalTheory.note Thm.definitionK ((name, atts), [prove lthy2 th]);
135 val lhs' = Morphism.term (LocalTheory.target_morphism lthy3) lhs;
136 val _ = print_consts lthy3 (member (op =) (Term.add_frees lhs' [])) [(x, T)];
137 in ((lhs, (b, th')), lthy3) end;
139 val definition = gen_def read_specification;
140 val definition_i = gen_def cert_specification;
145 fun gen_abbrev prep mode (raw_var, raw_prop) lthy =
147 val ((vars, [(_, [prop])]), _) =
148 prep (the_list raw_var) [(("", []), [raw_prop])]
149 (lthy |> ProofContext.set_expand_abbrevs false);
150 val ((x, T), rhs) = LocalDefs.abs_def (#2 (LocalDefs.cert_def lthy prop));
151 val mx = (case vars of [] => NoSyn | [((y, _), mx)] =>
153 else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y));
155 |> ProofContext.set_syntax_mode mode (* FIXME !? *)
156 |> LocalTheory.abbrev mode ((x, mx), rhs) |> snd
157 |> ProofContext.restore_syntax_mode lthy;
158 val _ = print_consts lthy' (K false) [(x, T)]
161 val abbreviation = gen_abbrev read_specification;
162 val abbreviation_i = gen_abbrev cert_specification;
167 fun gen_notation prep_const mode args lthy =
168 lthy |> LocalTheory.notation mode (map (apfst (prep_const lthy)) args);
170 val notation = gen_notation ProofContext.read_const;
171 val notation_i = gen_notation (K I);
174 (* fact statements *)
176 fun gen_theorems prep_thms prep_att kind raw_facts lthy =
178 val attrib = prep_att (ProofContext.theory_of lthy);
179 val facts = raw_facts |> map (fn ((name, atts), bs) =>
180 ((name, map attrib atts),
181 bs |> map (fn (b, more_atts) => (prep_thms lthy b, map attrib more_atts))));
182 val (res, lthy') = lthy |> LocalTheory.notes kind facts;
183 val _ = present_results' lthy' kind res;
186 val theorems = gen_theorems ProofContext.get_thms Attrib.intern_src;
187 val theorems_i = gen_theorems (K I) (K I);
190 (* complex goal statements *)
194 fun subtract_prems ctxt1 ctxt2 =
195 Library.drop (length (Assumption.prems_of ctxt1), Assumption.prems_of ctxt2);
197 fun prep_statement prep_att prep_stmt elems concl ctxt =
199 Element.Shows shows =>
201 val (_, loc_ctxt, elems_ctxt, propp) = prep_stmt elems (map snd shows) ctxt;
202 val prems = subtract_prems loc_ctxt elems_ctxt;
203 val stmt = Attrib.map_specs prep_att (map fst shows ~~ propp);
204 val goal_ctxt = fold (fold (Variable.auto_fixes o fst)) propp elems_ctxt;
205 in ((prems, stmt, []), goal_ctxt) end
206 | Element.Obtains obtains =>
208 val case_names = obtains |> map_index
209 (fn (i, ("", _)) => string_of_int (i + 1) | (_, (name, _)) => name);
210 val constraints = obtains |> map (fn (_, (vars, _)) =>
211 Locale.Elem (Element.Constrains
212 (vars |> map_filter (fn (x, SOME T) => SOME (x, T) | _ => NONE))));
214 val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props);
215 val (_, loc_ctxt, elems_ctxt, propp) = prep_stmt (elems @ constraints) raw_propp ctxt;
217 val thesis = ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) AutoBind.thesisN;
219 fun assume_case ((name, (vars, _)), asms) ctxt' =
221 val xs = map fst vars;
222 val props = map fst asms;
223 val (parms, _) = ctxt'
224 |> fold Variable.declare_term props
225 |> fold_map ProofContext.inferred_param xs;
226 val asm = Term.list_all_free (parms, Logic.list_implies (props, thesis));
228 ctxt' |> (snd o ProofContext.add_fixes_i (map (fn x => (x, NONE, NoSyn)) xs));
229 ctxt' |> Variable.auto_fixes asm
230 |> ProofContext.add_assms_i Assumption.assume_export
231 [((name, [ContextRules.intro_query NONE]), [(asm, [])])]
232 |>> (fn [(_, [th])] => th)
235 val atts = map (Attrib.internal o K)
236 [RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names];
237 val prems = subtract_prems loc_ctxt elems_ctxt;
238 val stmt = [(("", atts), [(thesis, [])])];
240 val (facts, goal_ctxt) = elems_ctxt
241 |> (snd o ProofContext.add_fixes_i [(AutoBind.thesisN, NONE, NoSyn)])
242 |> fold_map assume_case (obtains ~~ propp)
243 |-> (fn ths => ProofContext.note_thmss_i Thm.assumptionK
244 [((Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths);
245 in ((prems, stmt, facts), goal_ctxt) end);
247 fun gen_theorem prep_att prep_stmt
248 kind before_qed after_qed (name, raw_atts) raw_elems raw_concl lthy0 =
250 val _ = LocalTheory.assert lthy0;
251 val thy = ProofContext.theory_of lthy0;
253 val (loc, ctxt, lthy) =
254 (case (TheoryTarget.peek lthy0, exists (fn Locale.Expr _ => true | _ => false) raw_elems) of
255 (SOME loc, true) => (* workaround non-modularity of in/includes *) (* FIXME *)
256 (SOME loc, ProofContext.init thy, LocalTheory.restore lthy0)
257 | _ => (NONE, lthy0, lthy0));
259 val attrib = prep_att thy;
260 val atts = map attrib raw_atts;
262 val elems = raw_elems |> map (Locale.map_elem (Element.map_ctxt_attrib attrib));
263 val ((prems, stmt, facts), goal_ctxt) =
264 prep_statement attrib (prep_stmt loc) elems raw_concl ctxt;
266 fun after_qed' results goal_ctxt' =
268 burrow (map Goal.norm_result o ProofContext.export goal_ctxt' lthy) results
271 |> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results')
272 |> (fn (res, lthy') =>
273 (present_results lthy' ((kind, name), res);
274 if name = "" andalso null atts then lthy'
275 else #2 (LocalTheory.notes kind [((name, atts), [(maps #2 res, [])])] lthy')))
276 |> after_qed results'
280 |> ProofContext.note_thmss_i Thm.assumptionK [((AutoBind.assmsN, []), [(prems, [])])]
282 |> Proof.theorem_i before_qed after_qed' (map snd stmt)
283 |> Proof.refine_insert facts
288 val theorem = gen_theorem Attrib.intern_src Locale.read_context_statement_i;
289 val theorem_i = gen_theorem (K I) Locale.cert_context_statement;