| author | Fabian Huch <huch@in.tum.de> | 
| Mon, 10 Jun 2024 18:45:12 +0200 | |
| changeset 80341 | b061568ae52d | 
| parent 74888 | 1c50ddcf6a01 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 59975 | 1  | 
(* Title: HOL/Library/Rewrite.thy  | 
2  | 
Author: Christoph Traut, Lars Noschinski, TU Muenchen  | 
|
3  | 
||
4  | 
Proof method "rewrite" with support for subterm-selection based on patterns.  | 
|
| 74888 | 5  | 
|
6  | 
Documentation: https://arxiv.org/abs/2111.04082  | 
|
| 59975 | 7  | 
*)  | 
8  | 
||
| 59739 | 9  | 
theory Rewrite  | 
10  | 
imports Main  | 
|
11  | 
begin  | 
|
12  | 
||
| 61383 | 13  | 
consts rewrite_HOLE :: "'a::{}"  ("\<hole>")
 | 
| 59739 | 14  | 
|
15  | 
lemma eta_expand:  | 
|
| 59975 | 16  | 
  fixes f :: "'a::{} \<Rightarrow> 'b::{}"
 | 
17  | 
shows "f \<equiv> \<lambda>x. f x" .  | 
|
| 59739 | 18  | 
|
| 
60054
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
19  | 
lemma imp_cong_eq:  | 
| 61383 | 20  | 
"(PROP A \<Longrightarrow> (PROP B \<Longrightarrow> PROP C) \<equiv> (PROP B' \<Longrightarrow> PROP C')) \<equiv>  | 
21  | 
((PROP B \<Longrightarrow> PROP A \<Longrightarrow> PROP C) \<equiv> (PROP B' \<Longrightarrow> PROP A \<Longrightarrow> PROP C'))"  | 
|
| 
60054
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
22  | 
apply (intro Pure.equal_intr_rule)  | 
| 
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
23  | 
apply (drule (1) cut_rl; drule Pure.equal_elim_rule1 Pure.equal_elim_rule2; assumption)+  | 
| 
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
24  | 
apply (drule Pure.equal_elim_rule1 Pure.equal_elim_rule2; assumption)+  | 
| 
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
25  | 
done  | 
| 
 
ef4878146485
rewrite: with asm pattern, propagate also remaining assumptions to new subgoals
 
noschinl 
parents: 
60047 
diff
changeset
 | 
26  | 
|
| 69605 | 27  | 
ML_file \<open>cconv.ML\<close>  | 
28  | 
ML_file \<open>rewrite.ML\<close>  | 
|
| 59739 | 29  | 
|
30  | 
end  |