| author | nipkow | 
| Mon, 09 Mar 2009 09:24:50 +0100 | |
| changeset 30371 | b3a60d828de0 | 
| parent 30344 | 10a67c5ddddb | 
| child 33090 | 2f05a1feac21 | 
| 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: 
29606diff
changeset | 11 | val typedecl: binding * string list * mixfix -> theory -> typ * theory | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 12 | val add_judgment: binding * typ * mixfix -> theory -> theory | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
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: 
23540diff
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: 
23586diff
changeset | 25 | val atomize: conv | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 26 | val atomize_prems: conv | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
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 | |
| 16449 | 50 | structure ObjectLogicData = TheoryDataFun | 
| 22846 | 51 | ( | 
| 25497 | 52 | type T = data; | 
| 53 | val empty = make_data (NONE, NONE, ([], [])); | |
| 11897 | 54 | val copy = I; | 
| 16449 | 55 | val extend = I; | 
| 11897 | 56 | |
| 25497 | 57 | fun merge_opt eq (SOME x, SOME y) = | 
| 58 | if eq (x, y) then SOME x else error "Attempt to merge different object-logics" | |
| 59 | | merge_opt _ (x, y) = if is_some x then x else y; | |
| 11897 | 60 | |
| 25497 | 61 | fun merge _ | 
| 62 |      (Data {base_sort = base_sort1, judgment = judgment1, atomize_rulify = (atomize1, rulify1)},
 | |
| 63 |       Data {base_sort = base_sort2, judgment = judgment2, atomize_rulify = (atomize2, rulify2)}) =
 | |
| 64 | 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: 
23602diff
changeset | 65 | (Thm.merge_thms (atomize1, atomize2), Thm.merge_thms (rulify1, rulify2))); | 
| 22846 | 66 | ); | 
| 15801 | 67 | |
| 25497 | 68 | fun map_data f = ObjectLogicData.map (fn (Data {base_sort, judgment, atomize_rulify}) =>
 | 
| 69 | make_data (f (base_sort, judgment, atomize_rulify))); | |
| 70 | ||
| 71 | fun get_data thy = ObjectLogicData.get thy |> (fn Data args => args); | |
| 72 | ||
| 11897 | 73 | |
| 74 | ||
| 75 | (** generic treatment of judgments -- with a single argument only **) | |
| 76 | ||
| 25497 | 77 | (* base_sort *) | 
| 78 | ||
| 79 | val get_base_sort = #base_sort o get_data; | |
| 80 | ||
| 81 | fun add_base_sort S = map_data (fn (base_sort, judgment, atomize_rulify) => | |
| 82 | if is_some base_sort then error "Attempt to redeclare object-logic base sort" | |
| 83 | else (SOME S, judgment, atomize_rulify)); | |
| 84 | ||
| 85 | ||
| 86 | (* typedecl *) | |
| 87 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 88 | fun typedecl (a, vs, mx) thy = | 
| 25497 | 89 | let | 
| 90 | val base_sort = get_base_sort thy; | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 91 | val b = Binding.map_name (Syntax.type_name mx) a; | 
| 25497 | 92 | val _ = has_duplicates (op =) vs andalso | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 93 |       error ("Duplicate parameters in type declaration: " ^ quote (Binding.str_of b));
 | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 94 | val name = Sign.full_name thy b; | 
| 25497 | 95 | val n = length vs; | 
| 96 | val T = Type (name, map (fn v => TFree (v, [])) vs); | |
| 97 | in | |
| 98 | thy | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 99 | |> Sign.add_types [(a, n, mx)] | 
| 25497 | 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 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 109 | fun gen_add_judgment add_consts (b, T, mx) thy = | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 110 | let val c = Sign.full_name thy (Binding.map_name (Syntax.const_name mx) b) in | 
| 14226 | 111 | thy | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 112 | |> add_consts [(b, T, mx)] | 
| 25018 
fac2ceba75b4
replaced obsolete Theory.add_finals_i by Theory.add_deps;
 wenzelm parents: 
24848diff
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 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 121 | val add_judgment = gen_add_judgment Sign.add_consts_i; | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 122 | val add_judgment_cmd = gen_add_judgment Sign.add_consts; | 
| 11897 | 123 | |
| 124 | end; | |
| 125 | ||
| 126 | ||
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
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: 
23540diff
changeset | 161 | fun judgment_conv cv ct = | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
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: 
23540diff
changeset | 163 | then Conv.arg_conv cv ct | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 164 |   else raise CTERM ("judgment_conv", [ct]);
 | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
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 | ||
| 28620 | 196 | val _ = Context.>> (Context.map_theory (fold add_rulify Drule.norm_hhf_eqs)); | 
| 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: 
23586diff
changeset | 204 | fun atomize ct = | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
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: 
23586diff
changeset | 207 | fun atomize_prems ct = | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
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: 
23586diff
changeset | 211 | else Conv.all_conv ct; | 
| 11897 | 212 | |
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 213 | val atomize_prems_tac = CONVERSION atomize_prems; | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
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; |