author | haftmann |
Thu, 20 Nov 2008 14:55:28 +0100 | |
changeset 28862 | 53f13f763d4f |
parent 28858 | 855e61829e22 |
child 28877 | 9ff624bd4fe1 |
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 print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
12 |
val check_specification: (Name.binding * typ option * mixfix) list -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
13 |
(Attrib.binding * term list) list list -> Proof.context -> |
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
14 |
(((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
15 |
val read_specification: (Name.binding * string option * mixfix) list -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
16 |
(Attrib.binding * string list) list list -> Proof.context -> |
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
17 |
(((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
18 |
val check_free_specification: (Name.binding * typ option * mixfix) list -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
19 |
(Attrib.binding * term list) list -> Proof.context -> |
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
20 |
(((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
21 |
val read_free_specification: (Name.binding * string option * mixfix) list -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
22 |
(Attrib.binding * string list) list -> Proof.context -> |
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
23 |
(((Name.binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
24 |
val axiomatization: (Name.binding * typ option * mixfix) list -> |
28114 | 25 |
(Attrib.binding * term list) list -> theory -> |
26 |
(term list * (string * thm list) list) * theory |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
27 |
val axiomatization_cmd: (Name.binding * string option * mixfix) list -> |
28114 | 28 |
(Attrib.binding * string list) list -> theory -> |
29 |
(term list * (string * thm list) list) * theory |
|
18954 | 30 |
val definition: |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
31 |
(Name.binding * typ option * mixfix) option * (Attrib.binding * term) -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
32 |
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
|
33 |
val definition_cmd: |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
34 |
(Name.binding * string option * mixfix) option * (Attrib.binding * string) -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
35 |
local_theory -> (term * (string * thm)) * local_theory |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
36 |
val abbreviation: Syntax.mode -> (Name.binding * typ option * mixfix) option * term -> |
21532 | 37 |
local_theory -> local_theory |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
38 |
val abbreviation_cmd: Syntax.mode -> (Name.binding * string option * mixfix) option * string -> |
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
39 |
local_theory -> local_theory |
24949 | 40 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory |
41 |
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
|
42 |
val theorems: string -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
43 |
(Attrib.binding * (thm list * Attrib.src list) list) list -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
44 |
local_theory -> (string * thm list) list * local_theory |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
25371
diff
changeset
|
45 |
val theorems_cmd: string -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
46 |
(Attrib.binding * (Facts.ref * Attrib.src list) list) list -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
47 |
local_theory -> (string * thm list) list * local_theory |
24986 | 48 |
val theorem: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
49 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
28710 | 50 |
Element.context_i list -> Element.statement_i -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
51 |
bool -> local_theory -> Proof.state |
24986 | 52 |
val theorem_cmd: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
53 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
28710 | 54 |
Element.context list -> Element.statement -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
55 |
bool -> local_theory -> Proof.state |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
56 |
val add_theorem_hook: (bool -> Proof.state -> Proof.state) -> Context.generic -> Context.generic |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
57 |
end; |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
58 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
59 |
structure Specification: SPECIFICATION = |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
60 |
struct |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
61 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
62 |
|
20890 | 63 |
(* diagnostics *) |
64 |
||
65 |
fun print_consts _ _ [] = () |
|
24986 | 66 |
| print_consts ctxt pred cs = Pretty.writeln (ProofDisplay.pretty_consts ctxt pred cs); |
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 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
71 |
local |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
72 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
73 |
fun close_forms ctxt i xs As = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
74 |
let |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
75 |
fun add_free (Free (x, _)) = if Variable.is_fixed ctxt x then I else insert (op =) x |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
76 |
| add_free _ = I; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
77 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
78 |
val commons = map #1 xs; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
79 |
val _ = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
80 |
(case duplicates (op =) commons of [] => () |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
81 |
| dups => error ("Duplicate local variables " ^ commas_quote dups)); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
82 |
val frees = rev ((fold o fold_aterms) add_free As (rev commons)); |
24848 | 83 |
val types = map (TypeInfer.param i o rpair []) (Name.invents Name.context Name.aT (length frees)); |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
84 |
val uniform_typing = the o AList.lookup (op =) (frees ~~ types); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
85 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
86 |
fun abs_body lev y (Abs (x, T, b)) = Abs (x, T, abs_body (lev + 1) y b) |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
87 |
| abs_body lev y (t $ u) = abs_body lev y t $ abs_body lev y u |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
88 |
| abs_body lev y (t as Free (x, T)) = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
89 |
if x = y then TypeInfer.constrain (uniform_typing x) (TypeInfer.constrain T (Bound lev)) |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
90 |
else t |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
91 |
| abs_body _ _ a = a; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
92 |
fun close (y, U) B = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
93 |
let val B' = abs_body 0 y (Term.incr_boundvars 1 B) |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
94 |
in if Term.loose_bvar1 (B', 0) then Term.all dummyT $ Abs (y, U, B') else B end; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
95 |
fun close_form A = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
96 |
let |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
97 |
val occ_frees = rev (fold_aterms add_free A []); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
98 |
val bounds = xs @ map (rpair dummyT) (subtract (op =) commons occ_frees); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
99 |
in fold_rev close bounds A end; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
100 |
in map close_form As end; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
101 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
102 |
fun prep_spec prep_vars 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
|
103 |
let |
18670 | 104 |
val thy = ProofContext.theory_of ctxt; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
105 |
|
18670 | 106 |
val (vars, vars_ctxt) = ctxt |> prep_vars raw_vars; |
107 |
val (xs, params_ctxt) = vars_ctxt |> ProofContext.add_fixes_i vars; |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
108 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
109 |
val Asss = (map o map) (map (parse_prop params_ctxt) o snd) raw_specss; |
24734 | 110 |
val names = Variable.names_of (params_ctxt |> (fold o fold o fold) Variable.declare_term Asss) |
111 |
|> fold Name.declare xs; |
|
112 |
val Asss' = #1 ((fold_map o fold_map o fold_map) Term.free_dummy_patterns Asss names); |
|
113 |
val idx = (fold o fold o fold) Term.maxidx_term Asss' ~1 + 1; |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
114 |
val specs = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
115 |
(if do_close then |
24734 | 116 |
#1 (fold_map |
117 |
(fn Ass => fn i => (burrow (close_forms params_ctxt i []) Ass, i + 1)) Asss' idx) |
|
118 |
else Asss') |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
119 |
|> 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
|
120 |
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
|
121 |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
122 |
val Ts = specs_ctxt |> fold_map ProofContext.inferred_param xs |> fst; |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
123 |
val params = map2 (fn (b, _, mx) => fn T => ((b, T), mx)) vars Ts; |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
124 |
val name_atts = map (fn ((name, atts), _) => (name, map (prep_att thy) atts)) (flat raw_specss); |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
125 |
in ((params, name_atts ~~ specs), specs_ctxt) end; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
126 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
127 |
fun read_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src x; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
128 |
fun check_spec x = prep_spec ProofContext.cert_vars (K I) (K I) x; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
129 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
130 |
in |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21359
diff
changeset
|
131 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
132 |
fun read_specification x = read_spec true x; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
133 |
fun check_specification x = check_spec true x; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
134 |
fun read_free_specification vars specs = read_spec false vars [specs]; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
135 |
fun check_free_specification vars specs = check_spec false vars [specs]; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
136 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
137 |
end; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
138 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
139 |
|
28114 | 140 |
(* axiomatization -- within global theory *) |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
141 |
|
28114 | 142 |
fun gen_axioms do_print prep raw_vars raw_specs thy = |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
143 |
let |
28114 | 144 |
val ((vars, specs), _) = prep raw_vars [raw_specs] (ProofContext.init thy); |
145 |
val xs = map (fn ((b, T), _) => (Name.name_of b, T)) vars; |
|
146 |
||
147 |
(*consts*) |
|
148 |
val (consts, consts_thy) = thy |> fold_map (Theory.specify_const []) vars; |
|
149 |
val subst = Term.subst_atomic (map Free xs ~~ consts); |
|
150 |
||
151 |
(*axioms*) |
|
152 |
val (axioms, axioms_thy) = consts_thy |> fold_map (fn ((b, atts), props) => |
|
153 |
fold_map Thm.add_axiom (PureThy.name_multi (Name.name_of b) (map subst props)) |
|
154 |
#>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])]))) specs; |
|
155 |
val (facts, thy') = axioms_thy |> PureThy.note_thmss Thm.axiomK |
|
156 |
(Attrib.map_facts (Attrib.attribute_i axioms_thy) axioms); |
|
26595 | 157 |
val _ = |
158 |
if not do_print then () |
|
28114 | 159 |
else print_consts (ProofContext.init thy') (K false) xs; |
160 |
in ((consts, facts), thy') end; |
|
18786 | 161 |
|
26595 | 162 |
val axiomatization = gen_axioms false check_specification; |
163 |
val axiomatization_cmd = gen_axioms true read_specification; |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
164 |
|
18786 | 165 |
|
166 |
(* definition *) |
|
167 |
||
26595 | 168 |
fun gen_def do_print prep (raw_var, (raw_a, raw_prop)) lthy = |
18786 | 169 |
let |
21705 | 170 |
val (vars, [((raw_name, atts), [prop])]) = |
171 |
fst (prep (the_list raw_var) [(raw_a, [raw_prop])] lthy); |
|
172 |
val (((x, T), rhs), prove) = LocalDefs.derived_def lthy true prop; |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
173 |
val name = Name.map_name (Thm.def_name_optional x) raw_name; |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
174 |
val var = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
175 |
(case vars of |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
176 |
[] => (Name.binding x, NoSyn) |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
177 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
178 |
let |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
179 |
val y = Name.name_of b; |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
180 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
181 |
error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
182 |
Position.str_of (Name.pos_of b)); |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
183 |
in (b, mx) end); |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
184 |
val ((lhs, (_, th)), lthy2) = lthy |> LocalTheory.define Thm.definitionK |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
185 |
(var, ((Name.map_name (suffix "_raw") name, []), rhs)); |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
186 |
val ((def_name, [th']), lthy3) = lthy2 |> LocalTheory.note Thm.definitionK |
28703 | 187 |
((name, Code.add_default_eqn_attrib :: atts), [prove lthy2 th]); |
18786 | 188 |
|
21705 | 189 |
val lhs' = Morphism.term (LocalTheory.target_morphism lthy3) lhs; |
26595 | 190 |
val _ = |
191 |
if not do_print then () |
|
192 |
else print_consts lthy3 (member (op =) (Term.add_frees lhs' [])) [(x, T)]; |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
193 |
in ((lhs, (def_name, th')), lthy3) end; |
18786 | 194 |
|
26595 | 195 |
val definition = gen_def false check_free_specification; |
196 |
val definition_cmd = gen_def true read_free_specification; |
|
18786 | 197 |
|
19080 | 198 |
|
199 |
(* abbreviation *) |
|
200 |
||
26595 | 201 |
fun gen_abbrev do_print prep mode (raw_var, raw_prop) lthy = |
19080 | 202 |
let |
21705 | 203 |
val ((vars, [(_, [prop])]), _) = |
204 |
prep (the_list raw_var) [(("", []), [raw_prop])] |
|
24676 | 205 |
(lthy |> ProofContext.set_mode ProofContext.mode_abbrev); |
21705 | 206 |
val ((x, T), rhs) = LocalDefs.abs_def (#2 (LocalDefs.cert_def lthy prop)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
207 |
val var = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
208 |
(case vars of |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
209 |
[] => (Name.binding x, NoSyn) |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
210 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
211 |
let |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
212 |
val y = Name.name_of b; |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
213 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
214 |
error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
215 |
Position.str_of (Name.pos_of b)); |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
216 |
in (b, mx) end); |
21705 | 217 |
val lthy' = lthy |
24986 | 218 |
|> ProofContext.set_syntax_mode mode (* FIXME ?!? *) |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
219 |
|> LocalTheory.abbrev mode (var, rhs) |> snd |
21705 | 220 |
|> ProofContext.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
|
221 |
|
26595 | 222 |
val _ = if not do_print then () else print_consts lthy' (K false) [(x, T)]; |
21532 | 223 |
in lthy' end; |
19080 | 224 |
|
26595 | 225 |
val abbreviation = gen_abbrev false check_free_specification; |
226 |
val abbreviation_cmd = gen_abbrev true read_free_specification; |
|
19080 | 227 |
|
19664 | 228 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
229 |
(* notation *) |
19664 | 230 |
|
24949 | 231 |
fun gen_notation prep_const add mode args lthy = |
25371
26d349416c4f
notation: improved ProofContext.read_const does the job;
wenzelm
parents:
25211
diff
changeset
|
232 |
lthy |> LocalTheory.notation add mode (map (apfst (prep_const lthy)) args); |
19664 | 233 |
|
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
234 |
val notation = gen_notation (K I); |
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
235 |
val notation_cmd = gen_notation ProofContext.read_const; |
19664 | 236 |
|
20914 | 237 |
|
21036 | 238 |
(* fact statements *) |
20914 | 239 |
|
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
240 |
fun gen_theorems prep_fact prep_att kind raw_facts lthy = |
20914 | 241 |
let |
242 |
val attrib = prep_att (ProofContext.theory_of lthy); |
|
243 |
val facts = raw_facts |> map (fn ((name, atts), bs) => |
|
244 |
((name, map attrib atts), |
|
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
245 |
bs |> map (fn (b, more_atts) => (prep_fact lthy b, map attrib more_atts)))); |
21435 | 246 |
val (res, lthy') = lthy |> LocalTheory.notes kind facts; |
28093 | 247 |
val _ = ProofDisplay.print_results true lthy' ((kind, ""), res); |
20914 | 248 |
in (res, lthy') end; |
249 |
||
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
250 |
val theorems = gen_theorems (K I) (K I); |
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
251 |
val theorems_cmd = gen_theorems ProofContext.get_fact Attrib.intern_src; |
20914 | 252 |
|
21036 | 253 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
254 |
(* complex goal statements *) |
21036 | 255 |
|
256 |
local |
|
257 |
||
21450 | 258 |
fun subtract_prems ctxt1 ctxt2 = |
259 |
Library.drop (length (Assumption.prems_of ctxt1), Assumption.prems_of ctxt2); |
|
260 |
||
21236 | 261 |
fun prep_statement prep_att prep_stmt elems concl ctxt = |
262 |
(case concl of |
|
263 |
Element.Shows shows => |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
264 |
let |
21450 | 265 |
val (_, loc_ctxt, elems_ctxt, propp) = prep_stmt elems (map snd shows) ctxt; |
266 |
val prems = subtract_prems loc_ctxt elems_ctxt; |
|
267 |
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:
21359
diff
changeset
|
268 |
val goal_ctxt = fold (fold (Variable.auto_fixes o fst)) propp elems_ctxt; |
21450 | 269 |
in ((prems, stmt, []), goal_ctxt) end |
21236 | 270 |
| Element.Obtains obtains => |
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
271 |
let |
28862 | 272 |
val case_names = obtains |> map_index (fn (i, (b, _)) => |
273 |
let val name = Name.name_of b |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
274 |
in if name = "" then string_of_int (i + 1) else name end); |
21236 | 275 |
val constraints = obtains |> map (fn (_, (vars, _)) => |
28710 | 276 |
Element.Constrains |
277 |
(vars |> map_filter (fn (x, SOME T) => SOME (Name.name_of x, T) | _ => NONE))); |
|
21036 | 278 |
|
21236 | 279 |
val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props); |
21450 | 280 |
val (_, loc_ctxt, elems_ctxt, propp) = prep_stmt (elems @ constraints) raw_propp ctxt; |
21236 | 281 |
|
282 |
val thesis = ObjectLogic.fixed_judgment (ProofContext.theory_of ctxt) AutoBind.thesisN; |
|
21036 | 283 |
|
21236 | 284 |
fun assume_case ((name, (vars, _)), asms) ctxt' = |
285 |
let |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
286 |
val bs = map fst vars; |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
287 |
val xs = map Name.name_of bs; |
21236 | 288 |
val props = map fst asms; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
289 |
val (Ts, _) = ctxt' |
21236 | 290 |
|> fold Variable.declare_term props |
291 |
|> fold_map ProofContext.inferred_param xs; |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
292 |
val asm = Term.list_all_free (xs ~~ Ts, Logic.list_implies (props, thesis)); |
21236 | 293 |
in |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
294 |
ctxt' |> (snd o ProofContext.add_fixes_i (map (fn b => (b, NONE, NoSyn)) bs)); |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21359
diff
changeset
|
295 |
ctxt' |> Variable.auto_fixes asm |
21236 | 296 |
|> ProofContext.add_assms_i Assumption.assume_export |
297 |
[((name, [ContextRules.intro_query NONE]), [(asm, [])])] |
|
298 |
|>> (fn [(_, [th])] => th) |
|
299 |
end; |
|
300 |
||
21658 | 301 |
val atts = map (Attrib.internal o K) |
21236 | 302 |
[RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names]; |
21450 | 303 |
val prems = subtract_prems loc_ctxt elems_ctxt; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
304 |
val stmt = [((Name.no_binding, atts), [(thesis, [])])]; |
21236 | 305 |
|
306 |
val (facts, goal_ctxt) = elems_ctxt |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
307 |
|> (snd o ProofContext.add_fixes_i [(Name.binding AutoBind.thesisN, NONE, NoSyn)]) |
21236 | 308 |
|> fold_map assume_case (obtains ~~ propp) |
21435 | 309 |
|-> (fn ths => ProofContext.note_thmss_i Thm.assumptionK |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
310 |
[((Name.binding Obtain.thatN, []), [(ths, [])])] #> #2 #> pair ths); |
21450 | 311 |
in ((prems, stmt, facts), goal_ctxt) end); |
21036 | 312 |
|
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
313 |
structure TheoremHook = GenericDataFun |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
314 |
( |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
315 |
type T = ((bool -> Proof.state -> Proof.state) * stamp) list; |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
316 |
val empty = []; |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
317 |
val extend = I; |
24464 | 318 |
fun merge _ hooks : T = Library.merge (eq_snd op =) hooks; |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
319 |
); |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
320 |
|
21036 | 321 |
fun gen_theorem prep_att prep_stmt |
28858 | 322 |
kind before_qed after_qed (name, raw_atts) raw_elems raw_concl int lthy = |
21036 | 323 |
let |
28858 | 324 |
val _ = LocalTheory.affirm lthy; |
325 |
val thy = ProofContext.theory_of lthy; |
|
21036 | 326 |
|
21435 | 327 |
val attrib = prep_att thy; |
328 |
val atts = map attrib raw_atts; |
|
28710 | 329 |
val elems = raw_elems |> map (Element.map_ctxt_attrib attrib); |
21450 | 330 |
val ((prems, stmt, facts), goal_ctxt) = |
28858 | 331 |
prep_statement attrib (prep_stmt NONE) elems raw_concl lthy; |
21036 | 332 |
|
333 |
fun after_qed' results goal_ctxt' = |
|
21602 | 334 |
let val results' = |
335 |
burrow (map Goal.norm_result o ProofContext.export goal_ctxt' lthy) results |
|
336 |
in |
|
21036 | 337 |
lthy |
21435 | 338 |
|> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') |
21036 | 339 |
|> (fn (res, lthy') => |
28791 | 340 |
if Name.is_nothing name andalso null atts then |
28093 | 341 |
(ProofDisplay.print_results true lthy' ((kind, ""), res); lthy') |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
342 |
else |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
343 |
let |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
344 |
val ([(res_name, _)], lthy'') = lthy' |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
345 |
|> LocalTheory.notes kind [((name, atts), [(maps #2 res, [])])]; |
28093 | 346 |
val _ = ProofDisplay.print_results true lthy' ((kind, res_name), res); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
347 |
in lthy'' end) |
21036 | 348 |
|> after_qed results' |
349 |
end; |
|
350 |
in |
|
351 |
goal_ctxt |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
352 |
|> ProofContext.note_thmss_i Thm.assumptionK |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
353 |
[((Name.binding AutoBind.assmsN, []), [(prems, [])])] |
21450 | 354 |
|> snd |
21435 | 355 |
|> Proof.theorem_i before_qed after_qed' (map snd stmt) |
24557 | 356 |
|> Proof.refine_insert facts |
24542 | 357 |
|> Library.apply (map (fn (f, _) => f int) (rev (TheoremHook.get (Context.Proof goal_ctxt)))) |
21036 | 358 |
end; |
359 |
||
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
360 |
in |
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
361 |
|
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
362 |
val theorem = gen_theorem (K I) Locale.cert_context_statement; |
28820 | 363 |
val theorem_cmd = gen_theorem Attrib.intern_src Locale.read_context_statement; |
21036 | 364 |
|
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
365 |
fun add_theorem_hook f = TheoremHook.map (cons (f, stamp ())); |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
366 |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
367 |
end; |
21036 | 368 |
|
369 |
end; |