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