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