src/HOL/ex/Rewrite_Examples.thy
author wenzelm
Sun, 12 Aug 2018 14:31:46 +0200
changeset 68744 64fb127e33f7
parent 66453 cc19f7ca2ed6
child 69597 ff784d5a5bfb
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     1
theory Rewrite_Examples
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 60117
diff changeset
     2
imports Main "HOL-Library.Rewrite"
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     3
begin
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     4
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     5
section \<open>The rewrite Proof Method by Example\<close>
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     6
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     7
(* This file is intended to give an overview over
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     8
   the features of the pattern-based rewrite proof method.
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
     9
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    10
   See also https://www21.in.tum.de/~noschinl/Pattern-2014/
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    11
*)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    12
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    13
  fixes a::int and b::int and c::int
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    14
  assumes "P (b + a)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    15
  shows "P (a + b)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    16
by (rewrite at "a + b" add.commute)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    17
   (rule assms)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    18
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    19
(* Selecting a specific subterm in a large, ambiguous term. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    20
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    21
  fixes a b c :: int
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    22
  assumes "f (a - a + (a - a)) + f (   0    + c) = f 0 + f c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    23
  shows   "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    24
  by (rewrite in "f _ + f \<hole> = _" diff_self) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    25
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    26
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    27
  fixes a b c :: int
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    28
  assumes "f (a - a +    0   ) + f ((a - a) + c) = f 0 + f c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    29
  shows   "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    30
  by (rewrite at "f (_ + \<hole>) + f _ = _" diff_self) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    31
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    32
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    33
  fixes a b c :: int
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    34
  assumes "f (  0   + (a - a)) + f ((a - a) + c) = f 0 + f c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    35
  shows   "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    36
  by (rewrite in "f (\<hole> + _) + _ = _" diff_self) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    37
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    38
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    39
  fixes a b c :: int
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    40
  assumes "f (a - a +    0   ) + f ((a - a) + c) = f 0 + f c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    41
  shows   "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    42
  by (rewrite in "f (_ + \<hole>) + _ = _" diff_self) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    43
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    44
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    45
  fixes x y :: nat
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    46
  shows"x + y > c \<Longrightarrow> y + x > c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    47
  by (rewrite at "\<hole> > c" add.commute) assumption
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    48
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    49
(* We can also rewrite in the assumptions.  *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    50
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    51
  fixes x y :: nat
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    52
  assumes "y + x > c \<Longrightarrow> y + x > c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    53
  shows   "x + y > c \<Longrightarrow> y + x > c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    54
  by (rewrite in asm add.commute) fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    55
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    56
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    57
  fixes x y :: nat
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    58
  assumes "y + x > c \<Longrightarrow> y + x > c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    59
  shows   "x + y > c \<Longrightarrow> y + x > c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    60
  by (rewrite in "x + y > c" at asm add.commute) fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    61
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    62
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    63
  fixes x y :: nat
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    64
  assumes "y + x > c \<Longrightarrow> y + x > c"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    65
  shows   "x + y > c \<Longrightarrow> y + x > c"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    66
  by (rewrite at "\<hole> > c" at asm  add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    67
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    68
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    69
  assumes "P {x::int. y + 1 = 1 + x}"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    70
  shows   "P {x::int. y + 1 = x + 1}"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    71
  by (rewrite at "x+1" in "{x::int. \<hole> }" add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    72
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    73
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    74
  assumes "P {x::int. y + 1 = 1 + x}"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    75
  shows   "P {x::int. y + 1 = x + 1}"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    76
  by (rewrite at "any_identifier_will_work+1" in "{any_identifier_will_work::int. \<hole> }" add.commute)
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    77
   fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    78
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    79
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    80
  assumes "P {(x::nat, y::nat, z). x + z * 3 = Q (\<lambda>s t. s * t + y - 3)}"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    81
  shows   "P {(x::nat, y::nat, z). x + z * 3 = Q (\<lambda>s t. y + s * t - 3)}"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
    82
  by (rewrite at "b + d * e" in "\<lambda>(a, b, c). _ = Q (\<lambda>d e. \<hole>)" add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
    83
60052
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    84
(* This is not limited to the first assumption *)
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    85
lemma
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    86
  assumes "PROP P \<equiv> PROP Q"
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    87
  shows "PROP R \<Longrightarrow> PROP P \<Longrightarrow> PROP Q"
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    88
    by (rewrite at asm assms)
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
    89
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
    90
lemma
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
    91
  assumes "PROP P \<equiv> PROP Q"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
    92
  shows "PROP R \<Longrightarrow> PROP R \<Longrightarrow> PROP P \<Longrightarrow> PROP Q"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
    93
    by (rewrite at asm assms)
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
    94
60053
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
    95
(* Rewriting "at asm" selects each full assumption, not any parts *)
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
    96
lemma
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
    97
  assumes "(PROP P \<Longrightarrow> PROP Q) \<equiv> (PROP S \<Longrightarrow> PROP R)"
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
    98
  shows "PROP S \<Longrightarrow> (PROP P \<Longrightarrow> PROP Q) \<Longrightarrow> PROP R"
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
    99
  apply (rewrite at asm assms)
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
   100
  apply assumption
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
   101
  done
0e9895ffab1d rewrite: do not descend into conclusion of premise with asm pattern
noschinl
parents: 60052
diff changeset
   102
60052
616a17640229 rewrite: with asm pattern, try all premises for rewriting, not only the first
noschinl
parents: 60050
diff changeset
   103
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   104
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   105
(* Rewriting with conditional rewriting rules works just as well. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   106
lemma test_theorem:
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   107
  fixes x :: nat
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   108
  shows "x \<le> y \<Longrightarrow> x \<ge> y \<Longrightarrow> x = y"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   109
  by (rule Orderings.order_antisym)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   110
60050
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   111
(* Premises of the conditional rule yield new subgoals. The
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   112
   assumptions of the goal are propagated into these subgoals
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   113
*)
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   114
lemma
60050
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   115
  fixes f :: "nat \<Rightarrow> nat"
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   116
  shows "f x \<le> 0 \<Longrightarrow> f x \<ge> 0 \<Longrightarrow> f x = 0"
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   117
  apply (rewrite at "f x" to "0" test_theorem)
60050
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   118
  apply assumption
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   119
  apply assumption
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   120
  apply (rule refl)
60050
dc6ac152d864 rewrite: propagate premises to new subgoals
noschinl
parents: 60047
diff changeset
   121
  done
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   122
60054
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   123
(* This holds also for rewriting in assumptions. The order of assumptions is preserved *)
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   124
lemma
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   125
  assumes rewr: "PROP P \<Longrightarrow> PROP Q \<Longrightarrow> PROP R \<equiv> PROP R'"
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   126
  assumes A1: "PROP S \<Longrightarrow> PROP T \<Longrightarrow> PROP U \<Longrightarrow> PROP P"
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   127
  assumes A2: "PROP S \<Longrightarrow> PROP T \<Longrightarrow> PROP U \<Longrightarrow> PROP Q"
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   128
  assumes C: "PROP S \<Longrightarrow> PROP R' \<Longrightarrow> PROP T \<Longrightarrow> PROP U \<Longrightarrow> PROP V"
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   129
  shows "PROP S \<Longrightarrow> PROP R \<Longrightarrow> PROP T \<Longrightarrow> PROP U \<Longrightarrow> PROP V"
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   130
  apply (rewrite at asm rewr)
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   131
  apply (fact A1)
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   132
  apply (fact A2)
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   133
  apply (fact C)
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   134
  done
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   135
ef4878146485 rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
noschinl
parents: 60053
diff changeset
   136
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   137
(*
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   138
   Instantiation.
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   139
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   140
   Since all rewriting is now done via conversions,
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   141
   instantiation becomes fairly easy to do.
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   142
*)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   143
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   144
(* We first introduce a function f and an extended
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   145
   version of f that is annotated with an invariant. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   146
fun f :: "nat \<Rightarrow> nat" where "f n = n"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   147
definition "f_inv (I :: nat \<Rightarrow> bool) n \<equiv> f n"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   148
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   149
lemma annotate_f: "f = f_inv I"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   150
  by (simp add: f_inv_def fun_eq_iff)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   151
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   152
(* We have a lemma with a bound variable n, and
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   153
   want to add an invariant to f. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   154
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   155
  assumes "P (\<lambda>n. f_inv (\<lambda>_. True) n + 1) = x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   156
  shows "P (\<lambda>n. f n + 1) = x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   157
  by (rewrite to "f_inv (\<lambda>_. True)" annotate_f) fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   158
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   159
(* We can also add an invariant that contains the variable n bound in the outer context.
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   160
   For this, we need to bind this variable to an identifier. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   161
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   162
  assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   163
  shows "P (\<lambda>n. f n + 1) = x"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
   164
  by (rewrite in "\<lambda>n. \<hole>" to "f_inv (\<lambda>x. n < x + 1)" annotate_f) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   165
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   166
(* Any identifier will work *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   167
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   168
  assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   169
  shows "P (\<lambda>n. f n + 1) = x"
60047
58e5b16cbd94 enable \<hole> syntax for rewrite
noschinl
parents: 59739
diff changeset
   170
  by (rewrite in "\<lambda>abc. \<hole>" to "f_inv (\<lambda>x. abc < x + 1)" annotate_f) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   171
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   172
(* The "for" keyword. *)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   173
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   174
  assumes "P (2 + 1)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   175
  shows "\<And>x y. P (1 + 2 :: nat)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   176
by (rewrite in "P (1 + 2)" at for (x) add.commute) fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   177
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   178
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   179
  assumes "\<And>x y. P (y + x)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   180
  shows "\<And>x y. P (x + y :: nat)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   181
by (rewrite in "P (x + _)" at for (x y) add.commute) fact
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   182
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   183
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   184
  assumes "\<And>x y z. y + x + z = z + y + (x::int)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   185
  shows   "\<And>x y z. x + y + z = z + y + (x::int)"
60108
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   186
by (rewrite at "x + y" in "x + y + z" in for (x y z) add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   187
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   188
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   189
  assumes "\<And>x y z. z + (x + y) = z + y + (x::int)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   190
  shows   "\<And>x y z. x + y + z = z + y + (x::int)"
60108
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   191
by (rewrite at "(_ + y) + z" in for (y z) add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   192
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   193
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   194
  assumes "\<And>x y z. x + y + z = y + z + (x::int)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   195
  shows   "\<And>x y z. x + y + z = z + y + (x::int)"
60108
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   196
by (rewrite at "\<hole> + _" at "_ = \<hole>" in for () add.commute) fact
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   197
60108
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   198
lemma
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   199
  assumes eq: "\<And>x. P x \<Longrightarrow> g x = x"
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   200
  assumes f1: "\<And>x. Q x \<Longrightarrow> P x"
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   201
  assumes f2: "\<And>x. Q x \<Longrightarrow> x"
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   202
  shows "\<And>x. Q x \<Longrightarrow> g x"
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   203
  apply (rewrite at "g x" in for (x) eq)
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   204
  apply (fact f1)
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   205
  apply (fact f2)
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   206
  done
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   207
d7fe3e0aca85 rewrite: add default pattern "in concl" for more cases
noschinl
parents: 60088
diff changeset
   208
(* The for keyword can be used anywhere in the pattern where there is an \<And>-Quantifier. *)
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   209
lemma
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   210
  assumes "(\<And>(x::int). x < 1 + x)"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   211
  and     "(x::int) + 1 > x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   212
  shows   "(\<And>(x::int). x + 1 > x) \<Longrightarrow> (x::int) + 1 > x"
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   213
by (rewrite at "x + 1" in for (x) at asm add.commute)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   214
   (rule assms)
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   215
60079
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   216
(* The rewrite method also has an ML interface *)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   217
lemma
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   218
  assumes "\<And>a b. P ((a + 1) * (1 + b)) "
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   219
  shows "\<And>a b :: nat. P ((a + 1) * (b + 1))"
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   220
  apply (tactic \<open>
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   221
    let
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   222
      val (x, ctxt) = yield_singleton Variable.add_fixes "x" @{context}
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   223
      (* Note that the pattern order is reversed *)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   224
      val pat = [
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   225
        Rewrite.For [(x, SOME @{typ nat})],
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   226
        Rewrite.In,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   227
        Rewrite.Term (@{const plus(nat)} $ Free (x, @{typ nat}) $ @{term "1 :: nat"}, [])]
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   228
      val to = NONE
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   229
    in CCONVERSION (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) 1 end
60079
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   230
  \<close>)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   231
  apply (fact assms)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   232
  done
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   233
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   234
lemma
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   235
  assumes "Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. a + b))"
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   236
  shows "Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. b + a))"
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   237
  apply (tactic \<open>
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   238
    let
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   239
      val (x, ctxt) = yield_singleton Variable.add_fixes "x" @{context}
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   240
      val pat = [
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   241
        Rewrite.Concl,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   242
        Rewrite.In,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   243
        Rewrite.Term (Free ("Q", (@{typ "int"} --> TVar (("'b",0), [])) --> @{typ bool})
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   244
          $ Abs ("x", @{typ int}, Rewrite.mk_hole 1 (@{typ int} --> TVar (("'b",0), [])) $ Bound 0), [(x, @{typ int})]),
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   245
        Rewrite.In,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   246
        Rewrite.Term (@{const plus(int)} $ Free (x, @{typ int}) $ Var (("c", 0), @{typ int}), [])
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   247
        ]
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   248
      val to = NONE
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   249
    in CCONVERSION (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) 1 end
60079
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   250
  \<close>)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   251
  apply (fact assms)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   252
  done
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   253
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   254
(* There is also conversion-like rewrite function: *)
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   255
ML \<open>
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   256
  val ct = @{cprop "Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. b + a))"}
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   257
  val (x, ctxt) = yield_singleton Variable.add_fixes "x" @{context}
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   258
  val pat = [
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   259
    Rewrite.Concl,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   260
    Rewrite.In,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   261
    Rewrite.Term (Free ("Q", (@{typ "int"} --> TVar (("'b",0), [])) --> @{typ bool})
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   262
      $ Abs ("x", @{typ int}, Rewrite.mk_hole 1 (@{typ int} --> TVar (("'b",0), [])) $ Bound 0), [(x, @{typ int})]),
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   263
    Rewrite.In,
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   264
    Rewrite.Term (@{const plus(int)} $ Free (x, @{typ int}) $ Var (("c", 0), @{typ int}), [])
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   265
    ]
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   266
  val to = NONE
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   267
  val th = Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute} ct
60079
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   268
\<close>
ef4fe30e9ef1 rewrite: add ML interface
noschinl
parents: 60054
diff changeset
   269
60088
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   270
section \<open>Regression tests\<close>
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   271
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   272
ML \<open>
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   273
  val ct = @{cterm "(\<lambda>b :: int. (\<lambda>a. b + a))"}
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   274
  val (x, ctxt) = yield_singleton Variable.add_fixes "x" @{context}
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   275
  val pat = [
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   276
    Rewrite.In,
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   277
    Rewrite.Term (@{const plus(int)} $ Var (("c", 0), @{typ int}) $ Var (("c", 0), @{typ int}), [])
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   278
    ]
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   279
  val to = NONE
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   280
  val _ =
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   281
    case try (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) ct of
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   282
      NONE => ()
60088
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   283
    | _ => error "should not have matched anything"
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   284
\<close>
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   285
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   286
ML \<open>
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   287
  Rewrite.params_pconv (Conv.all_conv |> K |> K) @{context} (Vartab.empty, []) @{cterm "\<And>x. PROP A"}
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   288
\<close>
60088
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   289
60117
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   290
lemma
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   291
  assumes eq: "PROP A \<Longrightarrow> PROP B \<equiv> PROP C"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   292
  assumes f1: "PROP D \<Longrightarrow> PROP A"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   293
  assumes f2: "PROP D \<Longrightarrow> PROP C"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   294
  shows "\<And>x. PROP D \<Longrightarrow> PROP B"
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   295
  apply (rewrite eq)
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   296
  apply (fact f1)
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   297
  apply (fact f2)
2712f40d6309 rewrite: work purely conversion-based
noschinl
parents: 60108
diff changeset
   298
  done
60088
0a064330a885 rewrite: use distinct names for unnamed abstractions
noschinl
parents: 60079
diff changeset
   299
59739
4ed50ebf5d36 added proof method rewrite
noschinl
parents:
diff changeset
   300
end