(*****************************************************************
Isabelle Tutorial
-----------------
2st June 2009, Beijing
*)
theory Lec3
imports "Main"
begin
definition
lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _")
where
"L1 ; L2 = {s1@s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
fun
lang_pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _")
where
"L \<up> 0 = {[]}"
| "L \<up> (Suc i) = L ; (L \<up> i)"
definition
lang_star :: "string set \<Rightarrow> string set" ("_\<star>")
where
"L\<star> \<equiv> \<Union>i. (L \<up> i)"
lemma lang_seq_cases:
shows "(s \<in> L1 ; L2) = (\<exists>s1 s2. s = s1@s2 \<and> s1\<in>L1 \<and> s2\<in>L2)"
by (simp add: lang_seq_def)
lemma lang_seq_union:
shows "(L1 \<union> L2);L3 = (L1;L3) \<union> (L2;L3)"
and "L1;(L2 \<union> L3) = (L1;L2) \<union> (L1;L3)"
unfolding lang_seq_def by auto
lemma lang_seq_empty:
shows "{[]} ; L = L"
and "L ; {[]} = L"
unfolding lang_seq_def by auto
lemma lang_seq_assoc:
shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)"
by (simp add: lang_seq_def Collect_def mem_def expand_fun_eq)
(metis append_assoc)
lemma silly:
shows "[] \<in> L \<up> 0"
by simp
lemma lang_star_empty:
shows "{[]} \<union> (L\<star>) = L\<star>"
unfolding lang_star_def
by (auto intro: silly)
lemma lang_star_in_empty:
shows "[] \<in> L\<star>"
unfolding lang_star_def
by (auto intro: silly)
lemma lang_seq_subseteq:
shows "L \<subseteq> (L'\<star>) ; L"
and "L \<subseteq> L ; (L'\<star>)"
proof -
have "L = {[]} ; L" using lang_seq_empty by simp
also have "\<dots> \<subseteq> ({[]} ; L) \<union> ((L'\<star>) ; L)" by auto
also have "\<dots> = ({[]} \<union> (L'\<star>)) ; L" by (simp add: lang_seq_union[symmetric])
also have "\<dots> = (L'\<star>); L" using lang_star_empty by simp
finally show "L \<subseteq> (L'\<star>); L" by simp
next
show "L \<subseteq> L ; (L'\<star>)" sorry
qed
lemma lang_star_subseteq:
shows "L ; (L\<star>) \<subseteq> (L\<star>)"
unfolding lang_star_def lang_seq_def
apply(auto)
apply(rule_tac x="Suc xa" in exI)
apply(auto simp add: lang_seq_def)
done
(* regular expressions *)
datatype rexp =
EMPTY
| CHAR char
| SEQ rexp rexp
| ALT rexp rexp
| STAR rexp
fun
L :: "rexp \<Rightarrow> string set"
where
"L(EMPTY) = {[]}"
| "L(CHAR c) = {[c]}"
| "L(SEQ r1 r2) = (L r1) ; (L r2)"
| "L(ALT r1 r2) = (L r1) \<union> (L r2)"
| "L(STAR r) = (L r)\<star>"
definition
Ls :: "rexp set \<Rightarrow> string set"
where
"Ls R = (\<Union>r\<in>R. (L r))"
lemma
shows "Ls {} = {}"
unfolding Ls_def by simp
lemma Ls_union:
"Ls (R1 \<union> R2) = (Ls R1) \<union> (Ls R2)"
unfolding Ls_def by auto
function
dagger :: "rexp \<Rightarrow> char \<Rightarrow> rexp set" ("_ \<dagger> _")
where
r1: "(EMPTY) \<dagger> c = {}"
| r2: "(CHAR c') \<dagger> c = (if c = c' then {EMPTY} else {})"
| r3: "(ALT r1 r2) \<dagger> c = r1 \<dagger> c \<union> r2 \<dagger> c"
| r4: "(SEQ EMPTY r2) \<dagger> c = r2 \<dagger> c"
| r5: "(SEQ (CHAR c') r2) \<dagger> c = (if c= c' then {r2} else {})"
| r6: "(SEQ (SEQ r11 r12) r2) \<dagger> c = (SEQ r11 (SEQ r12 r2)) \<dagger> c"
| r7: "(SEQ (ALT r11 r12) r2) \<dagger> c = (SEQ r11 r2) \<dagger> c \<union> (SEQ r12 r2) \<dagger> c"
| r8: "(SEQ (STAR r1) r2) \<dagger> c =
r2 \<dagger> c \<union> {SEQ (SEQ r' (STAR r1)) r2 | r'. r' \<in> r1 \<dagger> c}"
| r9: "(STAR r) \<dagger> c = {SEQ r' (STAR r) | r'. r' \<in> r \<dagger> c}"
by (pat_completeness) (auto)
termination
dagger sorry
definition
OR :: "bool set \<Rightarrow> bool"
where
"OR S \<equiv> (\<exists>b\<in>S. b)"
function
matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool" ("_ ! _")
where
s01: "EMPTY ! s = (s =[])"
| s02: "CHAR c ! s = (s = [c])"
| s03: "ALT r1 r2 ! s = (r1 ! s \<or> r2 ! s)"
| s04: "STAR r ! [] = True"
| s05: "STAR r ! c#s = (False \<or> OR {SEQ (r') (STAR r)!s | r'. r' \<in> r \<dagger> c})"
| s06: "SEQ r1 r2 ! [] = (r1 ! [] \<and> r2 ! [])"
| s07: "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)"
| s08: "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)"
| s09: "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)"
| s10: "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \<or> (SEQ r12 r2) ! (c#s))"
| s11: "SEQ (STAR r1) r2 ! (c#s) =
(r2 ! (c#s) \<or> OR {SEQ r' (SEQ (STAR r1) r2) ! s | r'. r' \<in> r1 \<dagger> c})"
by (pat_completeness) (auto)
termination
matcher sorry
lemma "(CHAR a) ! [a]" by auto
lemma "\<not>(CHAR a) ! [a,a]" by auto
lemma "(STAR (CHAR a)) ! []" by auto
lemma "(STAR (CHAR a)) ! [a,a]" by (auto simp add: OR_def)
lemma "(SEQ (CHAR a) (SEQ (STAR (CHAR b)) (CHAR c))) ! [a,b,b,b,c]"
by (auto simp add: OR_def)
lemma holes:
assumes a: "Ls (r \<dagger> c) = {s. c#s \<in> L r}"
shows "Ls (r \<dagger> c) ; L (STAR r) = {s''. c#s'' \<in> L (STAR r)}"
proof -
have "Ls (r \<dagger> c) ; L (STAR r) = {s. c#s \<in> L r} ; L (STAR r)" by (simp add: a)
also have "\<dots> = {s'. c#s' \<in> (L r ; L (STAR r))}" sorry
also have "\<dots> = {s''. c#s'' \<in> L (STAR r)}" sorry
finally show "Ls (r \<dagger> c) ; L (STAR r) = {s''. c#s'' \<in> L (STAR r)}" by simp
qed
lemma eq:
shows "Ls (STAR r) \<dagger> c = (Ls (r \<dagger> c) ; L (STAR r))"
proof
show "Ls STAR r \<dagger> c \<subseteq> Ls r \<dagger> c ; L (STAR r)"
by (auto simp add: lang_star_def lang_seq_def Ls_def) (blast)
next
show "Ls r \<dagger> c ; L (STAR r) \<subseteq> Ls STAR r \<dagger> c"
apply(auto simp add: lang_star_def lang_seq_def Ls_def)
apply(rule_tac x="SEQ xa (STAR r)" in exI)
apply(simp add: lang_star_def lang_seq_def)
apply(blast)
done
qed
(* correctness of the matcher *)
lemma dagger_holes:
"Ls (r \<dagger> c) = {s. c#s \<in> L r}"
proof (induct rule: dagger.induct)
case (1 c)
show "Ls (EMPTY \<dagger> c) = {s. c#s \<in> L EMPTY}"
by (simp add: Ls_def)
next
case (2 c' c)
show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}"
proof (cases "c=c'")
assume "c=c'"
then show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}"
by (simp add: Ls_def)
next
assume "c\<noteq>c'"
then show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}"
by (simp add: Ls_def)
qed
next
case (3 r1 r2 c)
have ih1: "Ls r1 \<dagger> c = {s. c#s \<in> L r1}" by fact
have ih2: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact
show "Ls (ALT r1 r2) \<dagger> c = {s. c#s \<in> L (ALT r1 r2)}"
by (auto simp add: Ls_union ih1 ih2)
next
case (4 r2 c)
have ih: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact
show "Ls (SEQ EMPTY r2) \<dagger> c = {s. c#s \<in> L (SEQ EMPTY r2)}"
by (simp add: ih lang_seq_empty)
next
case (5 c' r2 c)
show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}"
proof (cases "c=c'")
assume "c=c'"
then show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}"
by (simp add: Ls_def lang_seq_def)
next
assume "c\<noteq>c'"
then show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}"
by (simp add: Ls_def lang_seq_def)
qed
next
case (6 r11 r12 r2 c)
have ih: "Ls (SEQ r11 (SEQ r12 r2)) \<dagger> c = {s. c#s \<in> L (SEQ r11 (SEQ r12 r2))}" by fact
show "Ls (SEQ (SEQ r11 r12) r2) \<dagger> c = {s. c # s \<in> L (SEQ (SEQ r11 r12) r2)}"
by (simp add: ih lang_seq_assoc)
next
case (7 r11 r12 r2 c)
have ih1: "Ls (SEQ r11 r2) \<dagger> c = {s. c#s \<in> L (SEQ r11 r2)}" by fact
have ih2: "Ls (SEQ r12 r2) \<dagger> c = {s. c#s \<in> L (SEQ r12 r2)}" by fact
show "Ls (SEQ (ALT r11 r12) r2) \<dagger> c = {s. c#s \<in> L (SEQ (ALT r11 r12) r2)}"
by (auto simp add: Ls_union ih1 ih2 lang_seq_union)
next
case (8 r1 r2 c)
have ih1: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact
have ih2: "Ls r1 \<dagger> c = {s. c#s \<in> L r1}" by fact
show "Ls (SEQ (STAR r1) r2) \<dagger> c = {s. c#s \<in> L (SEQ (STAR r1) r2)}"
sorry
next
case (9 r c)
have ih: "Ls r \<dagger> c = {s. c#s \<in> L r}" by fact
show "Ls (STAR r) \<dagger> c = {s. c#s \<in> L (STAR r)}"
by (simp only: eq holes[OF ih] del: r9)
qed
(* correctness of the matcher *)
lemma macher_holes:
shows "r ! s \<Longrightarrow> s \<in> L r"
and "\<not> r ! s \<Longrightarrow> s \<notin> L r"
proof (induct rule: matcher.induct)
case (1 s)
{ case 1
have "EMPTY ! s" by fact
then show "s \<in> L EMPTY" by simp
next
case 2
have "\<not> EMPTY ! s" by fact
then show "s \<notin> L EMPTY" by simp
}
next
case (2 c s)
{ case 1
have "CHAR c ! s" by fact
then show "s \<in> L (CHAR c)" by simp
next
case 2
have "\<not> CHAR c ! s" by fact
then show "s \<notin> L (CHAR c)" by simp
}
next
case (3 r1 r2 s)
have ih1: "r1 ! s \<Longrightarrow> s \<in> L r1" by fact
have ih2: "\<not> r1 ! s \<Longrightarrow> s \<notin> L r1" by fact
have ih3: "r2 ! s \<Longrightarrow> s \<in> L r2" by fact
have ih4: "\<not> r2 ! s \<Longrightarrow> s \<notin> L r2" by fact
{ case 1
have "ALT r1 r2 ! s" by fact
then show "s \<in> L (ALT r1 r2)" by (auto simp add: ih1 ih3)
next
case 2
have "\<not> ALT r1 r2 ! s" by fact
then show "s \<notin> L (ALT r1 r2)" by (simp add: ih2 ih4)
}
next
case (4 r)
{ case 1
have "STAR r ! []" by fact
then show "[] \<in> L (STAR r)" by (simp add: lang_star_in_empty)
next
case 2
have "\<not> STAR r ! []" by fact
then show "[] \<notin> L (STAR r)" by (simp)
}
next
case (5 r c s)
have ih1: "\<And>rx. SEQ rx (STAR r) ! s \<Longrightarrow> s \<in> L (SEQ rx (STAR r))" by fact
have ih2: "\<And>rx. \<not>SEQ rx (STAR r) ! s \<Longrightarrow> s \<notin> L (SEQ rx (STAR r))" by fact
{ case 1
have as: "STAR r ! c#s" by fact
then have "\<exists>r' \<in> r \<dagger> c. SEQ r' (STAR r) ! s" by (auto simp add: OR_def)
then obtain r' where imp1: "r' \<in> r \<dagger> c" and imp2: "SEQ r' (STAR r) ! s" by blast
from imp2 have "s \<in> L (SEQ r' (STAR r))" using ih1 by simp
then have "s \<in> L r' ; L (STAR r)" by simp
then have "c#s \<in> {[c]} ; (L r' ; L (STAR r))" by (simp add: lang_seq_def)
also have "\<dots> \<subseteq> L r ; L (STAR r)" using imp1
apply(auto simp add: lang_seq_def) sorry
also have "\<dots> \<subseteq> L (STAR r)" by (simp add: lang_star_subseteq)
finally show "c#s \<in> L (STAR r)" by simp
next
case 2
have as: "\<not> STAR r ! c#s" by fact
then have "\<forall>r'\<in> r \<dagger> c. \<not> (SEQ r' (STAR r) ! s)"
by (auto simp add: OR_def)
then have "\<forall>r'\<in> r \<dagger> c. s \<notin> L (SEQ r' (STAR r))" using ih2 by auto
then obtain r' where "r'\<in> r \<dagger> c \<Longrightarrow> s \<notin> L (SEQ r' (STAR r))" by auto
show "c#s \<notin> L (STAR r)" sorry
}
next
case (6 r1 r2)
have ih1: "r1 ! [] \<Longrightarrow> [] \<in> L r1" by fact
have ih2: "\<not> r1 ! [] \<Longrightarrow> [] \<notin> L r1" by fact
have ih3: "r2 ! [] \<Longrightarrow> [] \<in> L r2" by fact
have ih4: "\<not> r2 ! [] \<Longrightarrow> [] \<notin> L r2" by fact
{ case 1
have as: "SEQ r1 r2 ! []" by fact
then have "r1 ! [] \<and> r2 ! []" by (simp)
then show "[] \<in> L (SEQ r1 r2)" using ih1 ih3 by (simp add: lang_seq_def)
next
case 2
have "\<not> SEQ r1 r2 ! []" by fact
then have "(\<not> r1 ! []) \<or> (\<not> r2 ! [])" by (simp)
then show "[] \<notin> L (SEQ r1 r2)" using ih2 ih4
by (auto simp add: lang_seq_def)
}
next
case (7 r2 c s)
have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact
have ih2: "\<not> r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact
{ case 1
have "SEQ EMPTY r2 ! c#s" by fact
then show "c#s \<in> L (SEQ EMPTY r2)"
using ih1 by (simp add: lang_seq_def)
next
case 2
have "\<not> SEQ EMPTY r2 ! c#s" by fact
then show "c#s \<notin> L (SEQ EMPTY r2)"
using ih2 by (simp add: lang_seq_def)
}
next
case (8 c' r2 c s)
have ih1: "\<lbrakk>c' = c; r2 ! s\<rbrakk> \<Longrightarrow> s \<in> L r2" by fact
have ih2: "\<lbrakk>c' = c; \<not>r2 ! s\<rbrakk> \<Longrightarrow> s \<notin> L r2" by fact
{ case 1
have "SEQ (CHAR c') r2 ! c#s" by fact
then show "c#s \<in> L (SEQ (CHAR c') r2)"
using ih1 by (auto simp add: lang_seq_def split: if_splits)
next
case 2
have "\<not> SEQ (CHAR c') r2 ! c#s" by fact
then show "c#s \<notin> L (SEQ (CHAR c') r2)"
using ih2 by (auto simp add: lang_seq_def)
}
next
case (9 r11 r12 r2 c s)
have ih1: "SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 (SEQ r12 r2))" by fact
have ih2: "\<not> SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 (SEQ r12 r2))" by fact
{ case 1
have "SEQ (SEQ r11 r12) r2 ! c#s" by fact
then show "c#s \<in> L (SEQ (SEQ r11 r12) r2)"
using ih1
apply(auto simp add: lang_seq_def)
apply(rule_tac x="s1@s1a" in exI)
apply(rule_tac x="s2a" in exI)
apply(simp)
apply(blast)
done
next
case 2
have "\<not> SEQ (SEQ r11 r12) r2 ! c#s" by fact
then show "c#s \<notin> L (SEQ (SEQ r11 r12) r2)"
using ih2 by (auto simp add: lang_seq_def)
}
next
case (10 r11 r12 r2 c s)
have ih1: "SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 r2)" by fact
have ih2: "\<not> SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 r2)" by fact
have ih3: "SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r12 r2)" by fact
have ih4: "\<not> SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r12 r2)" by fact
{ case 1
have "SEQ (ALT r11 r12) r2 ! c#s" by fact
then show "c#s \<in> L (SEQ (ALT r11 r12) r2)"
using ih1 ih3 by (auto simp add: lang_seq_union)
next
case 2
have "\<not> SEQ (ALT r11 r12) r2 ! c#s" by fact
then show " c#s \<notin> L (SEQ (ALT r11 r12) r2)"
using ih2 ih4 by (simp add: lang_seq_union)
}
next
case (11 r1 r2 c s)
have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact
have ih2: "\<not>r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact
{ case 1
have "SEQ (STAR r1) r2 ! c#s" by fact
then show "c#s \<in> L (SEQ (STAR r1) r2)"
using ih1 sorry
next
case 2
have "\<not> SEQ (STAR r1) r2 ! c#s" by fact
then show "c#s \<notin> L (SEQ (STAR r1) r2)"
using ih2 sorry
}
qed
end