16417
|
1 |
theory Isar imports Sets begin
|
11235
|
2 |
|
|
3 |
section{*A Taste of Structured Proofs in Isar*}
|
|
4 |
|
|
5 |
lemma [intro?]: "mono f \<Longrightarrow> x \<in> f (lfp f) \<Longrightarrow> x \<in> lfp f"
|
|
6 |
by(simp only: lfp_unfold [symmetric])
|
|
7 |
|
|
8 |
lemma "lfp(\<lambda>T. A \<union> M\<inverse> `` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
|
|
9 |
(is "lfp ?F = ?toA")
|
|
10 |
proof
|
|
11 |
show "lfp ?F \<subseteq> ?toA"
|
12815
|
12 |
by (blast intro!: lfp_lowerbound intro: rtrancl_trans)
|
13249
|
13 |
next
|
11235
|
14 |
show "?toA \<subseteq> lfp ?F"
|
|
15 |
proof
|
|
16 |
fix s assume "s \<in> ?toA"
|
11238
|
17 |
then obtain t where stM: "(s,t) \<in> M\<^sup>*" and tA: "t \<in> A"
|
|
18 |
by blast
|
11235
|
19 |
from stM show "s \<in> lfp ?F"
|
|
20 |
proof (rule converse_rtrancl_induct)
|
|
21 |
from tA have "t \<in> ?F (lfp ?F)" by blast
|
|
22 |
with mono_ef show "t \<in> lfp ?F" ..
|
13249
|
23 |
next
|
11238
|
24 |
fix s s' assume "(s,s') \<in> M" and "(s',t) \<in> M\<^sup>*"
|
|
25 |
and "s' \<in> lfp ?F"
|
11235
|
26 |
then have "s \<in> ?F (lfp ?F)" by blast
|
|
27 |
with mono_ef show "s \<in> lfp ?F" ..
|
|
28 |
qed
|
|
29 |
qed
|
|
30 |
qed
|
|
31 |
|
|
32 |
end
|
|
33 |
|