| author | wenzelm | 
| Mon, 25 Jun 2007 00:36:42 +0200 | |
| changeset 23493 | a056eefb76e5 | 
| parent 22923 | 6384c43da028 | 
| child 24219 | e558fe311376 | 
| permissions | -rw-r--r-- | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/Isar/specification.ML | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 3 | Author: Makarius | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 4 | |
| 21036 | 5 | Derived local theory specifications --- with type-inference and | 
| 18810 | 6 | toplevel polymorphism. | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 7 | *) | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 8 | |
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 9 | signature SPECIFICATION = | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 10 | sig | 
| 20890 | 11 | val quiet_mode: bool ref | 
| 12 | val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit | |
| 18771 | 13 | val read_specification: (string * string option * mixfix) list -> | 
| 18954 | 14 | ((string * Attrib.src list) * string list) list -> local_theory -> | 
| 18771 | 15 | (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) * | 
| 18954 | 16 | local_theory | 
| 18771 | 17 | val cert_specification: (string * typ option * mixfix) list -> | 
| 18954 | 18 | ((string * Attrib.src list) * term list) list -> local_theory -> | 
| 18771 | 19 | (((string * typ) * mixfix) list * ((string * Attrib.src list) * term list) list) * | 
| 18954 | 20 | local_theory | 
| 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 | |
| 27 | val definition: | |
| 21705 | 28 | (string * string option * mixfix) option * ((string * Attrib.src list) * string) -> | 
| 29 | local_theory -> (term * (bstring * thm)) * local_theory | |
| 18954 | 30 | val definition_i: | 
| 21705 | 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 -> | |
| 21532 | 34 | local_theory -> local_theory | 
| 21705 | 35 | val abbreviation_i: Syntax.mode -> (string * typ option * mixfix) option * term -> | 
| 21532 | 36 | local_theory -> local_theory | 
| 21206 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 37 | val notation: Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory | 
| 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 38 | val notation_i: Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory | 
| 20914 | 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 | |
| 21036 | 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 | |
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 49 | end; | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 50 | |
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 51 | structure Specification: SPECIFICATION = | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 52 | struct | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 53 | |
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 54 | |
| 20890 | 55 | (* diagnostics *) | 
| 56 | ||
| 57 | val quiet_mode = ref false; | |
| 58 | ||
| 59 | fun print_consts _ _ [] = () | |
| 60 | | print_consts ctxt pred cs = | |
| 61 | if ! quiet_mode then () else Pretty.writeln (ProofDisplay.pretty_consts ctxt pred cs); | |
| 62 | ||
| 21036 | 63 | fun present_results ctxt res = | 
| 64 | if ! quiet_mode then () else ProofDisplay.present_results ctxt res; | |
| 65 | ||
| 66 | fun present_results' ctxt kind res = present_results ctxt ((kind, ""), res); | |
| 20914 | 67 | |
| 19664 | 68 | |
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 69 | (* prepare specification *) | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 70 | |
| 18828 | 71 | fun prep_specification prep_vars prep_propp prep_att raw_vars raw_specs ctxt = | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 72 | let | 
| 18670 | 73 | val thy = ProofContext.theory_of ctxt; | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 74 | |
| 18670 | 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) = | |
| 19585 | 78 | prep_propp (params_ctxt, map (map (rpair []) o snd) raw_specs) | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 79 | |> swap |>> map (map fst) | 
| 18771 | 80 | ||>> fold_map ProofContext.inferred_param xs; | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 81 | |
| 18771 | 82 | val params = vs ~~ map #3 vars; | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 83 | val names = map (fst o fst) raw_specs; | 
| 18670 | 84 | val atts = map (map (prep_att thy) o snd o fst) raw_specs; | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 85 | in ((params, (names ~~ atts) ~~ specs), specs_ctxt) end; | 
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 86 | |
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 87 | fun read_specification x = | 
| 18771 | 88 | prep_specification ProofContext.read_vars ProofContext.read_propp Attrib.intern_src x; | 
| 21370 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 wenzelm parents: 
21359diff
changeset | 89 | |
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 90 | fun cert_specification x = | 
| 18670 | 91 | prep_specification ProofContext.cert_vars ProofContext.cert_propp (K I) x; | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 92 | |
| 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 93 | |
| 18771 | 94 | (* axiomatization *) | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 95 | |
| 20890 | 96 | fun gen_axioms prep raw_vars raw_specs lthy = | 
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 97 | let | 
| 20890 | 98 | val (vars, specs) = fst (prep raw_vars raw_specs lthy); | 
| 18828 | 99 | val cs = map fst vars; | 
| 18880 
b8a1c3cdf739
axiomatization: retrict parameters to occurrences in specs;
 wenzelm parents: 
18828diff
changeset | 100 | val spec_frees = member (op =) (fold (fold Term.add_frees o snd) specs []); | 
| 18786 | 101 | |
| 20890 | 102 | val ((consts, axioms), lthy') = lthy | 
| 103 | |> LocalTheory.consts spec_frees vars | |
| 21370 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 wenzelm parents: 
21359diff
changeset | 104 | ||> fold (fold Variable.auto_fixes o snd) specs (* FIXME !? *) | 
| 21435 | 105 | ||>> LocalTheory.axioms Thm.axiomK specs; | 
| 18786 | 106 | |
| 20890 | 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); | |
| 110 | ||
| 111 | val _ = print_consts lthy' spec_frees cs; | |
| 112 | in ((map #1 consts, axioms), lthy'') end; | |
| 18786 | 113 | |
| 18954 | 114 | val axiomatization = gen_axioms read_specification; | 
| 115 | val axiomatization_i = gen_axioms cert_specification; | |
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 116 | |
| 18786 | 117 | |
| 118 | (* definition *) | |
| 119 | ||
| 21705 | 120 | fun gen_def prep (raw_var, (raw_a, raw_prop)) lthy = | 
| 18786 | 121 | let | 
| 21705 | 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)] => | |
| 127 | if x = x' then mx | |
| 128 |       else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote x'));
 | |
| 129 | val ((lhs, (_, th)), lthy2) = lthy | |
| 22923 | 130 | |> LocalTheory.def Thm.definitionK ((x, mx), ((name ^ "_raw", []), rhs)); | 
| 21705 | 131 | val ((b, [th']), lthy3) = lthy2 | 
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
21860diff
changeset | 132 | |> LocalTheory.note Thm.definitionK | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
21860diff
changeset | 133 | ((name, CodegenData.add_func_attr false :: atts), [prove lthy2 th]); | 
| 18786 | 134 | |
| 21705 | 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; | |
| 18786 | 138 | |
| 21705 | 139 | val definition = gen_def read_specification; | 
| 140 | val definition_i = gen_def cert_specification; | |
| 18786 | 141 | |
| 19080 | 142 | |
| 143 | (* abbreviation *) | |
| 144 | ||
| 21705 | 145 | fun gen_abbrev prep mode (raw_var, raw_prop) lthy = | 
| 19080 | 146 | let | 
| 21705 | 147 | val ((vars, [(_, [prop])]), _) = | 
| 148 |       prep (the_list raw_var) [(("", []), [raw_prop])]
 | |
| 21729 | 149 | (lthy |> ProofContext.set_expand_abbrevs false); | 
| 21705 | 150 | val ((x, T), rhs) = LocalDefs.abs_def (#2 (LocalDefs.cert_def lthy prop)); | 
| 151 | val mx = (case vars of [] => NoSyn | [((y, _), mx)] => | |
| 152 | if x = y then mx | |
| 153 |       else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y));
 | |
| 154 | val lthy' = lthy | |
| 155 | |> ProofContext.set_syntax_mode mode (* FIXME !? *) | |
| 21793 | 156 | |> LocalTheory.abbrev mode ((x, mx), rhs) |> snd | 
| 21705 | 157 | |> ProofContext.restore_syntax_mode lthy; | 
| 158 | val _ = print_consts lthy' (K false) [(x, T)] | |
| 21532 | 159 | in lthy' end; | 
| 19080 | 160 | |
| 21705 | 161 | val abbreviation = gen_abbrev read_specification; | 
| 162 | val abbreviation_i = gen_abbrev cert_specification; | |
| 19080 | 163 | |
| 19664 | 164 | |
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 165 | (* notation *) | 
| 19664 | 166 | |
| 21206 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 167 | fun gen_notation prep_const mode args lthy = | 
| 21746 | 168 | lthy |> LocalTheory.notation mode (map (apfst (prep_const lthy)) args); | 
| 19664 | 169 | |
| 21206 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 170 | val notation = gen_notation ProofContext.read_const; | 
| 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 171 | val notation_i = gen_notation (K I); | 
| 19664 | 172 | |
| 20914 | 173 | |
| 21036 | 174 | (* fact statements *) | 
| 20914 | 175 | |
| 176 | fun gen_theorems prep_thms prep_att kind raw_facts lthy = | |
| 177 | let | |
| 178 | val attrib = prep_att (ProofContext.theory_of lthy); | |
| 179 | val facts = raw_facts |> map (fn ((name, atts), bs) => | |
| 180 | ((name, map attrib atts), | |
| 21435 | 181 | bs |> map (fn (b, more_atts) => (prep_thms lthy b, map attrib more_atts)))); | 
| 182 | val (res, lthy') = lthy |> LocalTheory.notes kind facts; | |
| 21036 | 183 | val _ = present_results' lthy' kind res; | 
| 20914 | 184 | in (res, lthy') end; | 
| 185 | ||
| 186 | val theorems = gen_theorems ProofContext.get_thms Attrib.intern_src; | |
| 187 | val theorems_i = gen_theorems (K I) (K I); | |
| 188 | ||
| 21036 | 189 | |
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 190 | (* complex goal statements *) | 
| 21036 | 191 | |
| 192 | local | |
| 193 | ||
| 21450 | 194 | fun subtract_prems ctxt1 ctxt2 = | 
| 195 | Library.drop (length (Assumption.prems_of ctxt1), Assumption.prems_of ctxt2); | |
| 196 | ||
| 21236 | 197 | fun prep_statement prep_att prep_stmt elems concl ctxt = | 
| 198 | (case concl of | |
| 199 | Element.Shows shows => | |
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 200 | let | 
| 21450 | 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); | |
| 21370 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 wenzelm parents: 
21359diff
changeset | 204 | val goal_ctxt = fold (fold (Variable.auto_fixes o fst)) propp elems_ctxt; | 
| 21450 | 205 | in ((prems, stmt, []), goal_ctxt) end | 
| 21236 | 206 | | Element.Obtains obtains => | 
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 207 | let | 
| 21236 | 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)))); | |
| 21036 | 213 | |
| 21236 | 214 | val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props); | 
| 21450 | 215 | val (_, loc_ctxt, elems_ctxt, propp) = prep_stmt (elems @ constraints) raw_propp ctxt; | 
| 21236 | 216 | |
| 217 | val thesis = ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) AutoBind.thesisN; | |
| 21036 | 218 | |
| 21236 | 219 | fun assume_case ((name, (vars, _)), asms) ctxt' = | 
| 220 | let | |
| 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)); | |
| 227 | in | |
| 228 | ctxt' |> (snd o ProofContext.add_fixes_i (map (fn x => (x, NONE, NoSyn)) xs)); | |
| 21370 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 wenzelm parents: 
21359diff
changeset | 229 | ctxt' |> Variable.auto_fixes asm | 
| 21236 | 230 | |> ProofContext.add_assms_i Assumption.assume_export | 
| 231 | [((name, [ContextRules.intro_query NONE]), [(asm, [])])] | |
| 232 | |>> (fn [(_, [th])] => th) | |
| 233 | end; | |
| 234 | ||
| 21658 | 235 | val atts = map (Attrib.internal o K) | 
| 21236 | 236 | [RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names]; | 
| 21450 | 237 | val prems = subtract_prems loc_ctxt elems_ctxt; | 
| 21236 | 238 |         val stmt = [(("", atts), [(thesis, [])])];
 | 
| 239 | ||
| 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) | |
| 21435 | 243 | |-> (fn ths => ProofContext.note_thmss_i Thm.assumptionK | 
| 244 | [((Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths); | |
| 21450 | 245 | in ((prems, stmt, facts), goal_ctxt) end); | 
| 21036 | 246 | |
| 247 | fun gen_theorem prep_att prep_stmt | |
| 248 | kind before_qed after_qed (name, raw_atts) raw_elems raw_concl lthy0 = | |
| 249 | let | |
| 21860 | 250 | val _ = LocalTheory.affirm lthy0; | 
| 21036 | 251 | val thy = ProofContext.theory_of lthy0; | 
| 21206 
2af4c7b3f7ef
replaced const_syntax by notation, which operates on terms;
 wenzelm parents: 
21036diff
changeset | 252 | |
| 21036 | 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 *) | |
| 21275 | 256 | (SOME loc, ProofContext.init thy, LocalTheory.restore lthy0) | 
| 21036 | 257 | | _ => (NONE, lthy0, lthy0)); | 
| 258 | ||
| 21435 | 259 | val attrib = prep_att thy; | 
| 260 | val atts = map attrib raw_atts; | |
| 261 | ||
| 21532 | 262 | val elems = raw_elems |> map (Locale.map_elem (Element.map_ctxt_attrib attrib)); | 
| 21450 | 263 | val ((prems, stmt, facts), goal_ctxt) = | 
| 264 | prep_statement attrib (prep_stmt loc) elems raw_concl ctxt; | |
| 21036 | 265 | |
| 266 | fun after_qed' results goal_ctxt' = | |
| 21602 | 267 | let val results' = | 
| 268 | burrow (map Goal.norm_result o ProofContext.export goal_ctxt' lthy) results | |
| 269 | in | |
| 21036 | 270 | lthy | 
| 21435 | 271 | |> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') | 
| 21036 | 272 | |> (fn (res, lthy') => | 
| 273 | (present_results lthy' ((kind, name), res); | |
| 21435 | 274 | if name = "" andalso null atts then lthy' | 
| 275 | else #2 (LocalTheory.notes kind [((name, atts), [(maps #2 res, [])])] lthy'))) | |
| 21036 | 276 | |> after_qed results' | 
| 277 | end; | |
| 278 | in | |
| 279 | goal_ctxt | |
| 21450 | 280 | |> ProofContext.note_thmss_i Thm.assumptionK [((AutoBind.assmsN, []), [(prems, [])])] | 
| 281 | |> snd | |
| 21435 | 282 | |> Proof.theorem_i before_qed after_qed' (map snd stmt) | 
| 21036 | 283 | |> Proof.refine_insert facts | 
| 284 | end; | |
| 285 | ||
| 21230 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 286 | in | 
| 
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
 wenzelm parents: 
21206diff
changeset | 287 | |
| 21036 | 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; | |
| 290 | ||
| 18620 
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
 wenzelm parents: diff
changeset | 291 | end; | 
| 21036 | 292 | |
| 293 | end; |