| author | Fabian Huch <huch@in.tum.de> | 
| Thu, 23 Nov 2023 20:30:45 +0100 | |
| changeset 79027 | d08fb157e300 | 
| parent 76007 | 08288b406005 | 
| permissions | -rw-r--r-- | 
| 59739 | 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 | 3  | 
begin  | 
4  | 
||
5  | 
section \<open>The rewrite Proof Method by Example\<close>  | 
|
6  | 
||
| 74888 | 7  | 
text\<open>  | 
8  | 
This theory gives an overview over the features of the pattern-based rewrite proof method.  | 
|
| 59739 | 9  | 
|
| 76007 | 10  | 
Documentation: \<^url>\<open>https://arxiv.org/abs/2111.04082\<close>  | 
| 74888 | 11  | 
\<close>  | 
12  | 
||
| 59739 | 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"  | 
|
| 60047 | 25  | 
by (rewrite in "f _ + f \<hole> = _" diff_self) fact  | 
| 59739 | 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"  | 
|
| 60047 | 31  | 
by (rewrite at "f (_ + \<hole>) + f _ = _" diff_self) fact  | 
| 59739 | 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"  | 
|
| 60047 | 37  | 
by (rewrite in "f (\<hole> + _) + _ = _" diff_self) fact  | 
| 59739 | 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"  | 
|
| 60047 | 43  | 
by (rewrite in "f (_ + \<hole>) + _ = _" diff_self) fact  | 
| 59739 | 44  | 
|
45  | 
lemma  | 
|
46  | 
fixes x y :: nat  | 
|
47  | 
shows"x + y > c \<Longrightarrow> y + x > c"  | 
|
| 60047 | 48  | 
by (rewrite at "\<hole> > c" add.commute) assumption  | 
| 59739 | 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"  | 
|
| 60047 | 67  | 
by (rewrite at "\<hole> > c" at asm add.commute) fact  | 
| 59739 | 68  | 
|
69  | 
lemma  | 
|
70  | 
  assumes "P {x::int. y + 1 = 1 + x}"
 | 
|
71  | 
  shows   "P {x::int. y + 1 = x + 1}"
 | 
|
| 60047 | 72  | 
  by (rewrite at "x+1" in "{x::int. \<hole> }" add.commute) fact
 | 
| 59739 | 73  | 
|
74  | 
lemma  | 
|
75  | 
  assumes "P {x::int. y + 1 = 1 + x}"
 | 
|
76  | 
  shows   "P {x::int. y + 1 = x + 1}"
 | 
|
| 60047 | 77  | 
  by (rewrite at "any_identifier_will_work+1" in "{any_identifier_will_work::int. \<hole> }" add.commute)
 | 
| 59739 | 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)}"
 | 
|
| 60047 | 83  | 
by (rewrite at "b + d * e" in "\<lambda>(a, b, c). _ = Q (\<lambda>d e. \<hole>)" add.commute) fact  | 
| 59739 | 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 | 91  | 
lemma  | 
92  | 
assumes "PROP P \<equiv> PROP Q"  | 
|
93  | 
shows "PROP R \<Longrightarrow> PROP R \<Longrightarrow> PROP P \<Longrightarrow> PROP Q"  | 
|
94  | 
by (rewrite at asm assms)  | 
|
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 | 105  | 
|
106  | 
(* Rewriting with conditional rewriting rules works just as well. *)  | 
|
107  | 
lemma test_theorem:  | 
|
108  | 
fixes x :: nat  | 
|
109  | 
shows "x \<le> y \<Longrightarrow> x \<ge> y \<Longrightarrow> x = y"  | 
|
110  | 
by (rule Orderings.order_antisym)  | 
|
111  | 
||
| 60050 | 112  | 
(* Premises of the conditional rule yield new subgoals. The  | 
113  | 
assumptions of the goal are propagated into these subgoals  | 
|
114  | 
*)  | 
|
| 59739 | 115  | 
lemma  | 
| 60050 | 116  | 
fixes f :: "nat \<Rightarrow> nat"  | 
117  | 
shows "f x \<le> 0 \<Longrightarrow> f x \<ge> 0 \<Longrightarrow> f x = 0"  | 
|
| 59739 | 118  | 
apply (rewrite at "f x" to "0" test_theorem)  | 
| 60050 | 119  | 
apply assumption  | 
120  | 
apply assumption  | 
|
| 59739 | 121  | 
apply (rule refl)  | 
| 60050 | 122  | 
done  | 
| 59739 | 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 | 138  | 
(*  | 
139  | 
Instantiation.  | 
|
140  | 
||
141  | 
Since all rewriting is now done via conversions,  | 
|
142  | 
instantiation becomes fairly easy to do.  | 
|
143  | 
*)  | 
|
144  | 
||
145  | 
(* We first introduce a function f and an extended  | 
|
146  | 
version of f that is annotated with an invariant. *)  | 
|
147  | 
fun f :: "nat \<Rightarrow> nat" where "f n = n"  | 
|
148  | 
definition "f_inv (I :: nat \<Rightarrow> bool) n \<equiv> f n"  | 
|
149  | 
||
150  | 
lemma annotate_f: "f = f_inv I"  | 
|
151  | 
by (simp add: f_inv_def fun_eq_iff)  | 
|
152  | 
||
153  | 
(* We have a lemma with a bound variable n, and  | 
|
154  | 
want to add an invariant to f. *)  | 
|
155  | 
lemma  | 
|
156  | 
assumes "P (\<lambda>n. f_inv (\<lambda>_. True) n + 1) = x"  | 
|
157  | 
shows "P (\<lambda>n. f n + 1) = x"  | 
|
158  | 
by (rewrite to "f_inv (\<lambda>_. True)" annotate_f) fact  | 
|
159  | 
||
160  | 
(* We can also add an invariant that contains the variable n bound in the outer context.  | 
|
161  | 
For this, we need to bind this variable to an identifier. *)  | 
|
162  | 
lemma  | 
|
163  | 
assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"  | 
|
164  | 
shows "P (\<lambda>n. f n + 1) = x"  | 
|
| 60047 | 165  | 
by (rewrite in "\<lambda>n. \<hole>" to "f_inv (\<lambda>x. n < x + 1)" annotate_f) fact  | 
| 59739 | 166  | 
|
167  | 
(* Any identifier will work *)  | 
|
168  | 
lemma  | 
|
169  | 
assumes "P (\<lambda>n. f_inv (\<lambda>x. n < x + 1) n + 1) = x"  | 
|
170  | 
shows "P (\<lambda>n. f n + 1) = x"  | 
|
| 60047 | 171  | 
by (rewrite in "\<lambda>abc. \<hole>" to "f_inv (\<lambda>x. abc < x + 1)" annotate_f) fact  | 
| 59739 | 172  | 
|
173  | 
(* The "for" keyword. *)  | 
|
174  | 
lemma  | 
|
175  | 
assumes "P (2 + 1)"  | 
|
176  | 
shows "\<And>x y. P (1 + 2 :: nat)"  | 
|
177  | 
by (rewrite in "P (1 + 2)" at for (x) add.commute) fact  | 
|
178  | 
||
179  | 
lemma  | 
|
180  | 
assumes "\<And>x y. P (y + x)"  | 
|
181  | 
shows "\<And>x y. P (x + y :: nat)"  | 
|
182  | 
by (rewrite in "P (x + _)" at for (x y) add.commute) fact  | 
|
183  | 
||
184  | 
lemma  | 
|
185  | 
assumes "\<And>x y z. y + x + z = z + y + (x::int)"  | 
|
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 | 188  | 
|
189  | 
lemma  | 
|
190  | 
assumes "\<And>x y z. z + (x + y) = z + y + (x::int)"  | 
|
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 | 193  | 
|
194  | 
lemma  | 
|
195  | 
assumes "\<And>x y z. x + y + z = y + z + (x::int)"  | 
|
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 | 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 | 210  | 
lemma  | 
211  | 
assumes "(\<And>(x::int). x < 1 + x)"  | 
|
212  | 
and "(x::int) + 1 > x"  | 
|
213  | 
shows "(\<And>(x::int). x + 1 > x) \<Longrightarrow> (x::int) + 1 > x"  | 
|
214  | 
by (rewrite at "x + 1" in for (x) at asm add.commute)  | 
|
215  | 
(rule assms)  | 
|
216  | 
||
| 60079 | 217  | 
(* The rewrite method also has an ML interface *)  | 
218  | 
lemma  | 
|
219  | 
assumes "\<And>a b. P ((a + 1) * (1 + b)) "  | 
|
220  | 
shows "\<And>a b :: nat. P ((a + 1) * (b + 1))"  | 
|
221  | 
apply (tactic \<open>  | 
|
222  | 
let  | 
|
| 69597 | 223  | 
val (x, ctxt) = yield_singleton Variable.add_fixes "x" \<^context>  | 
| 60079 | 224  | 
(* Note that the pattern order is reversed *)  | 
225  | 
val pat = [  | 
|
| 74395 | 226  | 
Rewrite.For [(x, SOME \<^Type>\<open>nat\<close>)],  | 
| 60079 | 227  | 
Rewrite.In,  | 
| 74395 | 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 | 229  | 
val to = NONE  | 
| 60117 | 230  | 
    in CCONVERSION (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) 1 end
 | 
| 60079 | 231  | 
\<close>)  | 
232  | 
apply (fact assms)  | 
|
233  | 
done  | 
|
234  | 
||
235  | 
lemma  | 
|
236  | 
assumes "Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. a + b))"  | 
|
237  | 
shows "Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. b + a))"  | 
|
238  | 
apply (tactic \<open>  | 
|
239  | 
let  | 
|
| 69597 | 240  | 
val (x, ctxt) = yield_singleton Variable.add_fixes "x" \<^context>  | 
| 60079 | 241  | 
val pat = [  | 
242  | 
Rewrite.Concl,  | 
|
243  | 
Rewrite.In,  | 
|
| 74395 | 244  | 
        Rewrite.Term (Free ("Q", (\<^Type>\<open>int\<close> --> TVar (("'b",0), [])) --> \<^Type>\<open>bool\<close>)
 | 
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 | 246  | 
Rewrite.In,  | 
| 74395 | 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 | 248  | 
]  | 
249  | 
val to = NONE  | 
|
| 60117 | 250  | 
    in CCONVERSION (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) 1 end
 | 
| 60079 | 251  | 
\<close>)  | 
252  | 
apply (fact assms)  | 
|
253  | 
done  | 
|
254  | 
||
255  | 
(* There is also conversion-like rewrite function: *)  | 
|
256  | 
ML \<open>  | 
|
| 69597 | 257  | 
val ct = \<^cprop>\<open>Q (\<lambda>b :: int. P (\<lambda>a. a + b) (\<lambda>a. b + a))\<close>  | 
258  | 
val (x, ctxt) = yield_singleton Variable.add_fixes "x" \<^context>  | 
|
| 60079 | 259  | 
val pat = [  | 
260  | 
Rewrite.Concl,  | 
|
261  | 
Rewrite.In,  | 
|
| 69597 | 262  | 
    Rewrite.Term (Free ("Q", (\<^typ>\<open>int\<close> --> TVar (("'b",0), [])) --> \<^typ>\<open>bool\<close>)
 | 
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 | 264  | 
Rewrite.In,  | 
| 74395 | 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 | 266  | 
]  | 
267  | 
val to = NONE  | 
|
| 60117 | 268  | 
  val th = Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute} ct
 | 
| 60079 | 269  | 
\<close>  | 
270  | 
||
| 74889 | 271  | 
text \<open>Some regression tests\<close>  | 
| 
60088
 
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 | 274  | 
val ct = \<^cterm>\<open>(\<lambda>b :: int. (\<lambda>a. b + a))\<close>  | 
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 | 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 | 281  | 
val _ =  | 
282  | 
    case try (Rewrite.rewrite_conv ctxt (pat, to) @{thms add.commute}) ct of
 | 
|
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 | 287  | 
ML \<open>  | 
| 69597 | 288  | 
Rewrite.params_pconv (Conv.all_conv |> K |> K) \<^context> (Vartab.empty, []) \<^cterm>\<open>\<And>x. PROP A\<close>  | 
| 60117 | 289  | 
\<close>  | 
| 
60088
 
0a064330a885
rewrite: use distinct names for unnamed abstractions
 
noschinl 
parents: 
60079 
diff
changeset
 | 
290  | 
|
| 60117 | 291  | 
lemma  | 
292  | 
assumes eq: "PROP A \<Longrightarrow> PROP B \<equiv> PROP C"  | 
|
293  | 
assumes f1: "PROP D \<Longrightarrow> PROP A"  | 
|
294  | 
assumes f2: "PROP D \<Longrightarrow> PROP C"  | 
|
295  | 
shows "\<And>x. PROP D \<Longrightarrow> PROP B"  | 
|
296  | 
apply (rewrite eq)  | 
|
297  | 
apply (fact f1)  | 
|
298  | 
apply (fact f2)  | 
|
299  | 
done  | 
|
| 
60088
 
0a064330a885
rewrite: use distinct names for unnamed abstractions
 
noschinl 
parents: 
60079 
diff
changeset
 | 
300  | 
|
| 59739 | 301  | 
end  |