author | wenzelm |
Thu, 19 Jun 2008 21:14:30 +0200 | |
changeset 27283 | ebd0291ea79c |
parent 26568 | 3a3a83493f00 |
child 28620 | 9846d772b306 |
permissions | -rw-r--r-- |
11897 | 1 |
(* Title: Pure/Isar/object_logic.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Specifics about common object-logics. |
|
6 |
*) |
|
7 |
||
8 |
signature OBJECT_LOGIC = |
|
9 |
sig |
|
25497 | 10 |
val get_base_sort: theory -> sort option |
11 |
val add_base_sort: sort -> theory -> theory |
|
12 |
val typedecl: bstring * string list * mixfix -> theory -> typ * theory |
|
11897 | 13 |
val add_judgment: bstring * string * mixfix -> theory -> theory |
14 |
val add_judgment_i: bstring * typ * mixfix -> theory -> theory |
|
16449 | 15 |
val judgment_name: theory -> string |
16 |
val is_judgment: theory -> term -> bool |
|
17 |
val drop_judgment: theory -> term -> term |
|
18 |
val fixed_judgment: theory -> string -> term |
|
18121 | 19 |
val ensure_propT: theory -> term -> term |
23586 | 20 |
val dest_judgment: cterm -> cterm |
23566
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
21 |
val judgment_conv: conv -> conv |
19261 | 22 |
val is_elim: thm -> bool |
18728 | 23 |
val declare_atomize: attribute |
24 |
val declare_rulify: attribute |
|
16449 | 25 |
val atomize_term: theory -> term -> term |
23602
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
26 |
val atomize: conv |
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
27 |
val atomize_prems: conv |
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
28 |
val atomize_prems_tac: int -> tactic |
12829 | 29 |
val full_atomize_tac: int -> tactic |
18807 | 30 |
val rulify_term: theory -> term -> term |
31 |
val rulify_tac: int -> tactic |
|
11897 | 32 |
val rulify: thm -> thm |
33 |
val rulify_no_asm: thm -> thm |
|
18728 | 34 |
val rule_format: attribute |
35 |
val rule_format_no_asm: attribute |
|
11897 | 36 |
end; |
37 |
||
38 |
structure ObjectLogic: OBJECT_LOGIC = |
|
39 |
struct |
|
40 |
||
25497 | 41 |
(** theory data **) |
11897 | 42 |
|
25497 | 43 |
datatype data = Data of |
44 |
{base_sort: sort option, |
|
45 |
judgment: string option, |
|
46 |
atomize_rulify: thm list * thm list}; |
|
47 |
||
48 |
fun make_data (base_sort, judgment, atomize_rulify) = |
|
49 |
Data {base_sort = base_sort, judgment = judgment, atomize_rulify = atomize_rulify}; |
|
11897 | 50 |
|
16449 | 51 |
structure ObjectLogicData = TheoryDataFun |
22846 | 52 |
( |
25497 | 53 |
type T = data; |
54 |
val empty = make_data (NONE, NONE, ([], [])); |
|
11897 | 55 |
val copy = I; |
16449 | 56 |
val extend = I; |
11897 | 57 |
|
25497 | 58 |
fun merge_opt eq (SOME x, SOME y) = |
59 |
if eq (x, y) then SOME x else error "Attempt to merge different object-logics" |
|
60 |
| merge_opt _ (x, y) = if is_some x then x else y; |
|
11897 | 61 |
|
25497 | 62 |
fun merge _ |
63 |
(Data {base_sort = base_sort1, judgment = judgment1, atomize_rulify = (atomize1, rulify1)}, |
|
64 |
Data {base_sort = base_sort2, judgment = judgment2, atomize_rulify = (atomize2, rulify2)}) = |
|
65 |
make_data (merge_opt (op =) (base_sort1, base_sort2), merge_opt (op =) (judgment1, judgment2), |
|
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
23602
diff
changeset
|
66 |
(Thm.merge_thms (atomize1, atomize2), Thm.merge_thms (rulify1, rulify2))); |
22846 | 67 |
); |
15801 | 68 |
|
25497 | 69 |
fun map_data f = ObjectLogicData.map (fn (Data {base_sort, judgment, atomize_rulify}) => |
70 |
make_data (f (base_sort, judgment, atomize_rulify))); |
|
71 |
||
72 |
fun get_data thy = ObjectLogicData.get thy |> (fn Data args => args); |
|
73 |
||
11897 | 74 |
|
75 |
||
76 |
(** generic treatment of judgments -- with a single argument only **) |
|
77 |
||
25497 | 78 |
(* base_sort *) |
79 |
||
80 |
val get_base_sort = #base_sort o get_data; |
|
81 |
||
82 |
fun add_base_sort S = map_data (fn (base_sort, judgment, atomize_rulify) => |
|
83 |
if is_some base_sort then error "Attempt to redeclare object-logic base sort" |
|
84 |
else (SOME S, judgment, atomize_rulify)); |
|
85 |
||
86 |
||
87 |
(* typedecl *) |
|
88 |
||
89 |
fun typedecl (raw_name, vs, mx) thy = |
|
90 |
let |
|
91 |
val base_sort = get_base_sort thy; |
|
92 |
val name = Sign.full_name thy (Syntax.type_name raw_name mx); |
|
93 |
val _ = has_duplicates (op =) vs andalso |
|
94 |
error ("Duplicate parameters in type declaration: " ^ quote name); |
|
95 |
val n = length vs; |
|
96 |
val T = Type (name, map (fn v => TFree (v, [])) vs); |
|
97 |
in |
|
98 |
thy |
|
99 |
|> Sign.add_types [(raw_name, n, mx)] |
|
100 |
|> (case base_sort of NONE => I | SOME S => AxClass.axiomatize_arity (name, replicate n S, S)) |
|
101 |
|> pair T |
|
102 |
end; |
|
103 |
||
104 |
||
18825 | 105 |
(* add judgment *) |
11897 | 106 |
|
107 |
local |
|
108 |
||
16449 | 109 |
fun gen_add_judgment add_consts (bname, T, mx) thy = |
110 |
let val c = Sign.full_name thy (Syntax.const_name bname mx) in |
|
14226 | 111 |
thy |
16449 | 112 |
|> add_consts [(bname, T, mx)] |
25018
fac2ceba75b4
replaced obsolete Theory.add_finals_i by Theory.add_deps;
wenzelm
parents:
24848
diff
changeset
|
113 |
|> (fn thy' => Theory.add_deps c (c, Sign.the_const_type thy' c) [] thy') |
25497 | 114 |
|> map_data (fn (base_sort, judgment, atomize_rulify) => |
115 |
if is_some judgment then error "Attempt to redeclare object-logic judgment" |
|
116 |
else (base_sort, SOME c, atomize_rulify)) |
|
14226 | 117 |
end; |
11897 | 118 |
|
119 |
in |
|
120 |
||
22796 | 121 |
val add_judgment = gen_add_judgment Sign.add_consts; |
122 |
val add_judgment_i = gen_add_judgment Sign.add_consts_i; |
|
11897 | 123 |
|
124 |
end; |
|
125 |
||
126 |
||
23566
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
127 |
(* judgments *) |
11897 | 128 |
|
16449 | 129 |
fun judgment_name thy = |
25497 | 130 |
(case #judgment (get_data thy) of |
131 |
SOME name => name |
|
11897 | 132 |
| _ => raise TERM ("Unknown object-logic judgment", [])); |
133 |
||
16449 | 134 |
fun is_judgment thy (Const (c, _) $ _) = c = judgment_name thy |
11897 | 135 |
| is_judgment _ _ = false; |
136 |
||
16449 | 137 |
fun drop_judgment thy (Abs (x, T, t)) = Abs (x, T, drop_judgment thy t) |
138 |
| drop_judgment thy (tm as (Const (c, _) $ t)) = |
|
139 |
if (c = judgment_name thy handle TERM _ => false) then t else tm |
|
11897 | 140 |
| drop_judgment _ tm = tm; |
141 |
||
16449 | 142 |
fun fixed_judgment thy x = |
11897 | 143 |
let (*be robust wrt. low-level errors*) |
16449 | 144 |
val c = judgment_name thy; |
24848 | 145 |
val aT = TFree (Name.aT, []); |
11897 | 146 |
val T = |
18939 | 147 |
the_default (aT --> propT) (Sign.const_type thy c) |
11897 | 148 |
|> Term.map_type_tvar (fn ((x, _), S) => TFree (x, S)); |
149 |
val U = Term.domain_type T handle Match => aT; |
|
150 |
in Const (c, T) $ Free (x, U) end; |
|
151 |
||
18121 | 152 |
fun ensure_propT thy t = |
13376 | 153 |
let val T = Term.fastype_of t |
16449 | 154 |
in if T = propT then t else Const (judgment_name thy, T --> propT) $ t end; |
13376 | 155 |
|
23586 | 156 |
fun dest_judgment ct = |
157 |
if is_judgment (Thm.theory_of_cterm ct) (Thm.term_of ct) |
|
158 |
then Thm.dest_arg ct |
|
159 |
else raise CTERM ("dest_judgment", [ct]); |
|
160 |
||
23566
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
161 |
fun judgment_conv cv ct = |
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
162 |
if is_judgment (Thm.theory_of_cterm ct) (Thm.term_of ct) |
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
163 |
then Conv.arg_conv cv ct |
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
164 |
else raise CTERM ("judgment_conv", [ct]); |
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
wenzelm
parents:
23540
diff
changeset
|
165 |
|
11897 | 166 |
|
19261 | 167 |
(* elimination rules *) |
168 |
||
169 |
fun is_elim rule = |
|
170 |
let |
|
171 |
val thy = Thm.theory_of_thm rule; |
|
172 |
val concl = Thm.concl_of rule; |
|
173 |
in |
|
174 |
Term.is_Var (drop_judgment thy concl) andalso |
|
175 |
exists (fn prem => concl aconv Logic.strip_assums_concl prem) (Thm.prems_of rule) |
|
176 |
end; |
|
177 |
||
178 |
||
11897 | 179 |
|
180 |
(** treatment of meta-level connectives **) |
|
181 |
||
182 |
(* maintain rules *) |
|
183 |
||
25497 | 184 |
val get_atomize = #1 o #atomize_rulify o get_data; |
185 |
val get_rulify = #2 o #atomize_rulify o get_data; |
|
11897 | 186 |
|
25497 | 187 |
fun add_atomize th = map_data (fn (base_sort, judgment, (atomize, rulify)) => |
188 |
(base_sort, judgment, (Thm.add_thm th atomize, rulify))); |
|
189 |
||
190 |
fun add_rulify th = map_data (fn (base_sort, judgment, (atomize, rulify)) => |
|
191 |
(base_sort, judgment, (atomize, Thm.add_thm th rulify))); |
|
11897 | 192 |
|
22846 | 193 |
val declare_atomize = Thm.declaration_attribute (fn th => Context.mapping (add_atomize th) I); |
194 |
val declare_rulify = Thm.declaration_attribute (fn th => Context.mapping (add_rulify th) I); |
|
195 |
||
26463 | 196 |
val _ = Context.>> (Context.map_theory (add_rulify Drule.norm_hhf_eq)); |
11897 | 197 |
|
198 |
||
199 |
(* atomize *) |
|
200 |
||
16449 | 201 |
fun atomize_term thy = |
202 |
drop_judgment thy o MetaSimplifier.rewrite_term thy (get_atomize thy) []; |
|
12729 | 203 |
|
23602
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
204 |
fun atomize ct = |
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
205 |
MetaSimplifier.rewrite true (get_atomize (Thm.theory_of_cterm ct)) ct; |
14743 | 206 |
|
23602
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
207 |
fun atomize_prems ct = |
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
208 |
if Logic.has_meta_prems (Thm.term_of ct) then |
26568 | 209 |
Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize)) |
24832 | 210 |
(ProofContext.init (Thm.theory_of_cterm ct)) ct |
23602
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
211 |
else Conv.all_conv ct; |
11897 | 212 |
|
23602
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
213 |
val atomize_prems_tac = CONVERSION atomize_prems; |
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
wenzelm
parents:
23586
diff
changeset
|
214 |
val full_atomize_tac = CONVERSION atomize; |
11897 | 215 |
|
216 |
||
217 |
(* rulify *) |
|
218 |
||
18807 | 219 |
fun rulify_term thy = MetaSimplifier.rewrite_term thy (get_rulify thy) []; |
23540 | 220 |
fun rulify_tac i st = MetaSimplifier.rewrite_goal_tac (get_rulify (Thm.theory_of_thm st)) i st; |
18807 | 221 |
|
11897 | 222 |
fun gen_rulify full thm = |
21708 | 223 |
MetaSimplifier.simplify full (get_rulify (Thm.theory_of_thm thm)) thm |
20912 | 224 |
|> Drule.gen_all |> Thm.strip_shyps |> Drule.zero_var_indexes; |
11897 | 225 |
|
226 |
val rulify = gen_rulify true; |
|
227 |
val rulify_no_asm = gen_rulify false; |
|
228 |
||
18728 | 229 |
fun rule_format x = Thm.rule_attribute (fn _ => rulify) x; |
230 |
fun rule_format_no_asm x = Thm.rule_attribute (fn _ => rulify_no_asm) x; |
|
11897 | 231 |
|
232 |
end; |