author | wenzelm |
Sun, 05 Jun 2005 23:07:26 +0200 | |
changeset 16288 | df2b550a17f6 |
parent 15973 | 5fd94d84470f |
child 16449 | d0dc9a301e37 |
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 |
|
10 |
val add_judgment: bstring * string * mixfix -> theory -> theory |
|
11 |
val add_judgment_i: bstring * typ * mixfix -> theory -> theory |
|
12 |
val judgment_name: Sign.sg -> string |
|
13 |
val is_judgment: Sign.sg -> term -> bool |
|
14 |
val drop_judgment: Sign.sg -> term -> term |
|
15 |
val fixed_judgment: Sign.sg -> string -> term |
|
13376 | 16 |
val assert_propT: Sign.sg -> term -> term |
11897 | 17 |
val declare_atomize: theory attribute |
18 |
val declare_rulify: theory attribute |
|
12807 | 19 |
val atomize_term: Sign.sg -> term -> term |
14743 | 20 |
val atomize_thm: thm -> thm |
13376 | 21 |
val atomize_rule: Sign.sg -> cterm -> thm |
11897 | 22 |
val atomize_tac: int -> tactic |
12829 | 23 |
val full_atomize_tac: int -> tactic |
11897 | 24 |
val atomize_goal: int -> thm -> thm |
25 |
val rulify: thm -> thm |
|
26 |
val rulify_no_asm: thm -> thm |
|
27 |
val rule_format: 'a attribute |
|
28 |
val rule_format_no_asm: 'a attribute |
|
29 |
end; |
|
30 |
||
31 |
structure ObjectLogic: OBJECT_LOGIC = |
|
32 |
struct |
|
33 |
||
34 |
||
35 |
(** object-logic theory data **) |
|
36 |
||
37 |
(* data kind 'Pure/object-logic' *) |
|
38 |
||
39 |
structure ObjectLogicDataArgs = |
|
40 |
struct |
|
41 |
val name = "Pure/object-logic"; |
|
42 |
type T = string option * (thm list * thm list); |
|
43 |
||
15531 | 44 |
val empty = (NONE, ([], [Drule.norm_hhf_eq])); |
11897 | 45 |
val copy = I; |
46 |
val prep_ext = I; |
|
47 |
||
15531 | 48 |
fun merge_judgment (SOME x, SOME y) = |
49 |
if x = y then SOME x else error "Attempt to merge different object-logics" |
|
15973 | 50 |
| merge_judgment (j1, j2) = if is_some j1 then j1 else j2; |
11897 | 51 |
|
52 |
fun merge ((judgment1, (atomize1, rulify1)), (judgment2, (atomize2, rulify2))) = |
|
53 |
(merge_judgment (judgment1, judgment2), |
|
54 |
(Drule.merge_rules (atomize1, atomize2), Drule.merge_rules (rulify1, rulify2))); |
|
55 |
||
56 |
fun print _ _ = (); |
|
57 |
end; |
|
58 |
||
59 |
structure ObjectLogicData = TheoryDataFun(ObjectLogicDataArgs); |
|
15801 | 60 |
val _ = Context.add_setup [ObjectLogicData.init]; |
61 |
||
11897 | 62 |
|
63 |
||
64 |
(** generic treatment of judgments -- with a single argument only **) |
|
65 |
||
66 |
(* add_judgment(_i) *) |
|
67 |
||
68 |
local |
|
69 |
||
15531 | 70 |
fun new_judgment name (NONE, rules) = (SOME name, rules) |
71 |
| new_judgment _ (SOME _, _) = error "Attempt to redeclare object-logic judgment"; |
|
11897 | 72 |
|
14226 | 73 |
fun add_final name thy = |
74 |
let |
|
75 |
val typ = case Sign.const_type (sign_of thy) name of |
|
15531 | 76 |
SOME T => T |
77 |
| NONE => error "Internal error in ObjectLogic.gen_add_judgment"; |
|
14226 | 78 |
in |
79 |
Theory.add_finals_i false [Const(name,typ)] thy |
|
80 |
end; |
|
81 |
||
11897 | 82 |
fun gen_add_judgment add_consts (name, T, syn) thy = |
14226 | 83 |
let |
84 |
val fullname = Sign.full_name (Theory.sign_of thy) name; |
|
85 |
in |
|
86 |
thy |
|
87 |
|> add_consts [(name, T, syn)] |
|
88 |
|> add_final fullname |
|
89 |
|> ObjectLogicData.map (new_judgment fullname) |
|
90 |
end; |
|
11897 | 91 |
|
92 |
in |
|
93 |
||
94 |
val add_judgment = gen_add_judgment Theory.add_consts; |
|
95 |
val add_judgment_i = gen_add_judgment Theory.add_consts_i; |
|
96 |
||
97 |
end; |
|
98 |
||
99 |
||
100 |
(* term operations *) |
|
101 |
||
102 |
fun judgment_name sg = |
|
103 |
(case ObjectLogicData.get_sg sg of |
|
15531 | 104 |
(SOME name, _) => name |
11897 | 105 |
| _ => raise TERM ("Unknown object-logic judgment", [])); |
106 |
||
107 |
fun is_judgment sg (Const (c, _) $ _) = c = judgment_name sg |
|
108 |
| is_judgment _ _ = false; |
|
109 |
||
110 |
fun drop_judgment sg (Abs (x, T, t)) = Abs (x, T, drop_judgment sg t) |
|
12479
ed46612ad7ec
drop_judgment: be graceful about undeclared judgment;
wenzelm
parents:
12371
diff
changeset
|
111 |
| drop_judgment sg (tm as (Const (c, _) $ t)) = |
ed46612ad7ec
drop_judgment: be graceful about undeclared judgment;
wenzelm
parents:
12371
diff
changeset
|
112 |
if (c = judgment_name sg handle TERM _ => false) then t else tm |
11897 | 113 |
| drop_judgment _ tm = tm; |
114 |
||
115 |
fun fixed_judgment sg x = |
|
116 |
let (*be robust wrt. low-level errors*) |
|
117 |
val c = judgment_name sg; |
|
14854 | 118 |
val aT = TFree ("'a", []); |
11897 | 119 |
val T = |
15973 | 120 |
if_none (Sign.const_type sg c) (aT --> propT) |
11897 | 121 |
|> Term.map_type_tvar (fn ((x, _), S) => TFree (x, S)); |
122 |
val U = Term.domain_type T handle Match => aT; |
|
123 |
in Const (c, T) $ Free (x, U) end; |
|
124 |
||
13376 | 125 |
fun assert_propT sg t = |
126 |
let val T = Term.fastype_of t |
|
127 |
in if T = propT then t else Const (judgment_name sg, T --> propT) $ t end; |
|
128 |
||
11897 | 129 |
|
130 |
||
131 |
(** treatment of meta-level connectives **) |
|
132 |
||
133 |
(* maintain rules *) |
|
134 |
||
135 |
val get_atomize = #1 o #2 o ObjectLogicData.get_sg; |
|
136 |
val get_rulify = #2 o #2 o ObjectLogicData.get_sg; |
|
137 |
||
12371 | 138 |
val add_atomize = ObjectLogicData.map o Library.apsnd o Library.apfst o Drule.add_rule; |
139 |
val add_rulify = ObjectLogicData.map o Library.apsnd o Library.apsnd o Drule.add_rule; |
|
11897 | 140 |
|
12371 | 141 |
fun declare_atomize (thy, th) = (add_atomize th thy, th); |
142 |
fun declare_rulify (thy, th) = (add_rulify th thy, th); |
|
11897 | 143 |
|
144 |
||
145 |
(* atomize *) |
|
146 |
||
15001
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
skalberg
parents:
14981
diff
changeset
|
147 |
fun rewrite_prems_tac rews i = PRIMITIVE (Drule.fconv_rule |
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
skalberg
parents:
14981
diff
changeset
|
148 |
(Drule.goals_conv (Library.equal i) |
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
skalberg
parents:
14981
diff
changeset
|
149 |
(Drule.forall_conv |
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
skalberg
parents:
14981
diff
changeset
|
150 |
(Drule.goals_conv (K true) (Tactic.rewrite true rews))))); |
11897 | 151 |
|
12807 | 152 |
fun atomize_term sg = |
13197
0567f4fd1415
Changed interface of MetaSimplifier.rewrite_term.
berghofe
parents:
12829
diff
changeset
|
153 |
drop_judgment sg o MetaSimplifier.rewrite_term sg (get_atomize sg) []; |
12729 | 154 |
|
13376 | 155 |
fun atomize_rule sg = Tactic.rewrite true (get_atomize sg); |
156 |
||
14743 | 157 |
(*Convert a natural-deduction rule into a formula (probably in FOL)*) |
158 |
fun atomize_thm th = |
|
159 |
rewrite_rule (get_atomize (Thm.sign_of_thm th)) th; |
|
160 |
||
11897 | 161 |
fun atomize_tac i st = |
12807 | 162 |
if Logic.has_meta_prems (Thm.prop_of st) i then |
11897 | 163 |
(rewrite_prems_tac (get_atomize (Thm.sign_of_thm st)) i) st |
164 |
else all_tac st; |
|
165 |
||
12829 | 166 |
fun full_atomize_tac i st = |
167 |
rewrite_goal_tac (get_atomize (Thm.sign_of_thm st)) i st; |
|
168 |
||
11897 | 169 |
fun atomize_goal i st = |
15531 | 170 |
(case Seq.pull (atomize_tac i st) of NONE => st | SOME (st', _) => st'); |
11897 | 171 |
|
172 |
||
173 |
(* rulify *) |
|
174 |
||
175 |
fun gen_rulify full thm = |
|
176 |
Tactic.simplify full (get_rulify (Thm.sign_of_thm thm)) thm |
|
12725 | 177 |
|> Drule.gen_all |> Drule.strip_shyps_warning |> Drule.zero_var_indexes; |
11897 | 178 |
|
179 |
val rulify = gen_rulify true; |
|
180 |
val rulify_no_asm = gen_rulify false; |
|
181 |
||
182 |
fun rule_format x = Drule.rule_attribute (fn _ => rulify) x; |
|
183 |
fun rule_format_no_asm x = Drule.rule_attribute (fn _ => rulify_no_asm) x; |
|
184 |
||
185 |
||
186 |
end; |