| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 01 Dec 2023 20:32:34 +0100 | |
| changeset 79101 | 4e47b34fbb8e | 
| parent 74561 | 8e6c973003c8 | 
| child 81519 | cdc43c0fdbfc | 
| 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 | |
| 59970 | 9 | val get_base_sort: Proof.context -> sort option | 
| 25497 | 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 add_judgment: binding * typ * mixfix -> theory -> theory | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 12 | val add_judgment_cmd: binding * string * mixfix -> theory -> theory | 
| 59970 | 13 | val judgment_name: Proof.context -> string | 
| 74341 | 14 | val judgment_const: Proof.context -> string * typ | 
| 59970 | 15 | val is_judgment: Proof.context -> term -> bool | 
| 16 | val drop_judgment: Proof.context -> term -> term | |
| 17 | val fixed_judgment: Proof.context -> string -> term | |
| 18 | val ensure_propT: Proof.context -> term -> term | |
| 19 | val dest_judgment: Proof.context -> cterm -> cterm | |
| 20 | val judgment_conv: Proof.context -> conv -> conv | |
| 69988 | 21 | val is_propositional: Proof.context -> typ -> bool | 
| 59970 | 22 | val elim_concl: Proof.context -> thm -> term option | 
| 18728 | 23 | val declare_atomize: attribute | 
| 24 | val declare_rulify: attribute | |
| 59970 | 25 | val atomize_term: Proof.context -> term -> term | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 26 | val atomize: Proof.context -> conv | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 27 | val atomize_prems: Proof.context -> conv | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 28 | val atomize_prems_tac: Proof.context -> int -> tactic | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 29 | val full_atomize_tac: Proof.context -> int -> tactic | 
| 59970 | 30 | val rulify_term: Proof.context -> term -> term | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 31 | val rulify_tac: Proof.context -> int -> tactic | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 32 | val rulify: Proof.context -> thm -> thm | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 33 | val rulify_no_asm: Proof.context -> thm -> thm | 
| 18728 | 34 | val rule_format: attribute | 
| 35 | val rule_format_no_asm: attribute | |
| 11897 | 36 | end; | 
| 37 | ||
| 35625 | 38 | structure Object_Logic: OBJECT_LOGIC = | 
| 11897 | 39 | struct | 
| 40 | ||
| 59970 | 41 | (** context 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 | |
| 59970 | 51 | structure Data = Generic_Data | 
| 22846 | 52 | ( | 
| 25497 | 53 | type T = data; | 
| 54 | val empty = make_data (NONE, NONE, ([], [])); | |
| 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" | |
| 41493 | 58 | | merge_opt _ data = merge_options data; | 
| 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: 
23602diff
changeset | 64 | (Thm.merge_thms (atomize1, atomize2), Thm.merge_thms (rulify1, rulify2))); | 
| 22846 | 65 | ); | 
| 15801 | 66 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36610diff
changeset | 67 | fun map_data f = Data.map (fn (Data {base_sort, judgment, atomize_rulify}) =>
 | 
| 25497 | 68 | make_data (f (base_sort, judgment, atomize_rulify))); | 
| 69 | ||
| 59970 | 70 | fun get_data ctxt = Data.get (Context.Proof ctxt) |> (fn Data args => args); | 
| 25497 | 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 | ||
| 59970 | 80 | fun add_base_sort S = | 
| 81 | (Context.theory_map o 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)); | |
| 25497 | 84 | |
| 85 | ||
| 18825 | 86 | (* add judgment *) | 
| 11897 | 87 | |
| 88 | local | |
| 89 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
29606diff
changeset | 90 | fun gen_add_judgment add_consts (b, T, mx) thy = | 
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61246diff
changeset | 91 | let | 
| 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61246diff
changeset | 92 | val c = Sign.full_name thy b; | 
| 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61246diff
changeset | 93 | val thy' = thy |> add_consts [(b, T, mx)]; | 
| 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61246diff
changeset | 94 | in | 
| 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61246diff
changeset | 95 | thy' | 
| 74340 
e098fa45bfe0
proper Sign.const_typargs via Theory.add_deps_const/Theory.const_dep;
 wenzelm parents: 
70471diff
changeset | 96 | |> Theory.add_deps_const c | 
| 59970 | 97 | |> (Context.theory_map o map_data) (fn (base_sort, judgment, atomize_rulify) => | 
| 25497 | 98 | if is_some judgment then error "Attempt to redeclare object-logic judgment" | 
| 99 | else (base_sort, SOME c, atomize_rulify)) | |
| 14226 | 100 | end; | 
| 11897 | 101 | |
| 102 | in | |
| 103 | ||
| 56239 | 104 | val add_judgment = gen_add_judgment Sign.add_consts; | 
| 105 | val add_judgment_cmd = gen_add_judgment Sign.add_consts_cmd; | |
| 11897 | 106 | |
| 107 | end; | |
| 108 | ||
| 109 | ||
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 110 | (* judgments *) | 
| 11897 | 111 | |
| 59970 | 112 | fun judgment_name ctxt = | 
| 113 | (case #judgment (get_data ctxt) of | |
| 25497 | 114 | SOME name => name | 
| 11897 | 115 |   | _ => raise TERM ("Unknown object-logic judgment", []));
 | 
| 116 | ||
| 74341 | 117 | fun judgment_const ctxt = | 
| 118 | let | |
| 119 | val thy = Proof_Context.theory_of ctxt; | |
| 120 | val c = judgment_name ctxt; | |
| 121 | val T = Sign.the_const_type thy c; | |
| 122 | in (c, T) end; | |
| 123 | ||
| 74344 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 124 | fun is_judgment ctxt = | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 125 | let val name = judgment_name ctxt | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 126 | in fn Const (c, _) $ _ => c = name | _ => false end; | 
| 11897 | 127 | |
| 74344 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 128 | fun drop_judgment ctxt = | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 129 | let | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 130 | val name = judgment_name ctxt; | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 131 | fun drop (Abs (x, T, t)) = Abs (x, T, drop t) | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 132 | | drop (t as Const (c, _) $ u) = if c = name then u else t | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 133 | | drop t = t; | 
| 
1c2c0380d58b
clarified partial application: immediate check of object-logic, and avoidance of context within closure;
 wenzelm parents: 
74341diff
changeset | 134 | in drop end handle TERM _ => I; | 
| 11897 | 135 | |
| 59970 | 136 | fun fixed_judgment ctxt x = | 
| 11897 | 137 | let (*be robust wrt. low-level errors*) | 
| 59970 | 138 | val c = judgment_name ctxt; | 
| 70387 | 139 | val aT = Term.aT []; | 
| 11897 | 140 | val T = | 
| 59970 | 141 | the_default (aT --> propT) (Sign.const_type (Proof_Context.theory_of ctxt) c) | 
| 60443 | 142 | |> Term.map_type_tvar (fn ((a, _), S) => TFree (a, S)); | 
| 11897 | 143 | val U = Term.domain_type T handle Match => aT; | 
| 144 | in Const (c, T) $ Free (x, U) end; | |
| 145 | ||
| 59970 | 146 | fun ensure_propT ctxt t = | 
| 13376 | 147 | let val T = Term.fastype_of t | 
| 59970 | 148 | in if T = propT then t else Const (judgment_name ctxt, T --> propT) $ t end; | 
| 13376 | 149 | |
| 59970 | 150 | fun dest_judgment ctxt ct = | 
| 151 | if is_judgment ctxt (Thm.term_of ct) | |
| 23586 | 152 | then Thm.dest_arg ct | 
| 153 |   else raise CTERM ("dest_judgment", [ct]);
 | |
| 154 | ||
| 59970 | 155 | fun judgment_conv ctxt cv ct = | 
| 156 | if is_judgment ctxt (Thm.term_of ct) | |
| 23566 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 157 | then Conv.arg_conv cv ct | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 158 |   else raise CTERM ("judgment_conv", [ct]);
 | 
| 
b65692d4adcd
replaced HOLogic.Trueprop_conv by ObjectLogic.judgment_conv;
 wenzelm parents: 
23540diff
changeset | 159 | |
| 69988 | 160 | fun is_propositional ctxt T = | 
| 161 | T = propT orelse | |
| 162 |     let val x = Free (singleton (Variable.variant_frees ctxt []) ("x", T))
 | |
| 163 | in can (fn () => Syntax.check_term ctxt (ensure_propT ctxt x)) () end; | |
| 164 | ||
| 11897 | 165 | |
| 19261 | 166 | (* elimination rules *) | 
| 167 | ||
| 59970 | 168 | fun elim_concl ctxt rule = | 
| 19261 | 169 | let | 
| 170 | val concl = Thm.concl_of rule; | |
| 59970 | 171 | val C = drop_judgment ctxt concl; | 
| 19261 | 172 | in | 
| 41581 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 wenzelm parents: 
41493diff
changeset | 173 | if Term.is_Var C andalso | 
| 19261 | 174 | exists (fn prem => concl aconv Logic.strip_assums_concl prem) (Thm.prems_of rule) | 
| 41581 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 wenzelm parents: 
41493diff
changeset | 175 | then SOME C else NONE | 
| 19261 | 176 | end; | 
| 177 | ||
| 178 | ||
| 11897 | 179 | |
| 180 | (** treatment of meta-level connectives **) | |
| 181 | ||
| 182 | (* maintain rules *) | |
| 183 | ||
| 70471 | 184 | fun get_atomize_rulify f ctxt = map (Thm.transfer' ctxt) (f (#atomize_rulify (get_data ctxt))); | 
| 185 | val get_atomize = get_atomize_rulify #1; | |
| 186 | val get_rulify = get_atomize_rulify #2; | |
| 11897 | 187 | |
| 25497 | 188 | fun add_atomize th = map_data (fn (base_sort, judgment, (atomize, rulify)) => | 
| 61092 | 189 | (base_sort, judgment, (Thm.add_thm (Thm.trim_context th) atomize, rulify))); | 
| 25497 | 190 | |
| 191 | fun add_rulify th = map_data (fn (base_sort, judgment, (atomize, rulify)) => | |
| 61092 | 192 | (base_sort, judgment, (atomize, Thm.add_thm (Thm.trim_context th) rulify))); | 
| 11897 | 193 | |
| 59970 | 194 | val declare_atomize = Thm.declaration_attribute add_atomize; | 
| 195 | val declare_rulify = Thm.declaration_attribute add_rulify; | |
| 22846 | 196 | |
| 59970 | 197 | val _ = Theory.setup (fold (Context.theory_map o add_rulify) Drule.norm_hhf_eqs); | 
| 11897 | 198 | |
| 199 | ||
| 200 | (* atomize *) | |
| 201 | ||
| 59970 | 202 | fun atomize_term ctxt = | 
| 203 | drop_judgment ctxt o | |
| 204 | Raw_Simplifier.rewrite_term (Proof_Context.theory_of ctxt) (get_atomize ctxt) []; | |
| 12729 | 205 | |
| 59970 | 206 | fun atomize ctxt = Raw_Simplifier.rewrite ctxt true (get_atomize ctxt); | 
| 14743 | 207 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 208 | fun atomize_prems ctxt ct = | 
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 209 | if Logic.has_meta_prems (Thm.term_of ct) then | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 210 | Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize) ctxt ct | 
| 23602 
361e9c3a5e3a
tuned interfaces: atomize, atomize_prems, atomize_prems_tac;
 wenzelm parents: 
23586diff
changeset | 211 | else Conv.all_conv ct; | 
| 11897 | 212 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 213 | val atomize_prems_tac = CONVERSION o atomize_prems; | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 214 | val full_atomize_tac = CONVERSION o atomize; | 
| 11897 | 215 | |
| 216 | ||
| 217 | (* rulify *) | |
| 218 | ||
| 59970 | 219 | fun rulify_term ctxt = | 
| 220 | Raw_Simplifier.rewrite_term (Proof_Context.theory_of ctxt) (get_rulify ctxt) []; | |
| 221 | ||
| 222 | fun rulify_tac ctxt = rewrite_goal_tac ctxt (get_rulify ctxt); | |
| 18807 | 223 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 224 | fun gen_rulify full ctxt = | 
| 59970 | 225 | Conv.fconv_rule (Raw_Simplifier.rewrite ctxt full (get_rulify ctxt)) | 
| 60822 | 226 | #> Variable.gen_all ctxt | 
| 59647 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
56239diff
changeset | 227 | #> Thm.strip_shyps | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
56239diff
changeset | 228 | #> Drule.zero_var_indexes; | 
| 11897 | 229 | |
| 230 | val rulify = gen_rulify true; | |
| 231 | val rulify_no_asm = gen_rulify false; | |
| 232 | ||
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61255diff
changeset | 233 | val rule_format = Thm.rule_attribute [] (rulify o Context.proof_of); | 
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61255diff
changeset | 234 | val rule_format_no_asm = Thm.rule_attribute [] (rulify_no_asm o Context.proof_of); | 
| 11897 | 235 | |
| 236 | end; |