author | haftmann |
Wed, 17 Feb 2016 21:51:57 +0100 | |
changeset 62348 | 9a5f43dac883 |
parent 61950 | a2d4742b127f |
child 62770 | 6e6cacf8fe50 |
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 |
Author: Makarius |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
3 |
|
21036 | 4 |
Derived local theory specifications --- with type-inference and |
18810 | 5 |
toplevel polymorphism. |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
6 |
*) |
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 |
signature SPECIFICATION = |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
9 |
sig |
59844
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
10 |
val read_props: string list -> (binding * string option * mixfix) list -> Proof.context -> |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
11 |
term list * Proof.context |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
12 |
val read_prop: string -> (binding * string option * mixfix) list -> Proof.context -> |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
13 |
term * Proof.context |
61947 | 14 |
val check_free_spec: |
15 |
(binding * typ option * mixfix) list -> (Attrib.binding * term) list -> Proof.context -> |
|
16 |
((((binding * typ) * mixfix) list * (Attrib.binding * term) list) * (string -> Position.T)) |
|
17 |
* Proof.context |
|
18 |
val read_free_spec: |
|
19 |
(binding * string option * mixfix) list -> (Attrib.binding * string) list -> Proof.context -> |
|
20 |
((((binding * typ) * mixfix) list * (Attrib.binding * term) list) * (string -> Position.T)) |
|
21 |
* Proof.context |
|
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
22 |
val check_spec: |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
23 |
(binding * typ option * mixfix) list -> (Attrib.binding * term) list -> Proof.context -> |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
24 |
(((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
25 |
val read_spec: |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
26 |
(binding * string option * mixfix) list -> (Attrib.binding * string) list -> Proof.context -> |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
27 |
(((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context |
29581 | 28 |
val check_specification: (binding * typ option * mixfix) list -> |
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
29 |
(Attrib.binding * term list) list -> Proof.context -> |
29581 | 30 |
(((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
31 |
val read_specification: (binding * string option * mixfix) list -> |
|
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
32 |
(Attrib.binding * string list) list -> Proof.context -> |
29581 | 33 |
(((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
34 |
val axiomatization: (binding * typ option * mixfix) list -> |
|
28114 | 35 |
(Attrib.binding * term list) list -> theory -> |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
36 |
(term list * thm list list) * theory |
29581 | 37 |
val axiomatization_cmd: (binding * string option * mixfix) list -> |
28114 | 38 |
(Attrib.binding * string list) list -> theory -> |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
39 |
(term list * thm list list) * theory |
35894 | 40 |
val axiom: Attrib.binding * term -> theory -> thm * theory |
18954 | 41 |
val definition: |
29581 | 42 |
(binding * typ option * mixfix) option * (Attrib.binding * term) -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
43 |
local_theory -> (term * (string * thm)) * local_theory |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
44 |
val definition': |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
45 |
(binding * typ option * mixfix) option * (Attrib.binding * term) -> |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
46 |
bool -> local_theory -> (term * (string * thm)) * local_theory |
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
47 |
val definition_cmd: |
29581 | 48 |
(binding * string option * mixfix) option * (Attrib.binding * string) -> |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
49 |
bool -> local_theory -> (term * (string * thm)) * local_theory |
29581 | 50 |
val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option * term -> |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
51 |
bool -> local_theory -> local_theory |
29581 | 52 |
val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option * string -> |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
53 |
bool -> local_theory -> local_theory |
35413 | 54 |
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory |
55639 | 55 |
val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> |
56 |
local_theory -> local_theory |
|
24949 | 57 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory |
58 |
val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
59 |
val theorems: string -> |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
56932
diff
changeset
|
60 |
(Attrib.binding * (thm list * Token.src list) list) list -> |
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
61 |
(binding * typ option * mixfix) list -> |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
62 |
bool -> local_theory -> (string * thm list) list * local_theory |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
25371
diff
changeset
|
63 |
val theorems_cmd: string -> |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
56932
diff
changeset
|
64 |
(Attrib.binding * (Facts.ref * Token.src list) list) list -> |
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
65 |
(binding * string option * mixfix) list -> |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
66 |
bool -> local_theory -> (string * thm list) list * local_theory |
24986 | 67 |
val theorem: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
68 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
69 |
string list -> Element.context_i list -> Element.statement_i -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
70 |
bool -> local_theory -> Proof.state |
24986 | 71 |
val theorem_cmd: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
72 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
73 |
(xstring * Position.T) list -> Element.context list -> Element.statement -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
74 |
bool -> local_theory -> Proof.state |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
75 |
val schematic_theorem: string -> Method.text option -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
76 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
77 |
string list -> Element.context_i list -> Element.statement_i -> |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
78 |
bool -> local_theory -> Proof.state |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
79 |
val schematic_theorem_cmd: string -> Method.text option -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
80 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
81 |
(xstring * Position.T) list -> Element.context list -> Element.statement -> |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
82 |
bool -> local_theory -> Proof.state |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
83 |
end; |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
84 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
85 |
structure Specification: SPECIFICATION = |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
86 |
struct |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
87 |
|
59844
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
88 |
(* prepare propositions *) |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
89 |
|
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
90 |
fun read_props raw_props raw_fixes ctxt = |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
91 |
let |
60469 | 92 |
val (_, ctxt1) = ctxt |> Proof_Context.add_fixes_cmd raw_fixes; |
59844
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
93 |
val props1 = map (Syntax.parse_prop ctxt1) raw_props; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
94 |
val (props2, ctxt2) = ctxt1 |> fold_map Variable.fix_dummy_patterns props1; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
95 |
val props3 = Syntax.check_props ctxt2 props2; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
96 |
val ctxt3 = ctxt2 |> fold Variable.declare_term props3; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
97 |
in (props3, ctxt3) end; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
98 |
|
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
99 |
fun read_prop raw_prop raw_fixes ctxt = |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
100 |
let val ([prop], ctxt') = read_props [raw_prop] raw_fixes ctxt |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
101 |
in (prop, ctxt') end; |
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
102 |
|
c648efffea73
support for minimal specifications, with usual treatment of fixes and dummies;
wenzelm
parents:
58993
diff
changeset
|
103 |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
104 |
(* prepare specification *) |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
105 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
106 |
local |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
107 |
|
60445 | 108 |
fun close_forms ctxt i As = |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
109 |
let |
60454
a4c6b278f3a7
eliminated slightly odd Element.close_form: toplevel specifications have different policies than proof text elements;
wenzelm
parents:
60448
diff
changeset
|
110 |
val xs = rev (fold (Variable.add_free_names ctxt) As []); |
55639 | 111 |
val types = |
60454
a4c6b278f3a7
eliminated slightly odd Element.close_form: toplevel specifications have different policies than proof text elements;
wenzelm
parents:
60448
diff
changeset
|
112 |
map (Type_Infer.param i o rpair []) (Name.invent Name.context Name.aT (length xs)); |
a4c6b278f3a7
eliminated slightly odd Element.close_form: toplevel specifications have different policies than proof text elements;
wenzelm
parents:
60448
diff
changeset
|
113 |
val uniform_typing = AList.lookup (op =) (xs ~~ types); |
a4c6b278f3a7
eliminated slightly odd Element.close_form: toplevel specifications have different policies than proof text elements;
wenzelm
parents:
60448
diff
changeset
|
114 |
val close = fold_rev (Logic.dependent_all_constraint uniform_typing) (xs ~~ xs); |
a4c6b278f3a7
eliminated slightly odd Element.close_form: toplevel specifications have different policies than proof text elements;
wenzelm
parents:
60448
diff
changeset
|
115 |
in map close As end; |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
116 |
|
61947 | 117 |
fun get_positions ctxt x = |
118 |
let |
|
119 |
fun get Cs (Const ("_type_constraint_", C) $ t) = get (C :: Cs) t |
|
120 |
| get _ (t $ u) = get [] t @ get [] u |
|
121 |
| get _ (Abs (_, _, t)) = get [] t |
|
122 |
| get Cs (Free (y, T)) = |
|
123 |
if x = y then |
|
124 |
map_filter Term_Position.decode_positionT |
|
125 |
(T :: map (Type.constraint_type ctxt) Cs) |
|
126 |
else [] |
|
127 |
| get _ _ = []; |
|
128 |
in get [] end; |
|
129 |
||
60379 | 130 |
fun prepare prep_var parse_prop prep_att do_close raw_vars raw_specss ctxt = |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
131 |
let |
60379 | 132 |
val (vars, vars_ctxt) = ctxt |> fold_map prep_var raw_vars; |
42360 | 133 |
val (xs, params_ctxt) = vars_ctxt |> Proof_Context.add_fixes vars; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
134 |
|
42265 | 135 |
val Asss = |
136 |
(map o map) snd raw_specss |
|
58993
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
wenzelm
parents:
58848
diff
changeset
|
137 |
|> (burrow o burrow) (grouped 10 Par_List.map_independent (parse_prop params_ctxt)); |
24734 | 138 |
val names = Variable.names_of (params_ctxt |> (fold o fold o fold) Variable.declare_term Asss) |
139 |
|> fold Name.declare xs; |
|
140 |
val Asss' = #1 ((fold_map o fold_map o fold_map) Term.free_dummy_patterns Asss names); |
|
141 |
val idx = (fold o fold o fold) Term.maxidx_term Asss' ~1 + 1; |
|
61947 | 142 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
143 |
val specs = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
144 |
(if do_close then |
24734 | 145 |
#1 (fold_map |
60445 | 146 |
(fn Ass => fn i => (burrow (close_forms params_ctxt i) Ass, i + 1)) Asss' idx) |
24734 | 147 |
else Asss') |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
148 |
|> flat |> burrow (Syntax.check_props params_ctxt); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
149 |
val specs_ctxt = params_ctxt |> (fold o fold) Variable.declare_term specs; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
150 |
|
60407 | 151 |
val ps = specs_ctxt |> fold_map Proof_Context.inferred_param xs |> fst; |
152 |
val params = map2 (fn (b, _, mx) => fn (_, T) => ((b, T), mx)) vars ps; |
|
61947 | 153 |
val name_atts: Attrib.binding list = |
154 |
map (fn ((name, atts), _) => (name, map (prep_att ctxt) atts)) (flat raw_specss); |
|
155 |
||
156 |
fun get_pos x = |
|
157 |
if do_close then Position.none |
|
158 |
else |
|
159 |
(case (maps o maps o maps) (get_positions specs_ctxt x) Asss' of |
|
160 |
[] => Position.none |
|
161 |
| pos :: _ => pos); |
|
162 |
in (((params, name_atts ~~ specs), get_pos), specs_ctxt) end; |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
163 |
|
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
164 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
165 |
fun single_spec (a, prop) = [(a, [prop])]; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
166 |
fun the_spec (a, [prop]) = (a, prop); |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
167 |
|
60379 | 168 |
fun prep_spec prep_var parse_prop prep_att do_close vars specs = |
169 |
prepare prep_var parse_prop prep_att do_close |
|
61947 | 170 |
vars (map single_spec specs) #>> (apfst o apsnd) (map the_spec); |
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
171 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
172 |
in |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21359
diff
changeset
|
173 |
|
61947 | 174 |
fun check_free_spec vars specs = |
175 |
prep_spec Proof_Context.cert_var (K I) (K I) false vars specs; |
|
176 |
||
177 |
fun read_free_spec vars specs = |
|
178 |
prep_spec Proof_Context.read_var Syntax.parse_prop Attrib.check_src false vars specs; |
|
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
179 |
|
61947 | 180 |
fun check_spec vars specs = |
181 |
prep_spec Proof_Context.cert_var (K I) (K I) true vars specs #> apfst fst; |
|
182 |
||
183 |
fun read_spec vars specs = |
|
184 |
prep_spec Proof_Context.read_var Syntax.parse_prop Attrib.check_src true vars specs #> apfst fst; |
|
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
185 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
186 |
fun check_specification vars specs = |
61947 | 187 |
prepare Proof_Context.cert_var (K I) (K I) true vars [specs] #> apfst fst |
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
188 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
189 |
fun read_specification vars specs = |
61947 | 190 |
prepare Proof_Context.read_var Syntax.parse_prop Attrib.check_src true vars [specs] #> apfst fst; |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
191 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
192 |
end; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
193 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
194 |
|
28114 | 195 |
(* axiomatization -- within global theory *) |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
196 |
|
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
197 |
fun gen_axioms prep raw_vars raw_specs thy = |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
198 |
let |
42360 | 199 |
val ((vars, specs), _) = prep raw_vars raw_specs (Proof_Context.init_global thy); |
42494 | 200 |
val xs = map (fn ((b, T), _) => (Variable.check_name b, T)) vars; |
28114 | 201 |
|
202 |
(*consts*) |
|
33173
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents:
32856
diff
changeset
|
203 |
val (consts, consts_thy) = thy |> fold_map Theory.specify_const vars; |
28114 | 204 |
val subst = Term.subst_atomic (map Free xs ~~ consts); |
205 |
||
206 |
(*axioms*) |
|
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
207 |
val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), props) => |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42360
diff
changeset
|
208 |
fold_map Thm.add_axiom_global |
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
209 |
(map (apfst (fn a => Binding.map_name (K a) b)) |
47021
f35f654f297d
clarified Binding.name_of vs Name_Space.base_name vs Variable.check_name (see also 9bd8d4addd6e, 3305f573294e);
wenzelm
parents:
46992
diff
changeset
|
210 |
(Global_Theory.name_multi (Binding.name_of b) (map subst props))) |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35894
diff
changeset
|
211 |
#>> (fn ths => ((b, atts), [(map #2 ths, [])]))); |
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
212 |
|
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
213 |
(*facts*) |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
214 |
val (facts, facts_lthy) = axioms_thy |
38388 | 215 |
|> Named_Target.theory_init |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
216 |
|> Spec_Rules.add Spec_Rules.Unknown (consts, maps (maps #1 o #2) axioms) |
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
217 |
|> Local_Theory.notes axioms; |
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
218 |
|
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
219 |
in ((consts, map #2 facts), Local_Theory.exit_global facts_lthy) end; |
18786 | 220 |
|
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
221 |
val axiomatization = gen_axioms check_specification; |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
222 |
val axiomatization_cmd = gen_axioms read_specification; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
223 |
|
35894 | 224 |
fun axiom (b, ax) = axiomatization [] [(b, [ax])] #>> (hd o hd o snd); |
225 |
||
18786 | 226 |
|
227 |
(* definition *) |
|
228 |
||
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
229 |
fun gen_def prep (raw_var, raw_spec) int lthy = |
18786 | 230 |
let |
61947 | 231 |
val ((vars, [((raw_name, atts), prop)]), get_pos) = |
232 |
fst (prep (the_list raw_var) [raw_spec] lthy); |
|
35624 | 233 |
val (((x, T), rhs), prove) = Local_Defs.derived_def lthy true prop; |
56169
9b0dc5c704c9
reject internal names, notably from Term.free_dummy_patterns;
wenzelm
parents:
56026
diff
changeset
|
234 |
val _ = Name.reject_internal (x, []); |
30434 | 235 |
val var as (b, _) = |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
236 |
(case vars of |
61947 | 237 |
[] => (Binding.make (x, get_pos x), NoSyn) |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
238 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
239 |
let |
42494 | 240 |
val y = Variable.check_name b; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
241 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
242 |
error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ |
48992 | 243 |
Position.here (Binding.pos_of b)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
244 |
in (b, mx) end); |
61950 | 245 |
val name = Binding.reset_pos (Thm.def_binding_optional b raw_name); |
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
246 |
val ((lhs, (_, raw_th)), lthy2) = lthy |
46992
eeea81b86b70
refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents:
46989
diff
changeset
|
247 |
|> Local_Theory.define_internal (var, ((Binding.suffix_name "_raw" name, []), rhs)); |
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
248 |
|
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
249 |
val th = prove lthy2 raw_th; |
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
250 |
val lthy3 = lthy2 |> Spec_Rules.add Spec_Rules.Equational ([lhs], [th]); |
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
251 |
|
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33644
diff
changeset
|
252 |
val ([(def_name, [th'])], lthy4) = lthy3 |
47080 | 253 |
|> Local_Theory.notes [((name, Code.add_default_eqn_attrib :: atts), [([th], [])])]; |
18786 | 254 |
|
33671 | 255 |
val lhs' = Morphism.term (Local_Theory.target_morphism lthy4) lhs; |
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
256 |
|
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
257 |
val _ = |
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
258 |
Proof_Display.print_consts int (Position.thread_data ()) lthy4 |
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
259 |
(member (op =) (Term.add_frees lhs' [])) [(x, T)]; |
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
260 |
in ((lhs, (def_name, th')), lthy4) end; |
18786 | 261 |
|
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
262 |
val definition' = gen_def check_free_spec; |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
263 |
fun definition spec = definition' spec false; |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
264 |
val definition_cmd = gen_def read_free_spec; |
18786 | 265 |
|
19080 | 266 |
|
267 |
(* abbreviation *) |
|
268 |
||
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
269 |
fun gen_abbrev prep mode (raw_var, raw_prop) int lthy = |
19080 | 270 |
let |
55119
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
271 |
val lthy1 = lthy |
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
272 |
|> Proof_Context.set_syntax_mode mode; |
61947 | 273 |
val (((vars, [(_, prop)]), get_pos), _) = |
30482
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
274 |
prep (the_list raw_var) [(Attrib.empty_binding, raw_prop)] |
55119
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
275 |
(lthy1 |> Proof_Context.set_mode Proof_Context.mode_abbrev); |
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
276 |
val ((x, T), rhs) = Local_Defs.abs_def (#2 (Local_Defs.cert_def lthy1 prop)); |
56169
9b0dc5c704c9
reject internal names, notably from Term.free_dummy_patterns;
wenzelm
parents:
56026
diff
changeset
|
277 |
val _ = Name.reject_internal (x, []); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
278 |
val var = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
279 |
(case vars of |
61947 | 280 |
[] => (Binding.make (x, get_pos x), NoSyn) |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
281 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
282 |
let |
42494 | 283 |
val y = Variable.check_name b; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
284 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
285 |
error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ |
48992 | 286 |
Position.here (Binding.pos_of b)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
287 |
in (b, mx) end); |
55119
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
288 |
val lthy2 = lthy1 |
33671 | 289 |
|> Local_Theory.abbrev mode (var, rhs) |> snd |
42360 | 290 |
|> Proof_Context.restore_syntax_mode lthy; |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
291 |
|
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
292 |
val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy2 (K false) [(x, T)]; |
55119
9ca72949ebac
observe local syntax mode (according to e3a39dae2004, which was lost in 0f3ad56548bc), e.g. relevant for "abbreviation (output)" with non-terminating syntax;
wenzelm
parents:
54883
diff
changeset
|
293 |
in lthy2 end; |
19080 | 294 |
|
44192
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
295 |
val abbreviation = gen_abbrev check_free_spec; |
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents:
43329
diff
changeset
|
296 |
val abbreviation_cmd = gen_abbrev read_free_spec; |
19080 | 297 |
|
19664 | 298 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
299 |
(* notation *) |
19664 | 300 |
|
35413 | 301 |
local |
302 |
||
303 |
fun gen_type_notation prep_type add mode args lthy = |
|
304 |
lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args); |
|
305 |
||
24949 | 306 |
fun gen_notation prep_const add mode args lthy = |
33671 | 307 |
lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); |
19664 | 308 |
|
35413 | 309 |
in |
310 |
||
311 |
val type_notation = gen_type_notation (K I); |
|
55951
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents:
55639
diff
changeset
|
312 |
val type_notation_cmd = |
56002 | 313 |
gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); |
35413 | 314 |
|
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
315 |
val notation = gen_notation (K I); |
56002 | 316 |
val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); |
19664 | 317 |
|
35413 | 318 |
end; |
319 |
||
20914 | 320 |
|
21036 | 321 |
(* fact statements *) |
20914 | 322 |
|
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
323 |
local |
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
324 |
|
60469 | 325 |
fun gen_theorems prep_fact prep_att add_fixes |
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
326 |
kind raw_facts raw_fixes int lthy = |
20914 | 327 |
let |
328 |
val facts = raw_facts |> map (fn ((name, atts), bs) => |
|
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55955
diff
changeset
|
329 |
((name, map (prep_att lthy) atts), |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55955
diff
changeset
|
330 |
bs |> map (fn (b, more_atts) => (prep_fact lthy b, map (prep_att lthy) more_atts)))); |
60469 | 331 |
val (_, ctxt') = add_fixes raw_fixes lthy; |
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
332 |
|
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
333 |
val facts' = facts |
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
334 |
|> Attrib.partial_evaluation ctxt' |
58028
e4250d370657
tuned signature -- define some elementary operations earlier;
wenzelm
parents:
58011
diff
changeset
|
335 |
|> Attrib.transform_facts (Proof_Context.export_morphism ctxt' lthy); |
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
336 |
val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts'; |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
337 |
val _ = Proof_Display.print_results int (Position.thread_data ()) lthy' ((kind, ""), res); |
20914 | 338 |
in (res, lthy') end; |
339 |
||
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
340 |
in |
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
341 |
|
60469 | 342 |
val theorems = gen_theorems (K I) (K I) Proof_Context.add_fixes; |
343 |
val theorems_cmd = gen_theorems Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; |
|
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
344 |
|
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
45584
diff
changeset
|
345 |
end; |
20914 | 346 |
|
21036 | 347 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
348 |
(* complex goal statements *) |
21036 | 349 |
|
350 |
local |
|
351 |
||
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
352 |
fun prep_statement prep_att prep_stmt raw_elems raw_stmt ctxt = |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
353 |
let |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
354 |
val (stmt, elems_ctxt) = prep_stmt raw_elems raw_stmt ctxt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
355 |
val prems = Assumption.local_prems_of elems_ctxt ctxt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
356 |
val stmt_ctxt = fold (fold (Variable.auto_fixes o fst) o snd) stmt elems_ctxt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
357 |
in |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
358 |
(case raw_stmt of |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
359 |
Element.Shows _ => |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
360 |
let val stmt' = Attrib.map_specs (map prep_att) stmt |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
361 |
in (([], prems, stmt', NONE), stmt_ctxt) end |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
362 |
| Element.Obtains raw_obtains => |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
363 |
let |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
364 |
val asms_ctxt = stmt_ctxt |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
365 |
|> fold (fn ((name, _), asm) => |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
366 |
snd o Proof_Context.add_assms Assumption.assume_export |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
367 |
[((name, [Context_Rules.intro_query NONE]), asm)]) stmt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
368 |
val that = Assumption.local_prems_of asms_ctxt stmt_ctxt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
369 |
val ([(_, that')], that_ctxt) = asms_ctxt |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
370 |
|> Proof_Context.note_thmss "" [((Binding.name Auto_Bind.thatN, []), [(that, [])])]; |
21036 | 371 |
|
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
372 |
val more_atts = map (Attrib.internal o K) (Obtain.obtains_attributes raw_obtains); |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
373 |
val stmt' = [((Binding.empty, []), [(#2 (#1 (Obtain.obtain_thesis ctxt)), [])])]; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
374 |
in ((more_atts, prems, stmt', SOME that'), that_ctxt) end) |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60469
diff
changeset
|
375 |
end; |
21036 | 376 |
|
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
377 |
fun gen_theorem schematic bundle_includes prep_att prep_stmt |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
378 |
kind before_qed after_qed (name, raw_atts) raw_includes raw_elems raw_concl int lthy = |
21036 | 379 |
let |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47021
diff
changeset
|
380 |
val _ = Local_Theory.assert lthy; |
21036 | 381 |
|
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55955
diff
changeset
|
382 |
val elems = raw_elems |> map (Element.map_ctxt_attrib (prep_att lthy)); |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
383 |
val ((more_atts, prems, stmt, facts), goal_ctxt) = lthy |
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
384 |
|> bundle_includes raw_includes |
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55955
diff
changeset
|
385 |
|> prep_statement (prep_att lthy) prep_stmt elems raw_concl; |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55955
diff
changeset
|
386 |
val atts = more_atts @ map (prep_att lthy) raw_atts; |
21036 | 387 |
|
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
388 |
val pos = Position.thread_data (); |
21036 | 389 |
fun after_qed' results goal_ctxt' = |
45584
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
390 |
let |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
51313
diff
changeset
|
391 |
val results' = |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
51313
diff
changeset
|
392 |
burrow (map (Goal.norm_result lthy) o Proof_Context.export goal_ctxt' lthy) results; |
45584
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
393 |
val (res, lthy') = |
49062
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents:
49059
diff
changeset
|
394 |
if forall (Attrib.is_empty_binding o fst) stmt then (map (pair "") results', lthy) |
45584
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
395 |
else |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
396 |
Local_Theory.notes_kind kind |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
397 |
(map2 (fn (b, _) => fn ths => (b, [(ths, [])])) stmt results') lthy; |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
398 |
val lthy'' = |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
399 |
if Attrib.is_empty_binding (name, atts) then |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
400 |
(Proof_Display.print_results int pos lthy' ((kind, ""), res); lthy') |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
401 |
else |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
402 |
let |
45584
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
403 |
val ([(res_name, _)], lthy'') = |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
404 |
Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])] lthy'; |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
405 |
val _ = Proof_Display.print_results int pos lthy' ((kind, res_name), res); |
45584
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
406 |
in lthy'' end; |
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
wenzelm
parents:
45583
diff
changeset
|
407 |
in after_qed results' lthy'' end; |
21036 | 408 |
in |
409 |
goal_ctxt |
|
42360 | 410 |
|> Proof_Context.note_thmss "" [((Binding.name Auto_Bind.assmsN, []), [(prems, [])])] |
21450 | 411 |
|> snd |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36317
diff
changeset
|
412 |
|> Proof.theorem before_qed after_qed' (map snd stmt) |
32856
92d9555ac790
clarified Proof.refine_insert -- always "refine" to apply standard method treatment (of conjunctions);
wenzelm
parents:
32786
diff
changeset
|
413 |
|> (case facts of NONE => I | SOME ths => Proof.refine_insert ths) |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
414 |
|> tap (fn state => not schematic andalso Proof.schematic_goal state andalso |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
415 |
error "Illegal schematic goal statement") |
21036 | 416 |
end; |
417 |
||
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
418 |
in |
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
419 |
|
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
420 |
val theorem = |
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
421 |
gen_theorem false Bundle.includes (K I) Expression.cert_statement; |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
422 |
val theorem_cmd = |
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
423 |
gen_theorem false Bundle.includes_cmd Attrib.check_src Expression.read_statement; |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
424 |
|
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
425 |
val schematic_theorem = |
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
426 |
gen_theorem true Bundle.includes (K I) Expression.cert_statement; |
47067
4ef29b0c568f
optional 'includes' element for long theorem statements;
wenzelm
parents:
47066
diff
changeset
|
427 |
val schematic_theorem_cmd = |
56026
893fe12639bc
tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents:
56002
diff
changeset
|
428 |
gen_theorem true Bundle.includes_cmd Attrib.check_src Expression.read_statement; |
21036 | 429 |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
430 |
end; |
21036 | 431 |
|
432 |
end; |