|
59739
|
1 |
theory Rewrite_Examples
|
|
|
2 |
imports Main "~~/src/HOL/Library/Rewrite"
|
|
|
3 |
begin
|
|
|
4 |
|
|
|
5 |
section \<open>The rewrite Proof Method by Example\<close>
|
|
|
6 |
|
|
|
7 |
(* This file is intended to give an overview over
|
|
|
8 |
the features of the pattern-based rewrite proof method.
|
|
|
9 |
|
|
|
10 |
See also https://www21.in.tum.de/~noschinl/Pattern-2014/
|
|
|
11 |
*)
|
|
|
12 |
|
|
|
13 |
lemma
|
|
|
14 |
fixes a::int and b::int and c::int
|
|
|
15 |
assumes "P (b + a)"
|
|
|
16 |
shows "P (a + b)"
|
|
|
17 |
by (rewrite at "a + b" add.commute)
|
|
|
18 |
(rule assms)
|
|
|
19 |
|
|
|
20 |
(* Selecting a specific subterm in a large, ambiguous term. *)
|
|
|
21 |
lemma
|
|
|
22 |
fixes a b c :: int
|
|
|
23 |
assumes "f (a - a + (a - a)) + f ( 0 + c) = f 0 + f c"
|
|
|
24 |
shows "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
|
|
|
25 |
by (rewrite in "f _ + f \<box> = _" diff_self) fact
|
|
|
26 |
|
|
|
27 |
lemma
|
|
|
28 |
fixes a b c :: int
|
|
|
29 |
assumes "f (a - a + 0 ) + f ((a - a) + c) = f 0 + f c"
|
|
|
30 |
shows "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
|
|
|
31 |
by (rewrite at "f (_ + \<box>) + f _ = _" diff_self) fact
|
|
|
32 |
|
|
|
33 |
lemma
|
|
|
34 |
fixes a b c :: int
|
|
|
35 |
assumes "f ( 0 + (a - a)) + f ((a - a) + c) = f 0 + f c"
|
|
|
36 |
shows "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
|
|
|
37 |
by (rewrite in "f (\<box> + _) + _ = _" diff_self) fact
|
|
|
38 |
|
|
|
39 |
lemma
|
|
|
40 |
fixes a b c :: int
|
|
|
41 |
assumes "f (a - a + 0 ) + f ((a - a) + c) = f 0 + f c"
|
|
|
42 |
shows "f (a - a + (a - a)) + f ((a - a) + c) = f 0 + f c"
|
|
|
43 |
by (rewrite in "f (_ + \<box>) + _ = _" diff_self) fact
|
|
|
44 |
|
|
|
45 |
lemma
|
|
|
46 |
fixes x y :: nat
|
|
|
47 |
shows"x + y > c \<Longrightarrow> y + x > c"
|
|
|
48 |
by (rewrite at "\<box> > c" add.commute) assumption
|
|
|
49 |
|
|
|
50 |
(* We can also rewrite in the assumptions. *)
|
|
|
51 |
lemma
|
|
|
52 |
fixes x y :: nat
|
|
|
53 |
assumes "y + x > c \<Longrightarrow> y + x > c"
|
|
|
54 |
shows "x + y > c \<Longrightarrow> y + x > c"
|
|
|
55 |
by (rewrite in asm add.commute) fact
|
|
|
56 |
|
|
|
57 |
lemma
|
|
|
58 |
fixes x y :: nat
|
|
|
59 |
assumes "y + x > c \<Longrightarrow> y + x > c"
|
|
|
60 |
shows "x + y > c \<Longrightarrow> y + x > c"
|
|
|
61 |
by (rewrite in "x + y > c" at asm add.commute) fact
|
|
|
62 |
|
|
|
63 |
lemma
|
|
|
64 |
fixes x y :: nat
|
|
|
65 |
assumes "y + x > c \<Longrightarrow> y + x > c"
|
|
|
66 |
shows "x + y > c \<Longrightarrow> y + x > c"
|
|
|
67 |
by (rewrite at "\<box> > c" at asm add.commute) fact
|
|
|
68 |
|
|
|
69 |
lemma
|
|
|
70 |
assumes "P {x::int. y + 1 = 1 + x}"
|
|
|
71 |
shows "P {x::int. y + 1 = x + 1}"
|
|
|
72 |
by (rewrite at "x+1" in "{x::int. \<box> }" add.commute) fact
|
|
|
73 |
|
|
|
74 |
lemma
|
|
|
75 |
assumes "P {x::int. y + 1 = 1 + x}"
|
|
|
76 |
shows "P {x::int. y + 1 = x + 1}"
|
|
|
77 |
by (rewrite at "any_identifier_will_work+1" in "{any_identifier_will_work::int. \<box> }" add.commute)
|
|
|
78 |
fact
|
|
|
79 |
|
|
|
80 |
lemma
|
|
|
81 |
assumes "P {(x::nat, y::nat, z). x + z * 3 = Q (\<lambda>s t. s * t + y - 3)}"
|
|
|
82 |
shows "P {(x::nat, y::nat, z). x + z * 3 = Q (\<lambda>s t. y + s * t - 3)}"
|
|
|
83 |
by (rewrite at "b + d * e" in "\<lambda>(a, b, c). _ = Q (\<lambda>d e. \<box>)" add.commute) fact
|
|
|
84 |
|
|
|
85 |
|
|
|
86 |
(* Rewriting with conditional rewriting rules works just as well. *)
|
|
|
87 |
lemma test_theorem:
|
|
|
88 |
fixes x :: nat
|
|
|
89 |
shows "x \<le> y \<Longrightarrow> x \<ge> y \<Longrightarrow> x = y"
|
|
|
90 |
by (rule Orderings.order_antisym)
|
|
|
91 |
|
|
|
92 |
lemma
|
|
|
93 |
fixes f :: "nat \<Rightarrow> nat"
|
|
|
94 |
assumes "f x \<le> 0" "f x \<ge> 0"
|
|
|
95 |
shows "f x = 0"
|
|
|
96 |
apply (rewrite at "f x" to "0" test_theorem)
|
|
|
97 |
apply fact
|
|
|
98 |
apply fact
|
|
|
99 |
apply (rule refl)
|
|
|
100 |
done
|
|
|
101 |
|
|
|
102 |
(*
|
|
|
103 |
Instantiation.
|
|
|
104 |
|
|
|
105 |
Since all rewriting is now done via conversions,
|
|
|
106 |
instantiation becomes fairly easy to do.
|
|
|
107 |
*)
|
|
|
108 |
|
|
|
109 |
(* We first introduce a function f and an extended
|
|
|
110 |
version of f that is annotated with an invariant. *)
|
|
|
111 |
fun f :: "nat \<Rightarrow> nat" where "f n = n"
|
|
|
112 |
definition "f_inv (I :: nat \<Rightarrow> bool) n \<equiv> f n"
|
|
|
113 |
|
|
|
114 |
lemma annotate_f: "f = f_inv I"
|
|
|
115 |
by (simp add: f_inv_def fun_eq_iff)
|
|
|
116 |
|
|
|
117 |
(* We have a lemma with a bound variable n, and
|
|
|
118 |
want to add an invariant to f. *)
|
|
|
119 |
lemma
|
|
|
120 |
assumes "P (\<lambda>n. f_inv (\<lambda>_. True) n + 1) = x"
|
|
|
121 |
shows "P (\<lambda>n. f n + 1) = x"
|
|
|
122 |
by (rewrite to "f_inv (\<lambda>_. True)" annotate_f) fact
|
|
|
123 |
|
|
|
124 |
(* We can also add an invariant that contains the variable n bound in the outer context.
|
|
|
125 |
For this, we need to bind this variable to an identifier. *)
|
|
|
126 |
lemma
|
|
|
127 |
assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"
|
|
|
128 |
shows "P (\<lambda>n. f n + 1) = x"
|
|
|
129 |
by (rewrite in "\<lambda>n. \<box>" to "f_inv (\<lambda>x. n < x + 1)" annotate_f) fact
|
|
|
130 |
|
|
|
131 |
(* Any identifier will work *)
|
|
|
132 |
lemma
|
|
|
133 |
assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"
|
|
|
134 |
shows "P (\<lambda>n. f n + 1) = x"
|
|
|
135 |
by (rewrite in "\<lambda>abc. \<box>" to "f_inv (\<lambda>x. abc < x + 1)" annotate_f) fact
|
|
|
136 |
|
|
|
137 |
(* The "for" keyword. *)
|
|
|
138 |
lemma
|
|
|
139 |
assumes "P (2 + 1)"
|
|
|
140 |
shows "\<And>x y. P (1 + 2 :: nat)"
|
|
|
141 |
by (rewrite in "P (1 + 2)" at for (x) add.commute) fact
|
|
|
142 |
|
|
|
143 |
lemma
|
|
|
144 |
assumes "\<And>x y. P (y + x)"
|
|
|
145 |
shows "\<And>x y. P (x + y :: nat)"
|
|
|
146 |
by (rewrite in "P (x + _)" at for (x y) add.commute) fact
|
|
|
147 |
|
|
|
148 |
lemma
|
|
|
149 |
assumes "\<And>x y z. y + x + z = z + y + (x::int)"
|
|
|
150 |
shows "\<And>x y z. x + y + z = z + y + (x::int)"
|
|
|
151 |
by (rewrite at "x + y" in "x + y + z" in concl at for (x y z) add.commute) fact
|
|
|
152 |
|
|
|
153 |
lemma
|
|
|
154 |
assumes "\<And>x y z. z + (x + y) = z + y + (x::int)"
|
|
|
155 |
shows "\<And>x y z. x + y + z = z + y + (x::int)"
|
|
|
156 |
by (rewrite at "(_ + y) + z" in concl at for (y z) add.commute) fact
|
|
|
157 |
|
|
|
158 |
lemma
|
|
|
159 |
assumes "\<And>x y z. x + y + z = y + z + (x::int)"
|
|
|
160 |
shows "\<And>x y z. x + y + z = z + y + (x::int)"
|
|
|
161 |
by (rewrite at "\<box> + _" at "_ = \<box>" in concl at for () add.commute) fact
|
|
|
162 |
|
|
|
163 |
(* The all-keyword can be used anywhere in the pattern where there is an \<And>-Quantifier. *)
|
|
|
164 |
lemma
|
|
|
165 |
assumes "(\<And>(x::int). x < 1 + x)"
|
|
|
166 |
and "(x::int) + 1 > x"
|
|
|
167 |
shows "(\<And>(x::int). x + 1 > x) \<Longrightarrow> (x::int) + 1 > x"
|
|
|
168 |
by (rewrite at "x + 1" in for (x) at asm add.commute)
|
|
|
169 |
(rule assms)
|
|
|
170 |
|
|
|
171 |
end
|
|
|
172 |
|