| author | wenzelm | 
| Wed, 11 Jul 2007 00:29:52 +0200 | |
| changeset 23730 | 8866c87d1a16 | 
| parent 23602 | 361e9c3a5e3a | 
| child 24039 | 273698405054 | 
| 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 | |
| 16449 | 12 | val judgment_name: theory -> string | 
| 13 | val is_judgment: theory -> term -> bool | |
| 14 | val drop_judgment: theory -> term -> term | |
| 15 | val fixed_judgment: theory -> string -> term | |
| 18121 | 16 | val ensure_propT: theory -> term -> term | 
| 23586 | 17 | val dest_judgment: cterm -> cterm | 
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 18 | val judgment_conv: conv -> conv | 
| 19261 | 19 | val is_elim: thm -> bool | 
| 18728 | 20 | val declare_atomize: attribute | 
| 21 | val declare_rulify: attribute | |
| 16449 | 22 | val atomize_term: theory -> term -> term | 
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 23 | val atomize: conv | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 24 | val atomize_prems: conv | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 25 | val atomize_prems_tac: int -> tactic | 
| 12829 | 26 | val full_atomize_tac: int -> tactic | 
| 18807 | 27 | val rulify_term: theory -> term -> term | 
| 28 | val rulify_tac: int -> tactic | |
| 11897 | 29 | val rulify: thm -> thm | 
| 30 | val rulify_no_asm: thm -> thm | |
| 18728 | 31 | val rule_format: attribute | 
| 32 | val rule_format_no_asm: attribute | |
| 11897 | 33 | end; | 
| 34 | ||
| 35 | structure ObjectLogic: OBJECT_LOGIC = | |
| 36 | struct | |
| 37 | ||
| 38 | ||
| 18825 | 39 | (** theory data **) | 
| 11897 | 40 | |
| 16449 | 41 | structure ObjectLogicData = TheoryDataFun | 
| 22846 | 42 | ( | 
| 11897 | 43 | type T = string option * (thm list * thm list); | 
| 22846 | 44 | val empty = (NONE, ([], [])); | 
| 11897 | 45 | val copy = I; | 
| 16449 | 46 | val extend = I; | 
| 11897 | 47 | |
| 15531 | 48 | fun merge_judgment (SOME x, SOME y) = | 
| 23586 | 49 | if (x: string) = 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 | |
| 16449 | 52 | fun merge _ ((judgment1, (atomize1, rulify1)), (judgment2, (atomize2, rulify2))) = | 
| 11897 | 53 | (merge_judgment (judgment1, judgment2), | 
| 54 | (Drule.merge_rules (atomize1, atomize2), Drule.merge_rules (rulify1, rulify2))); | |
| 22846 | 55 | ); | 
| 15801 | 56 | |
| 11897 | 57 | |
| 58 | ||
| 59 | (** generic treatment of judgments -- with a single argument only **) | |
| 60 | ||
| 18825 | 61 | (* add judgment *) | 
| 11897 | 62 | |
| 63 | local | |
| 64 | ||
| 15531 | 65 | fun new_judgment name (NONE, rules) = (SOME name, rules) | 
| 66 | | new_judgment _ (SOME _, _) = error "Attempt to redeclare object-logic judgment"; | |
| 11897 | 67 | |
| 16449 | 68 | fun gen_add_judgment add_consts (bname, T, mx) thy = | 
| 69 | let val c = Sign.full_name thy (Syntax.const_name bname mx) in | |
| 14226 | 70 | thy | 
| 16449 | 71 | |> add_consts [(bname, T, mx)] | 
| 72 | |> (fn thy' => Theory.add_finals_i false [(Const (c, Sign.the_const_type thy' c))] thy') | |
| 73 | |> ObjectLogicData.map (new_judgment c) | |
| 14226 | 74 | end; | 
| 11897 | 75 | |
| 76 | in | |
| 77 | ||
| 22796 | 78 | val add_judgment = gen_add_judgment Sign.add_consts; | 
| 79 | val add_judgment_i = gen_add_judgment Sign.add_consts_i; | |
| 11897 | 80 | |
| 81 | end; | |
| 82 | ||
| 83 | ||
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 84 | (* judgments *) | 
| 11897 | 85 | |
| 16449 | 86 | fun judgment_name thy = | 
| 87 | (case ObjectLogicData.get thy of | |
| 15531 | 88 | (SOME name, _) => name | 
| 11897 | 89 |   | _ => raise TERM ("Unknown object-logic judgment", []));
 | 
| 90 | ||
| 16449 | 91 | fun is_judgment thy (Const (c, _) $ _) = c = judgment_name thy | 
| 11897 | 92 | | is_judgment _ _ = false; | 
| 93 | ||
| 16449 | 94 | fun drop_judgment thy (Abs (x, T, t)) = Abs (x, T, drop_judgment thy t) | 
| 95 | | drop_judgment thy (tm as (Const (c, _) $ t)) = | |
| 96 | if (c = judgment_name thy handle TERM _ => false) then t else tm | |
| 11897 | 97 | | drop_judgment _ tm = tm; | 
| 98 | ||
| 16449 | 99 | fun fixed_judgment thy x = | 
| 11897 | 100 | let (*be robust wrt. low-level errors*) | 
| 16449 | 101 | val c = judgment_name thy; | 
| 14854 | 102 |     val aT = TFree ("'a", []);
 | 
| 11897 | 103 | val T = | 
| 18939 | 104 | the_default (aT --> propT) (Sign.const_type thy c) | 
| 11897 | 105 | |> Term.map_type_tvar (fn ((x, _), S) => TFree (x, S)); | 
| 106 | val U = Term.domain_type T handle Match => aT; | |
| 107 | in Const (c, T) $ Free (x, U) end; | |
| 108 | ||
| 18121 | 109 | fun ensure_propT thy t = | 
| 13376 | 110 | let val T = Term.fastype_of t | 
| 16449 | 111 | in if T = propT then t else Const (judgment_name thy, T --> propT) $ t end; | 
| 13376 | 112 | |
| 23586 | 113 | fun dest_judgment ct = | 
| 114 | if is_judgment (Thm.theory_of_cterm ct) (Thm.term_of ct) | |
| 115 | then Thm.dest_arg ct | |
| 116 |   else raise CTERM ("dest_judgment", [ct]);
 | |
| 117 | ||
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 118 | fun judgment_conv cv ct = | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 119 | 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 | 120 | then Conv.arg_conv cv ct | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 121 |   else raise CTERM ("judgment_conv", [ct]);
 | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 122 | |
| 11897 | 123 | |
| 19261 | 124 | (* elimination rules *) | 
| 125 | ||
| 126 | fun is_elim rule = | |
| 127 | let | |
| 128 | val thy = Thm.theory_of_thm rule; | |
| 129 | val concl = Thm.concl_of rule; | |
| 130 | in | |
| 131 | Term.is_Var (drop_judgment thy concl) andalso | |
| 132 | exists (fn prem => concl aconv Logic.strip_assums_concl prem) (Thm.prems_of rule) | |
| 133 | end; | |
| 134 | ||
| 135 | ||
| 11897 | 136 | |
| 137 | (** treatment of meta-level connectives **) | |
| 138 | ||
| 139 | (* maintain rules *) | |
| 140 | ||
| 16449 | 141 | val get_atomize = #1 o #2 o ObjectLogicData.get; | 
| 142 | val get_rulify = #2 o #2 o ObjectLogicData.get; | |
| 11897 | 143 | |
| 22846 | 144 | val add_atomize = ObjectLogicData.map o apsnd o apfst o Drule.add_rule; | 
| 145 | val add_rulify = ObjectLogicData.map o apsnd o apsnd o Drule.add_rule; | |
| 11897 | 146 | |
| 22846 | 147 | val declare_atomize = Thm.declaration_attribute (fn th => Context.mapping (add_atomize th) I); | 
| 148 | val declare_rulify = Thm.declaration_attribute (fn th => Context.mapping (add_rulify th) I); | |
| 149 | ||
| 150 | val _ = Context.add_setup (add_rulify Drule.norm_hhf_eq); | |
| 11897 | 151 | |
| 152 | ||
| 153 | (* atomize *) | |
| 154 | ||
| 16449 | 155 | fun atomize_term thy = | 
| 156 | drop_judgment thy o MetaSimplifier.rewrite_term thy (get_atomize thy) []; | |
| 12729 | 157 | |
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 158 | fun atomize ct = | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 159 | MetaSimplifier.rewrite true (get_atomize (Thm.theory_of_cterm ct)) ct; | 
| 14743 | 160 | |
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 161 | fun atomize_prems ct = | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 162 | if Logic.has_meta_prems (Thm.term_of ct) then | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 163 | Conv.forall_conv ~1 (Conv.prems_conv ~1 atomize) ct | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 164 | else Conv.all_conv ct; | 
| 11897 | 165 | |
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 166 | val atomize_prems_tac = CONVERSION atomize_prems; | 
| 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 167 | val full_atomize_tac = CONVERSION atomize; | 
| 11897 | 168 | |
| 169 | ||
| 170 | (* rulify *) | |
| 171 | ||
| 18807 | 172 | fun rulify_term thy = MetaSimplifier.rewrite_term thy (get_rulify thy) []; | 
| 23540 | 173 | fun rulify_tac i st = MetaSimplifier.rewrite_goal_tac (get_rulify (Thm.theory_of_thm st)) i st; | 
| 18807 | 174 | |
| 11897 | 175 | fun gen_rulify full thm = | 
| 21708 | 176 | MetaSimplifier.simplify full (get_rulify (Thm.theory_of_thm thm)) thm | 
| 20912 | 177 | |> Drule.gen_all |> Thm.strip_shyps |> Drule.zero_var_indexes; | 
| 11897 | 178 | |
| 179 | val rulify = gen_rulify true; | |
| 180 | val rulify_no_asm = gen_rulify false; | |
| 181 | ||
| 18728 | 182 | fun rule_format x = Thm.rule_attribute (fn _ => rulify) x; | 
| 183 | fun rule_format_no_asm x = Thm.rule_attribute (fn _ => rulify_no_asm) x; | |
| 11897 | 184 | |
| 185 | end; |