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