author | wenzelm |
Thu, 19 Aug 2010 11:26:07 +0200 | |
changeset 38480 | e5eed57913d0 |
parent 38388 | 94d5624dd1f7 |
child 39288 | f1ae2493d93f |
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 |
20890 | 10 |
val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit |
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
|
11 |
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
|
12 |
(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
|
13 |
(((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
|
14 |
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
|
15 |
(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
|
16 |
(((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
|
17 |
val check_free_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
|
18 |
(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
|
19 |
(((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
|
20 |
val read_free_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
|
21 |
(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
|
22 |
(((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context |
29581 | 23 |
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
|
24 |
(Attrib.binding * term list) list -> Proof.context -> |
29581 | 25 |
(((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
26 |
val read_specification: (binding * string option * mixfix) list -> |
|
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
27 |
(Attrib.binding * string list) list -> Proof.context -> |
29581 | 28 |
(((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context |
29 |
val axiomatization: (binding * typ option * mixfix) list -> |
|
28114 | 30 |
(Attrib.binding * term list) list -> theory -> |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
31 |
(term list * thm list list) * theory |
29581 | 32 |
val axiomatization_cmd: (binding * string option * mixfix) list -> |
28114 | 33 |
(Attrib.binding * string list) list -> theory -> |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
34 |
(term list * thm list list) * theory |
35894 | 35 |
val axiom: Attrib.binding * term -> theory -> thm * theory |
36 |
val axiom_cmd: Attrib.binding * string -> theory -> thm * theory |
|
18954 | 37 |
val definition: |
29581 | 38 |
(binding * typ option * mixfix) option * (Attrib.binding * term) -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
39 |
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
|
40 |
val definition_cmd: |
29581 | 41 |
(binding * string option * mixfix) option * (Attrib.binding * string) -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
42 |
local_theory -> (term * (string * thm)) * local_theory |
29581 | 43 |
val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option * term -> |
21532 | 44 |
local_theory -> local_theory |
29581 | 45 |
val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option * string -> |
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
46 |
local_theory -> local_theory |
35413 | 47 |
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory |
48 |
val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory |
|
24949 | 49 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory |
50 |
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
|
51 |
val theorems: string -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
52 |
(Attrib.binding * (thm list * Attrib.src list) list) list -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
53 |
local_theory -> (string * thm list) list * local_theory |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
25371
diff
changeset
|
54 |
val theorems_cmd: string -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
55 |
(Attrib.binding * (Facts.ref * Attrib.src list) list) list -> |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
56 |
local_theory -> (string * thm list) list * local_theory |
24986 | 57 |
val theorem: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
58 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
28710 | 59 |
Element.context_i list -> Element.statement_i -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
60 |
bool -> local_theory -> Proof.state |
24986 | 61 |
val theorem_cmd: string -> Method.text option -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28080
diff
changeset
|
62 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
28710 | 63 |
Element.context list -> Element.statement -> |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
64 |
bool -> local_theory -> Proof.state |
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
65 |
val schematic_theorem: string -> Method.text option -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
66 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
67 |
Element.context_i list -> Element.statement_i -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
68 |
bool -> local_theory -> Proof.state |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
69 |
val schematic_theorem_cmd: string -> Method.text option -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
70 |
(thm list list -> local_theory -> local_theory) -> Attrib.binding -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
71 |
Element.context list -> Element.statement -> |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
72 |
bool -> local_theory -> Proof.state |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
73 |
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
|
74 |
end; |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
75 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
76 |
structure Specification: SPECIFICATION = |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
77 |
struct |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
78 |
|
20890 | 79 |
(* diagnostics *) |
80 |
||
81 |
fun print_consts _ _ [] = () |
|
33389 | 82 |
| print_consts ctxt pred cs = Pretty.writeln (Proof_Display.pretty_consts ctxt pred cs); |
20914 | 83 |
|
19664 | 84 |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
85 |
(* prepare specification *) |
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
86 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
87 |
local |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
88 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
89 |
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
|
90 |
let |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
91 |
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
|
92 |
| add_free _ = I; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
93 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
94 |
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
|
95 |
val _ = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
96 |
(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
|
97 |
| 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
|
98 |
val frees = rev ((fold o fold_aterms) add_free As (rev commons)); |
37145
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents:
36610
diff
changeset
|
99 |
val types = map (Type_Infer.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
|
100 |
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
|
101 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
102 |
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
|
103 |
| 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
|
104 |
| abs_body lev y (t as Free (x, T)) = |
37145
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents:
36610
diff
changeset
|
105 |
if x = y then Type_Infer.constrain (uniform_typing x) (Type_Infer.constrain T (Bound lev)) |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
106 |
else t |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
107 |
| abs_body _ _ a = a; |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
108 |
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
|
109 |
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
|
110 |
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
|
111 |
fun close_form A = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
112 |
let |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
|
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
|
118 |
fun prepare 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
|
119 |
let |
18670 | 120 |
val thy = ProofContext.theory_of ctxt; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
121 |
|
18670 | 122 |
val (vars, vars_ctxt) = ctxt |> prep_vars raw_vars; |
30763
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents:
30761
diff
changeset
|
123 |
val (xs, params_ctxt) = vars_ctxt |> ProofContext.add_fixes vars; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
124 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
125 |
val Asss = (map o map) (map (parse_prop params_ctxt) o snd) raw_specss; |
24734 | 126 |
val names = Variable.names_of (params_ctxt |> (fold o fold o fold) Variable.declare_term Asss) |
127 |
|> fold Name.declare xs; |
|
128 |
val Asss' = #1 ((fold_map o fold_map o fold_map) Term.free_dummy_patterns Asss names); |
|
129 |
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
|
130 |
val specs = |
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
131 |
(if do_close then |
24734 | 132 |
#1 (fold_map |
133 |
(fn Ass => fn i => (burrow (close_forms params_ctxt i []) Ass, i + 1)) Asss' idx) |
|
134 |
else Asss') |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
135 |
|> 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
|
136 |
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
|
137 |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
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
|
142 |
|
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
|
143 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
144 |
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
|
145 |
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
|
146 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
147 |
fun prep_spec prep_vars parse_prop prep_att do_close vars specs = |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
148 |
prepare prep_vars parse_prop prep_att do_close |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
149 |
vars (map single_spec specs) #>> apsnd (map the_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
|
150 |
|
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
151 |
in |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21359
diff
changeset
|
152 |
|
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
|
153 |
fun check_spec x = prep_spec ProofContext.cert_vars (K I) (K I) true x; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
154 |
fun read_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src true x; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
155 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
156 |
fun check_free_spec x = prep_spec ProofContext.cert_vars (K I) (K I) false x; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
157 |
fun read_free_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src false x; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
158 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
159 |
fun check_specification vars specs = |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
160 |
prepare ProofContext.cert_vars (K I) (K I) true vars [specs]; |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
161 |
|
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
162 |
fun read_specification vars specs = |
6e3c92de4a7d
simplified versions check_spec, read_spec, check_free_spec, read_free_spec: operate on list of singleton statements;
wenzelm
parents:
30472
diff
changeset
|
163 |
prepare ProofContext.read_vars Syntax.parse_prop Attrib.intern_src true vars [specs]; |
24724
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
164 |
|
0df74bb87a13
read/check_specification: proper type inference across multiple sections, result is in closed form;
wenzelm
parents:
24676
diff
changeset
|
165 |
end; |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
166 |
|
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
167 |
|
28114 | 168 |
(* axiomatization -- within global theory *) |
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
169 |
|
28114 | 170 |
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
|
171 |
let |
36610
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents:
36505
diff
changeset
|
172 |
val ((vars, specs), _) = prep raw_vars raw_specs (ProofContext.init_global thy); |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
173 |
val xs = map (fn ((b, T), _) => (Name.of_binding b, T)) vars; |
28114 | 174 |
|
175 |
(*consts*) |
|
33173
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents:
32856
diff
changeset
|
176 |
val (consts, consts_thy) = thy |> fold_map Theory.specify_const vars; |
28114 | 177 |
val subst = Term.subst_atomic (map Free xs ~~ consts); |
178 |
||
179 |
(*axioms*) |
|
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
180 |
val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), props) => |
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
181 |
fold_map Thm.add_axiom |
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
182 |
(map (apfst (fn a => Binding.map_name (K a) b)) |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
183 |
(PureThy.name_multi (Name.of_binding b) (map subst props))) |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35894
diff
changeset
|
184 |
#>> (fn ths => ((b, atts), [(map #2 ths, [])]))); |
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
185 |
|
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
186 |
(*facts*) |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
187 |
val (facts, facts_lthy) = axioms_thy |
38388 | 188 |
|> Named_Target.theory_init |
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
189 |
|> 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
|
190 |
|> Local_Theory.notes axioms; |
30470
9ae8f9d78cd3
axiomatization: more precise treatment of binding;
wenzelm
parents:
30434
diff
changeset
|
191 |
|
26595 | 192 |
val _ = |
193 |
if not do_print then () |
|
33703
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
194 |
else print_consts facts_lthy (K false) xs; |
50b6c07c0dd4
axiomatization: declare Spec_Rules, direct result;
wenzelm
parents:
33671
diff
changeset
|
195 |
in ((consts, map #2 facts), Local_Theory.exit_global facts_lthy) end; |
18786 | 196 |
|
26595 | 197 |
val axiomatization = gen_axioms false check_specification; |
198 |
val axiomatization_cmd = gen_axioms true read_specification; |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
199 |
|
35894 | 200 |
fun axiom (b, ax) = axiomatization [] [(b, [ax])] #>> (hd o hd o snd); |
201 |
fun axiom_cmd (b, ax) = axiomatization_cmd [] [(b, [ax])] #>> (hd o hd o snd); |
|
202 |
||
18786 | 203 |
|
204 |
(* definition *) |
|
205 |
||
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
|
206 |
fun gen_def do_print prep (raw_var, raw_spec) lthy = |
18786 | 207 |
let |
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
|
208 |
val (vars, [((raw_name, atts), prop)]) = fst (prep (the_list raw_var) [raw_spec] lthy); |
35624 | 209 |
val (((x, T), rhs), prove) = Local_Defs.derived_def lthy true prop; |
30434 | 210 |
val var as (b, _) = |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
211 |
(case vars of |
28965 | 212 |
[] => (Binding.name x, NoSyn) |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
213 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
214 |
let |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
215 |
val y = Name.of_binding b; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
216 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
217 |
error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ |
28941 | 218 |
Position.str_of (Binding.pos_of b)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
219 |
in (b, mx) end); |
30434 | 220 |
val name = Thm.def_binding_optional b raw_name; |
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
221 |
val ((lhs, (_, raw_th)), lthy2) = lthy |
33765
47b5db097646
Specification.definition: Thm.definitionK marks exactly the high-level end-user result;
wenzelm
parents:
33756
diff
changeset
|
222 |
|> Local_Theory.define (var, ((Binding.suffix_name "_raw" name, []), rhs)); |
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
223 |
|
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
224 |
val th = prove lthy2 raw_th; |
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
225 |
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
|
226 |
|
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33644
diff
changeset
|
227 |
val ([(def_name, [th'])], lthy4) = lthy3 |
33671 | 228 |
|> Local_Theory.notes_kind Thm.definitionK |
33756
47b7c9e0bf6e
replacing Predicate_Compile_Preproc_Const_Defs by more general Spec_Rules
bulwahn
parents:
33703
diff
changeset
|
229 |
[((name, Code.add_default_eqn_attrib :: atts), [([th], [])])]; |
18786 | 230 |
|
33671 | 231 |
val lhs' = Morphism.term (Local_Theory.target_morphism lthy4) lhs; |
26595 | 232 |
val _ = |
233 |
if not do_print then () |
|
33455
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
234 |
else print_consts lthy4 (member (op =) (Term.add_frees lhs' [])) [(x, T)]; |
4da71b27b3fe
declare Spec_Rules for most basic definitional packages;
wenzelm
parents:
33389
diff
changeset
|
235 |
in ((lhs, (def_name, th')), lthy4) end; |
18786 | 236 |
|
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
|
237 |
val definition = gen_def false check_free_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
|
238 |
val definition_cmd = gen_def true read_free_spec; |
18786 | 239 |
|
19080 | 240 |
|
241 |
(* abbreviation *) |
|
242 |
||
26595 | 243 |
fun gen_abbrev do_print prep mode (raw_var, raw_prop) lthy = |
19080 | 244 |
let |
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
|
245 |
val ((vars, [(_, 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
|
246 |
prep (the_list raw_var) [(Attrib.empty_binding, raw_prop)] |
24676 | 247 |
(lthy |> ProofContext.set_mode ProofContext.mode_abbrev); |
35624 | 248 |
val ((x, T), rhs) = Local_Defs.abs_def (#2 (Local_Defs.cert_def lthy prop)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
249 |
val var = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
250 |
(case vars of |
28965 | 251 |
[] => (Binding.name x, NoSyn) |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
252 |
| [((b, _), mx)] => |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
253 |
let |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
254 |
val y = Name.of_binding b; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
255 |
val _ = x = y orelse |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
256 |
error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ |
28941 | 257 |
Position.str_of (Binding.pos_of b)); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
258 |
in (b, mx) end); |
21705 | 259 |
val lthy' = lthy |
24986 | 260 |
|> ProofContext.set_syntax_mode mode (* FIXME ?!? *) |
33671 | 261 |
|> Local_Theory.abbrev mode (var, rhs) |> snd |
21705 | 262 |
|> 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
|
263 |
|
26595 | 264 |
val _ = if not do_print then () else print_consts lthy' (K false) [(x, T)]; |
21532 | 265 |
in lthy' end; |
19080 | 266 |
|
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
|
267 |
val abbreviation = gen_abbrev false check_free_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
|
268 |
val abbreviation_cmd = gen_abbrev true read_free_spec; |
19080 | 269 |
|
19664 | 270 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
271 |
(* notation *) |
19664 | 272 |
|
35413 | 273 |
local |
274 |
||
275 |
fun gen_type_notation prep_type add mode args lthy = |
|
276 |
lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args); |
|
277 |
||
24949 | 278 |
fun gen_notation prep_const add mode args lthy = |
33671 | 279 |
lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); |
19664 | 280 |
|
35413 | 281 |
in |
282 |
||
283 |
val type_notation = gen_type_notation (K I); |
|
284 |
val type_notation_cmd = gen_type_notation (fn ctxt => ProofContext.read_type_name ctxt false); |
|
285 |
||
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
286 |
val notation = gen_notation (K I); |
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36323
diff
changeset
|
287 |
val notation_cmd = gen_notation (fn ctxt => ProofContext.read_const ctxt false dummyT); |
19664 | 288 |
|
35413 | 289 |
end; |
290 |
||
20914 | 291 |
|
21036 | 292 |
(* fact statements *) |
20914 | 293 |
|
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
294 |
fun gen_theorems prep_fact prep_att kind raw_facts lthy = |
20914 | 295 |
let |
296 |
val attrib = prep_att (ProofContext.theory_of lthy); |
|
297 |
val facts = raw_facts |> map (fn ((name, atts), bs) => |
|
298 |
((name, map attrib atts), |
|
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
299 |
bs |> map (fn (b, more_atts) => (prep_fact lthy b, map attrib more_atts)))); |
33671 | 300 |
val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts; |
33389 | 301 |
val _ = Proof_Display.print_results true lthy' ((kind, ""), res); |
20914 | 302 |
in (res, lthy') end; |
303 |
||
24927
48e08f37ce92
Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
wenzelm
parents:
24848
diff
changeset
|
304 |
val theorems = gen_theorems (K I) (K I); |
26345
f70620a4cf81
renamed former get_thms(_silent) to get_fact(_silent);
wenzelm
parents:
26336
diff
changeset
|
305 |
val theorems_cmd = gen_theorems ProofContext.get_fact Attrib.intern_src; |
20914 | 306 |
|
21036 | 307 |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
308 |
(* complex goal statements *) |
21036 | 309 |
|
310 |
local |
|
311 |
||
21236 | 312 |
fun prep_statement prep_att prep_stmt elems concl ctxt = |
313 |
(case concl of |
|
314 |
Element.Shows shows => |
|
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
315 |
let |
28880 | 316 |
val (propp, elems_ctxt) = prep_stmt elems (map snd shows) ctxt; |
30472 | 317 |
val prems = Assumption.local_prems_of elems_ctxt ctxt; |
21450 | 318 |
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
|
319 |
val goal_ctxt = fold (fold (Variable.auto_fixes o fst)) propp elems_ctxt; |
32856
92d9555ac790
clarified Proof.refine_insert -- always "refine" to apply standard method treatment (of conjunctions);
wenzelm
parents:
32786
diff
changeset
|
320 |
in ((prems, stmt, NONE), goal_ctxt) end |
21236 | 321 |
| Element.Obtains obtains => |
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
322 |
let |
28862 | 323 |
val case_names = obtains |> map_index (fn (i, (b, _)) => |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
324 |
if Binding.is_empty b then string_of_int (i + 1) else Name.of_binding b); |
21236 | 325 |
val constraints = obtains |> map (fn (_, (vars, _)) => |
28710 | 326 |
Element.Constrains |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
327 |
(vars |> map_filter (fn (x, SOME T) => SOME (Name.of_binding x, T) | _ => NONE))); |
21036 | 328 |
|
21236 | 329 |
val raw_propp = obtains |> map (fn (_, (_, props)) => map (rpair []) props); |
28880 | 330 |
val (propp, elems_ctxt) = prep_stmt (elems @ constraints) raw_propp ctxt; |
21236 | 331 |
|
35625 | 332 |
val thesis = Object_Logic.fixed_judgment (ProofContext.theory_of ctxt) Auto_Bind.thesisN; |
21036 | 333 |
|
21236 | 334 |
fun assume_case ((name, (vars, _)), asms) ctxt' = |
335 |
let |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
336 |
val bs = map fst vars; |
30585
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents:
30482
diff
changeset
|
337 |
val xs = map Name.of_binding bs; |
21236 | 338 |
val props = map fst asms; |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
339 |
val (Ts, _) = ctxt' |
21236 | 340 |
|> fold Variable.declare_term props |
341 |
|> fold_map ProofContext.inferred_param xs; |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
342 |
val asm = Term.list_all_free (xs ~~ Ts, Logic.list_implies (props, thesis)); |
21236 | 343 |
in |
30763
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents:
30761
diff
changeset
|
344 |
ctxt' |> (snd o ProofContext.add_fixes (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
|
345 |
ctxt' |> Variable.auto_fixes asm |
21236 | 346 |
|> ProofContext.add_assms_i Assumption.assume_export |
33369 | 347 |
[((name, [Context_Rules.intro_query NONE]), [(asm, [])])] |
21236 | 348 |
|>> (fn [(_, [th])] => th) |
349 |
end; |
|
350 |
||
21658 | 351 |
val atts = map (Attrib.internal o K) |
33368 | 352 |
[Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names case_names]; |
30472 | 353 |
val prems = Assumption.local_prems_of elems_ctxt ctxt; |
28965 | 354 |
val stmt = [((Binding.empty, atts), [(thesis, [])])]; |
21236 | 355 |
|
356 |
val (facts, goal_ctxt) = elems_ctxt |
|
33386 | 357 |
|> (snd o ProofContext.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]) |
21236 | 358 |
|> fold_map assume_case (obtains ~~ propp) |
33644
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
wenzelm
parents:
33519
diff
changeset
|
359 |
|-> (fn ths => |
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
wenzelm
parents:
33519
diff
changeset
|
360 |
ProofContext.note_thmss "" [((Binding.name Obtain.thatN, []), [(ths, [])])] #> |
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
wenzelm
parents:
33519
diff
changeset
|
361 |
#2 #> pair ths); |
32856
92d9555ac790
clarified Proof.refine_insert -- always "refine" to apply standard method treatment (of conjunctions);
wenzelm
parents:
32786
diff
changeset
|
362 |
in ((prems, stmt, SOME facts), goal_ctxt) end); |
21036 | 363 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
364 |
structure Theorem_Hook = Generic_Data |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
365 |
( |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
366 |
type T = ((bool -> Proof.state -> Proof.state) * stamp) list; |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
367 |
val empty = []; |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
368 |
val extend = I; |
33519 | 369 |
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
|
370 |
); |
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
371 |
|
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
372 |
fun gen_theorem schematic prep_att prep_stmt |
28858 | 373 |
kind before_qed after_qed (name, raw_atts) raw_elems raw_concl int lthy = |
21036 | 374 |
let |
33671 | 375 |
val _ = Local_Theory.affirm lthy; |
28858 | 376 |
val thy = ProofContext.theory_of lthy; |
21036 | 377 |
|
21435 | 378 |
val attrib = prep_att thy; |
379 |
val atts = map attrib raw_atts; |
|
28710 | 380 |
val elems = raw_elems |> map (Element.map_ctxt_attrib attrib); |
21450 | 381 |
val ((prems, stmt, facts), goal_ctxt) = |
28880 | 382 |
prep_statement attrib prep_stmt elems raw_concl lthy; |
21036 | 383 |
|
384 |
fun after_qed' results goal_ctxt' = |
|
21602 | 385 |
let val results' = |
386 |
burrow (map Goal.norm_result o ProofContext.export goal_ctxt' lthy) results |
|
387 |
in |
|
21036 | 388 |
lthy |
33671 | 389 |
|> Local_Theory.notes_kind kind |
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33644
diff
changeset
|
390 |
(map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') |
21036 | 391 |
|> (fn (res, lthy') => |
28965 | 392 |
if Binding.is_empty name andalso null atts then |
33389 | 393 |
(Proof_Display.print_results true lthy' ((kind, ""), res); lthy') |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
394 |
else |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
395 |
let |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
396 |
val ([(res_name, _)], lthy'') = lthy' |
33671 | 397 |
|> Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])]; |
33389 | 398 |
val _ = Proof_Display.print_results true lthy' ((kind, res_name), res); |
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27858
diff
changeset
|
399 |
in lthy'' end) |
21036 | 400 |
|> after_qed results' |
401 |
end; |
|
402 |
in |
|
403 |
goal_ctxt |
|
33644
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
wenzelm
parents:
33519
diff
changeset
|
404 |
|> ProofContext.note_thmss "" [((Binding.name Auto_Bind.assmsN, []), [(prems, [])])] |
21450 | 405 |
|> snd |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36317
diff
changeset
|
406 |
|> 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
|
407 |
|> (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
|
408 |
|> 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
|
409 |
error "Illegal schematic goal statement") |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
410 |
|> Library.apply (map (fn (f, _) => f int) (rev (Theorem_Hook.get (Context.Proof goal_ctxt)))) |
21036 | 411 |
end; |
412 |
||
21230
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
413 |
in |
abfdce60b371
theorem statements: incorporate Obtain.statement, tuned;
wenzelm
parents:
21206
diff
changeset
|
414 |
|
36317
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
415 |
val theorem = gen_theorem false (K I) Expression.cert_statement; |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
416 |
val theorem_cmd = gen_theorem false Attrib.intern_src Expression.read_statement; |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
417 |
|
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
418 |
val schematic_theorem = gen_theorem true (K I) Expression.cert_statement; |
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
wenzelm
parents:
36106
diff
changeset
|
419 |
val schematic_theorem_cmd = gen_theorem true Attrib.intern_src Expression.read_statement; |
21036 | 420 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
421 |
fun add_theorem_hook f = Theorem_Hook.map (cons (f, stamp ())); |
24452
93b113c5ac33
- theorem(_i) now also takes "interactive" flag as argument
berghofe
parents:
24219
diff
changeset
|
422 |
|
18620
fc8b5f275359
Theory specifications --- with type-inference, but no internal polymorphism.
wenzelm
parents:
diff
changeset
|
423 |
end; |
21036 | 424 |
|
425 |
end; |