| author | wenzelm | 
| Sun, 29 Sep 2024 21:16:17 +0200 | |
| changeset 81008 | d0cd220d8e8b | 
| parent 77879 | dd222e2af01a | 
| child 81954 | 6f2bcdfa9a19 | 
| permissions | -rw-r--r-- | 
| 59975 | 1 | (* Title: HOL/Library/cconv.ML | 
| 2 | Author: Christoph Traut, Lars Noschinski, TU Muenchen | |
| 3 | ||
| 74549 | 4 | Conditional conversions. | 
| 59975 | 5 | *) | 
| 6 | ||
| 59739 | 7 | infix 1 then_cconv | 
| 8 | infix 0 else_cconv | |
| 9 | ||
| 10 | signature BASIC_CCONV = | |
| 11 | sig | |
| 74549 | 12 | type cconv = conv | 
| 59739 | 13 | val then_cconv: cconv * cconv -> cconv | 
| 14 | val else_cconv: cconv * cconv -> cconv | |
| 15 | val CCONVERSION: cconv -> int -> tactic | |
| 16 | end | |
| 17 | ||
| 18 | signature CCONV = | |
| 19 | sig | |
| 20 | include BASIC_CCONV | |
| 21 | val no_cconv: cconv | |
| 22 | val all_cconv: cconv | |
| 23 | val first_cconv: cconv list -> cconv | |
| 24 | val abs_cconv: (cterm * Proof.context -> cconv) -> Proof.context -> cconv | |
| 25 | val combination_cconv: cconv -> cconv -> cconv | |
| 26 | val comb_cconv: cconv -> cconv | |
| 27 | val arg_cconv: cconv -> cconv | |
| 28 | val fun_cconv: cconv -> cconv | |
| 29 | val arg1_cconv: cconv -> cconv | |
| 30 | val fun2_cconv: cconv -> cconv | |
| 31 | val rewr_cconv: thm -> cconv | |
| 32 | val rewrs_cconv: thm list -> cconv | |
| 33 | val params_cconv: int -> (Proof.context -> cconv) -> Proof.context -> cconv | |
| 34 | val prems_cconv: int -> cconv -> cconv | |
| 60054 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 35 | val with_prems_cconv: int -> cconv -> cconv | 
| 59739 | 36 | val concl_cconv: int -> cconv -> cconv | 
| 37 | val fconv_rule: cconv -> thm -> thm | |
| 38 | val gconv_rule: cconv -> int -> thm -> thm | |
| 39 | end | |
| 40 | ||
| 41 | structure CConv : CCONV = | |
| 42 | struct | |
| 43 | ||
| 74549 | 44 | type cconv = conv | 
| 45 | ||
| 59739 | 46 | val concl_lhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_lhs | 
| 47 | val concl_rhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_rhs | |
| 48 | ||
| 49 | fun transitive th1 th2 = Drule.transitive_thm OF [th1, th2] | |
| 50 | ||
| 51 | fun abstract_rule_thm n = | |
| 52 | let | |
| 74624 | 53 |     val eq = \<^cprop>\<open>\<And>x::'a::{}. (s::'a \<Rightarrow> 'b::{}) x \<equiv> t x\<close>
 | 
| 54 |     val x = \<^cterm>\<open>x::'a::{}\<close>
 | |
| 55 | val thm = | |
| 56 | Thm.assume eq | |
| 59739 | 57 | |> Thm.forall_elim x | 
| 58 | |> Thm.abstract_rule n x | |
| 59 | |> Thm.implies_intr eq | |
| 60 | in Drule.export_without_context thm end | |
| 61 | ||
| 62 | val no_cconv = Conv.no_conv | |
| 63 | val all_cconv = Conv.all_conv | |
| 74549 | 64 | val op else_cconv = Conv.else_conv | 
| 59739 | 65 | |
| 66 | fun (cv1 then_cconv cv2) ct = | |
| 67 | let | |
| 68 | val eq1 = cv1 ct | |
| 69 | val eq2 = cv2 (concl_rhs_of eq1) | |
| 70 | in | |
| 71 | if Thm.is_reflexive eq1 then eq2 | |
| 72 | else if Thm.is_reflexive eq2 then eq1 | |
| 73 | else transitive eq1 eq2 | |
| 74 | end | |
| 75 | ||
| 76 | fun first_cconv cvs = fold_rev (curry op else_cconv) cvs no_cconv | |
| 77 | ||
| 78 | fun rewr_cconv rule ct = | |
| 79 | let | |
| 80 | val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule | |
| 81 | val lhs = concl_lhs_of rule1 | |
| 82 | val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1 | |
| 74624 | 83 | val rule3 = | 
| 84 | Thm.instantiate (Thm.match (lhs, ct)) rule2 | |
| 85 |         handle Pattern.MATCH => raise CTERM ("rewr_cconv", [lhs, ct])
 | |
| 60048 
e9c30726ca8e
rewr_cconv: ignore premises when tuning conclusion
 noschinl parents: 
59975diff
changeset | 86 | val concl = rule3 |> Thm.cprop_of |> Drule.strip_imp_concl | 
| 59739 | 87 | val rule4 = | 
| 60048 
e9c30726ca8e
rewr_cconv: ignore premises when tuning conclusion
 noschinl parents: 
59975diff
changeset | 88 | if Thm.dest_equals_lhs concl aconvc ct then rule3 | 
| 74624 | 89 | else | 
| 90 | let val ceq = Thm.dest_fun2 concl | |
| 91 | in rule3 RS Thm.trivial (Thm.mk_binop ceq ct (Thm.dest_equals_rhs concl)) end | |
| 59739 | 92 | in | 
| 93 | transitive rule4 (Thm.beta_conversion true (concl_rhs_of rule4)) | |
| 94 | end | |
| 95 | ||
| 96 | fun rewrs_cconv rules = first_cconv (map rewr_cconv rules) | |
| 97 | ||
| 98 | fun combination_cconv cv1 cv2 cterm = | |
| 74624 | 99 | let val (l, r) = Thm.dest_comb cterm in | 
| 100 |     @{lemma \<open>f \<equiv> g \<Longrightarrow> s \<equiv> t \<Longrightarrow> f s \<equiv> g t\<close> for f g :: \<open>'a::{} \<Rightarrow> 'b::{}\<close> by simp}
 | |
| 101 | OF [cv1 l, cv2 r] | |
| 102 | end | |
| 59739 | 103 | |
| 104 | fun comb_cconv cv = combination_cconv cv cv | |
| 105 | ||
| 106 | fun fun_cconv conversion = | |
| 107 | combination_cconv conversion all_cconv | |
| 108 | ||
| 109 | fun arg_cconv conversion = | |
| 110 | combination_cconv all_cconv conversion | |
| 111 | ||
| 112 | fun abs_cconv cv ctxt ct = | |
| 113 | (case Thm.term_of ct of | |
| 114 | Abs (x, _, _) => | |
| 115 | let | |
| 116 | (* Instantiate the rule properly and apply it to the eq theorem. *) | |
| 74521 | 117 | fun abstract_rule v eq = | 
| 59739 | 118 | let | 
| 119 | (* Take a variable v and an equality theorem of form: | |
| 60049 | 120 | P1 \<Longrightarrow> ... \<Longrightarrow> Pn \<Longrightarrow> L v \<equiv> R v | 
| 59739 | 121 | And build a term of form: | 
| 60049 | 122 | \<And>v. (\<lambda>x. L x) v \<equiv> (\<lambda>x. R x) v *) | 
| 74521 | 123 | fun mk_concl eq = | 
| 59739 | 124 | let | 
| 74528 | 125 | fun abs t = lambda v t $ v | 
| 126 | fun equals_cong f = Logic.dest_equals #> apply2 f #> Logic.mk_equals | |
| 59739 | 127 | in | 
| 128 | Thm.concl_of eq | |
| 129 | |> equals_cong abs | |
| 74528 | 130 | |> Logic.all v | 
| 131 | |> Thm.cterm_of ctxt | |
| 59739 | 132 | end | 
| 133 | val rule = abstract_rule_thm x | |
| 74528 | 134 | val inst = Thm.match (hd (Drule.cprems_of rule), mk_concl eq) | 
| 77879 | 135 | val gen = (Names.empty, Names.make1_set (#1 (dest_Free v))) | 
| 59739 | 136 | in | 
| 74528 | 137 | (Drule.instantiate_normalize inst rule OF [Drule.generalize gen eq]) | 
| 59739 | 138 | |> Drule.zero_var_indexes | 
| 139 | end | |
| 140 | ||
| 141 | (* Destruct the abstraction and apply the conversion. *) | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74521diff
changeset | 142 | val ((v, ct'), ctxt') = Variable.dest_abs_cterm ct ctxt | 
| 59739 | 143 | val eq = cv (v, ctxt') ct' | 
| 144 | in | |
| 145 | if Thm.is_reflexive eq | |
| 146 | then all_cconv ct | |
| 74521 | 147 | else abstract_rule (Thm.term_of v) eq | 
| 59739 | 148 | end | 
| 149 |    | _ => raise CTERM ("abs_cconv", [ct]))
 | |
| 150 | ||
| 151 | val arg1_cconv = fun_cconv o arg_cconv | |
| 152 | val fun2_cconv = fun_cconv o fun_cconv | |
| 153 | ||
| 154 | (* conversions on HHF rules *) | |
| 155 | ||
| 60049 | 156 | (*rewrite B in \<And>x1 ... xn. B*) | 
| 59739 | 157 | fun params_cconv n cv ctxt ct = | 
| 158 | if n <> 0 andalso Logic.is_all (Thm.term_of ct) | |
| 159 | then arg_cconv (abs_cconv (params_cconv (n - 1) cv o #2) ctxt) ct | |
| 160 | else cv ctxt ct | |
| 161 | ||
| 74528 | 162 | (* FIXME: This code behaves not exactly like Conv.prems_cconv does. *) | 
| 60049 | 163 | (*rewrite the A's in A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> B*) | 
| 59739 | 164 | fun prems_cconv 0 cv ct = cv ct | 
| 165 | | prems_cconv n cv ct = | |
| 166 | (case ct |> Thm.term_of of | |
| 74624 | 167 | \<^Const_>\<open>Pure.imp for _ _\<close> => | 
| 59975 | 168 | ((if n = 1 then fun_cconv else I) o arg_cconv) (prems_cconv (n-1) cv) ct | 
| 59739 | 169 | | _ => cv ct) | 
| 170 | ||
| 74607 | 171 | fun imp_cong A = | 
| 172 | \<^instantiate>\<open>A in | |
| 173 | lemma (schematic) \<open>(PROP A \<Longrightarrow> PROP B \<equiv> PROP C) \<Longrightarrow> (PROP A \<Longrightarrow> PROP B) \<equiv> (PROP A \<Longrightarrow> PROP C)\<close> | |
| 174 | by (fact Pure.imp_cong)\<close> | |
| 60050 | 175 | |
| 60049 | 176 | (*rewrite B in A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> B*) | 
| 59739 | 177 | fun concl_cconv 0 cv ct = cv ct | 
| 178 | | concl_cconv n cv ct = | |
| 60050 | 179 | (case try Thm.dest_implies ct of | 
| 180 | NONE => cv ct | |
| 74607 | 181 | | SOME (A,B) => (concl_cconv (n-1) cv B) RS imp_cong A) | 
| 59739 | 182 | |
| 74624 | 183 | (* Rewrite A in A \<Longrightarrow> A1 \<Longrightarrow> An \<Longrightarrow> B. | 
| 60054 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 184 | The premises of the resulting theorem assume A1, ..., An | 
| 74624 | 185 | *) | 
| 74607 | 186 | local | 
| 187 | ||
| 188 | fun rewr_imp C = | |
| 189 | \<^instantiate>\<open>C in | |
| 190 | lemma (schematic) \<open>PROP A \<equiv> PROP B \<Longrightarrow> (PROP A \<Longrightarrow> PROP C) \<equiv> (PROP B \<Longrightarrow> PROP C)\<close> by simp\<close> | |
| 191 | ||
| 192 | fun cut_rl A = | |
| 193 | \<^instantiate>\<open>A in | |
| 194 | lemma (schematic) \<open>(PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP A \<Longrightarrow> PROP B\<close> by this\<close> | |
| 195 | ||
| 196 | in | |
| 197 | ||
| 60054 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 198 | fun with_prems_cconv n cv ct = | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 199 | let | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 200 | fun strip_prems 0 As B = (As, B) | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 201 | | strip_prems i As B = | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 202 | case try Thm.dest_implies B of | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 203 | NONE => (As, B) | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 204 | | SOME (A,B) => strip_prems (i - 1) (A::As) B | 
| 74624 | 205 | val (prem, (prems, concl)) = ct |> Thm.dest_implies ||> strip_prems n [] | 
| 74607 | 206 | val th1 = cv prem RS rewr_imp concl | 
| 60054 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 207 | val nprems = Thm.nprems_of th1 | 
| 74624 | 208 | fun f p th = | 
| 209 |       Conv.fconv_rule (Conv.concl_conv nprems (Conv.rewr_conv @{thm imp_cong_eq}))
 | |
| 210 | (th RS cut_rl p) | |
| 60054 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 211 | in fold f prems th1 end | 
| 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 noschinl parents: 
60050diff
changeset | 212 | |
| 74607 | 213 | end | 
| 214 | ||
| 59739 | 215 | (*forward conversion, cf. FCONV_RULE in LCF*) | 
| 216 | fun fconv_rule cv th = | |
| 217 | let | |
| 218 | val eq = cv (Thm.cprop_of th) | |
| 219 | in | |
| 220 | if Thm.is_reflexive eq then th | |
| 221 | else th COMP (Thm.permute_prems 0 (Thm.nprems_of eq) (eq RS Drule.equal_elim_rule1)) | |
| 222 | end | |
| 223 | ||
| 224 | (*goal conversion*) | |
| 225 | fun gconv_rule cv i th = | |
| 226 | (case try (Thm.cprem_of th) i of | |
| 227 | SOME ct => | |
| 228 | let | |
| 229 | val eq = cv ct | |
| 230 | ||
| 231 | (* Drule.with_subgoal assumes that there are no new premises generated | |
| 232 | and thus rotates the premises wrongly. *) | |
| 233 | fun with_subgoal i f thm = | |
| 234 | let | |
| 235 | val num_prems = Thm.nprems_of thm | |
| 236 | val rotate_to_front = rotate_prems (i - 1) | |
| 237 | fun rotate_back thm = rotate_prems (1 - i + num_prems - Thm.nprems_of thm) thm | |
| 238 | in | |
| 239 | thm |> rotate_to_front |> f |> rotate_back | |
| 240 | end | |
| 241 | in | |
| 242 | if Thm.is_reflexive eq then th | |
| 243 | else with_subgoal i (fconv_rule (arg1_cconv (K eq))) th | |
| 244 | end | |
| 245 |   | NONE => raise THM ("gconv_rule", i, [th]))
 | |
| 246 | ||
| 59975 | 247 | (* Conditional conversions as tactics. *) | 
| 59739 | 248 | fun CCONVERSION cv i st = Seq.single (gconv_rule cv i st) | 
| 249 | handle THM _ => Seq.empty | |
| 250 | | CTERM _ => Seq.empty | |
| 251 | | TERM _ => Seq.empty | |
| 252 | | TYPE _ => Seq.empty | |
| 253 | ||
| 254 | end | |
| 255 | ||
| 256 | structure Basic_CConv: BASIC_CCONV = CConv | |
| 257 | open Basic_CConv |