author | urbanc |
Thu, 28 Jul 2011 17:52:36 +0000 | |
changeset 180 | b755090d0f3d |
parent 170 | b1258b7d2789 |
child 203 | 5d724fe0e096 |
permissions | -rw-r--r-- |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
1 |
(* Author: Tobias Nipkow, Alex Krauss *) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
2 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
3 |
header "Regular sets" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
4 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
5 |
theory Regular_Set |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
6 |
imports Main |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
7 |
begin |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
8 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
9 |
type_synonym 'a lang = "'a list set" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
10 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
11 |
definition conc :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang" (infixr "@@" 75) where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
12 |
"A @@ B = {xs@ys | xs ys. xs:A & ys:B}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
13 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
14 |
overloading lang_pow == "compow :: nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
15 |
begin |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
16 |
primrec lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
17 |
"lang_pow 0 A = {[]}" | |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
18 |
"lang_pow (Suc n) A = A @@ (lang_pow n A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
19 |
end |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
20 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
21 |
definition star :: "'a lang \<Rightarrow> 'a lang" where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
22 |
"star A = (\<Union>n. A ^^ n)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
23 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
24 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
25 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
26 |
definition deriv :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
27 |
where "deriv x L = { xs. x#xs \<in> L }" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
28 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
29 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
30 |
coinductive bisimilar :: "'a list set \<Rightarrow> 'a list set \<Rightarrow> bool" where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
31 |
"([] \<in> K \<longleftrightarrow> [] \<in> L) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
32 |
\<Longrightarrow> (\<And>x. bisimilar (deriv x K) (deriv x L)) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
33 |
\<Longrightarrow> bisimilar K L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
34 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
35 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
36 |
subsection{* @{term "op @@"} *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
37 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
38 |
lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A @@ B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
39 |
by (auto simp add: conc_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
40 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
41 |
lemma concE[elim]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
42 |
assumes "w \<in> A @@ B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
43 |
obtains u v where "u \<in> A" "v \<in> B" "w = u@v" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
44 |
using assms by (auto simp: conc_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
45 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
46 |
lemma conc_mono: "A \<subseteq> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> A @@ B \<subseteq> C @@ D" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
47 |
by (auto simp: conc_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
48 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
49 |
lemma conc_empty[simp]: shows "{} @@ A = {}" and "A @@ {} = {}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
50 |
by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
51 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
52 |
lemma conc_epsilon[simp]: shows "{[]} @@ A = A" and "A @@ {[]} = A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
53 |
by (simp_all add:conc_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
54 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
55 |
lemma conc_assoc: "(A @@ B) @@ C = A @@ (B @@ C)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
56 |
by (auto elim!: concE) (simp only: append_assoc[symmetric] concI) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
57 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
58 |
lemma conc_Un_distrib: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
59 |
shows "A @@ (B \<union> C) = A @@ B \<union> A @@ C" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
60 |
and "(A \<union> B) @@ C = A @@ C \<union> B @@ C" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
61 |
by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
62 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
63 |
lemma conc_UNION_distrib: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
64 |
shows "A @@ UNION I M = UNION I (%i. A @@ M i)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
65 |
and "UNION I M @@ A = UNION I (%i. M i @@ A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
66 |
by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
67 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
68 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
69 |
subsection{* @{term "A ^^ n"} *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
70 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
71 |
lemma lang_pow_add: "A ^^ (n + m) = A ^^ n @@ A ^^ m" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
72 |
by (induct n) (auto simp: conc_assoc) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
73 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
74 |
lemma lang_pow_empty: "{} ^^ n = (if n = 0 then {[]} else {})" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
75 |
by (induct n) auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
76 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
77 |
lemma lang_pow_empty_Suc[simp]: "({}::'a lang) ^^ Suc n = {}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
78 |
by (simp add: lang_pow_empty) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
79 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
80 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
81 |
lemma length_lang_pow_ub: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
82 |
"ALL w : A. length w \<le> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<le> k*n" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
83 |
by(induct n arbitrary: w) (fastsimp simp: conc_def)+ |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
84 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
85 |
lemma length_lang_pow_lb: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
86 |
"ALL w : A. length w \<ge> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<ge> k*n" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
87 |
by(induct n arbitrary: w) (fastsimp simp: conc_def)+ |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
88 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
89 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
90 |
subsection{* @{const star} *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
91 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
92 |
lemma star_if_lang_pow[simp]: "w : A ^^ n \<Longrightarrow> w : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
93 |
by (auto simp: star_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
94 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
95 |
lemma Nil_in_star[iff]: "[] : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
96 |
proof (rule star_if_lang_pow) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
97 |
show "[] : A ^^ 0" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
98 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
99 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
100 |
lemma star_if_lang[simp]: assumes "w : A" shows "w : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
101 |
proof (rule star_if_lang_pow) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
102 |
show "w : A ^^ 1" using `w : A` by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
103 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
104 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
105 |
lemma append_in_starI[simp]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
106 |
assumes "u : star A" and "v : star A" shows "u@v : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
107 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
108 |
from `u : star A` obtain m where "u : A ^^ m" by (auto simp: star_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
109 |
moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
110 |
from `v : star A` obtain n where "v : A ^^ n" by (auto simp: star_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
111 |
ultimately have "u@v : A ^^ (m+n)" by (simp add: lang_pow_add) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
112 |
thus ?thesis by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
113 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
114 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
115 |
lemma conc_star_star: "star A @@ star A = star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
116 |
by (auto simp: conc_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
117 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
118 |
lemma star_induct[consumes 1, case_names Nil append, induct set: star]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
119 |
assumes "w : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
120 |
and "P []" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
121 |
and step: "!!u v. u : A \<Longrightarrow> v : star A \<Longrightarrow> P v \<Longrightarrow> P (u@v)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
122 |
shows "P w" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
123 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
124 |
{ fix n have "w : A ^^ n \<Longrightarrow> P w" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
125 |
by (induct n arbitrary: w) (auto intro: `P []` step star_if_lang_pow) } |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
126 |
with `w : star A` show "P w" by (auto simp: star_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
127 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
128 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
129 |
lemma star_empty[simp]: "star {} = {[]}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
130 |
by (auto elim: star_induct) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
131 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
132 |
lemma star_epsilon[simp]: "star {[]} = {[]}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
133 |
by (auto elim: star_induct) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
134 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
135 |
lemma star_idemp[simp]: "star (star A) = star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
136 |
by (auto elim: star_induct) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
137 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
138 |
lemma star_unfold_left: "star A = A @@ star A \<union> {[]}" (is "?L = ?R") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
139 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
140 |
show "?L \<subseteq> ?R" by (rule, erule star_induct) auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
141 |
qed auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
142 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
143 |
lemma concat_in_star: "set ws \<subseteq> A \<Longrightarrow> concat ws : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
144 |
by (induct ws) simp_all |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
145 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
146 |
lemma in_star_iff_concat: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
147 |
"w : star A = (EX ws. set ws \<subseteq> A & w = concat ws)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
148 |
(is "_ = (EX ws. ?R w ws)") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
149 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
150 |
assume "w : star A" thus "EX ws. ?R w ws" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
151 |
proof induct |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
152 |
case Nil have "?R [] []" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
153 |
thus ?case .. |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
154 |
next |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
155 |
case (append u v) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
156 |
moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
157 |
then obtain ws where "set ws \<subseteq> A \<and> v = concat ws" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
158 |
ultimately have "?R (u@v) (u#ws)" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
159 |
thus ?case .. |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
160 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
161 |
next |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
162 |
assume "EX us. ?R w us" thus "w : star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
163 |
by (auto simp: concat_in_star) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
164 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
165 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
166 |
lemma star_conv_concat: "star A = {concat ws|ws. set ws \<subseteq> A}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
167 |
by (fastsimp simp: in_star_iff_concat) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
168 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
169 |
lemma star_insert_eps[simp]: "star (insert [] A) = star(A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
170 |
proof- |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
171 |
{ fix us |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
172 |
have "set us \<subseteq> insert [] A \<Longrightarrow> EX vs. concat us = concat vs \<and> set vs \<subseteq> A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
173 |
(is "?P \<Longrightarrow> EX vs. ?Q vs") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
174 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
175 |
let ?vs = "filter (%u. u \<noteq> []) us" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
176 |
show "?P \<Longrightarrow> ?Q ?vs" by (induct us) auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
177 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
178 |
} thus ?thesis by (auto simp: star_conv_concat) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
179 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
180 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
181 |
lemma Arden: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
182 |
assumes "[] \<notin> A" and "X = A @@ X \<union> B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
183 |
shows "X = star A @@ B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
184 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
185 |
{ fix n have "X = A^^(n+1)@@X \<union> (\<Union>i\<le>n. A^^i@@B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
186 |
proof(induct n) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
187 |
case 0 show ?case using `X = A @@ X \<union> B` by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
188 |
next |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
189 |
case (Suc n) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
190 |
have "X = A@@X Un B" by(rule assms(2)) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
191 |
also have "\<dots> = A@@(A^^(n+1)@@X \<union> (\<Union>i\<le>n. A^^i@@B)) Un B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
192 |
by(subst Suc)(rule refl) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
193 |
also have "\<dots> = A^^(n+2)@@X \<union> (\<Union>i\<le>n. A^^(i+1)@@B) Un B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
194 |
by(simp add:conc_UNION_distrib conc_assoc conc_Un_distrib) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
195 |
also have "\<dots> = A^^(n+2)@@X \<union> (UN i : {1..n+1}. A^^i@@B) \<union> B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
196 |
by(subst UN_le_add_shift)(rule refl) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
197 |
also have "\<dots> = A^^(n+2)@@X \<union> (UN i : {1..n+1}. A^^i@@B) \<union> A^^0@@B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
198 |
by(simp) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
199 |
also have "\<dots> = A^^(n+2)@@X \<union> (\<Union>i\<le>n+1. A^^i@@B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
200 |
by(auto simp: UN_le_eq_Un0) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
201 |
finally show ?case by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
202 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
203 |
} note 1 = this |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
204 |
{ fix w assume "w : X" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
205 |
let ?n = "size w" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
206 |
from `[] \<notin> A` have "ALL u : A. length u \<ge> 1" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
207 |
by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
208 |
hence "ALL u : A^^(?n+1). length u \<ge> ?n+1" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
209 |
by (metis length_lang_pow_lb nat_mult_1) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
210 |
hence "ALL u : A^^(?n+1)@@X. length u \<ge> ?n+1" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
211 |
by(auto simp only: conc_def length_append) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
212 |
hence "w \<notin> A^^(?n+1)@@X" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
213 |
hence "w : star A @@ B" using `w : X` 1[of ?n] |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
214 |
by(auto simp add: star_def conc_UNION_distrib) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
215 |
} moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
216 |
{ fix w assume "w : star A @@ B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
217 |
hence "EX n. w : A^^n @@ B" by(auto simp: conc_def star_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
218 |
hence "w : X" using 1 by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
219 |
} ultimately show ?thesis by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
220 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
221 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
222 |
subsection{* @{const deriv} *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
223 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
224 |
lemma deriv_empty[simp]: "deriv a {} = {}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
225 |
and deriv_epsilon[simp]: "deriv a {[]} = {}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
226 |
and deriv_char[simp]: "deriv a {[b]} = (if a = b then {[]} else {})" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
227 |
and deriv_union[simp]: "deriv a (A \<union> B) = deriv a A \<union> deriv a B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
228 |
by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
229 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
230 |
lemma deriv_conc_subset: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
231 |
"deriv a A @@ B \<subseteq> deriv a (A @@ B)" (is "?L \<subseteq> ?R") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
232 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
233 |
fix w assume "w \<in> ?L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
234 |
then obtain u v where "w = u @ v" "a # u \<in> A" "v \<in> B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
235 |
by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
236 |
then have "a # w \<in> A @@ B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
237 |
by (auto intro: concI[of "a # u", simplified]) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
238 |
thus "w \<in> ?R" by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
239 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
240 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
241 |
lemma deriv_conc1: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
242 |
assumes "[] \<notin> A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
243 |
shows "deriv a (A @@ B) = deriv a A @@ B" (is "?L = ?R") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
244 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
245 |
show "?L \<subseteq> ?R" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
246 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
247 |
fix w assume "w \<in> ?L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
248 |
then have "a # w \<in> A @@ B" by (simp add: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
249 |
then obtain x y |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
250 |
where aw: "a # w = x @ y" "x \<in> A" "y \<in> B" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
251 |
with `[] \<notin> A` obtain x' where "x = a # x'" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
252 |
by (cases x) auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
253 |
with aw have "w = x' @ y" "x' \<in> deriv a A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
254 |
by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
255 |
with `y \<in> B` show "w \<in> ?R" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
256 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
257 |
show "?R \<subseteq> ?L" by (fact deriv_conc_subset) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
258 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
259 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
260 |
lemma deriv_conc2: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
261 |
assumes "[] \<in> A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
262 |
shows "deriv a (A @@ B) = deriv a A @@ B \<union> deriv a B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
263 |
(is "?L = ?R") |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
264 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
265 |
show "?L \<subseteq> ?R" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
266 |
proof |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
267 |
fix w assume "w \<in> ?L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
268 |
then have "a # w \<in> A @@ B" by (simp add: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
269 |
then obtain x y |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
270 |
where aw: "a # w = x @ y" "x \<in> A" "y \<in> B" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
271 |
show "w \<in> ?R" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
272 |
proof (cases x) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
273 |
case Nil |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
274 |
with aw have "w \<in> deriv a B" by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
275 |
thus ?thesis .. |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
276 |
next |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
277 |
case (Cons b x') |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
278 |
with aw have "w = x' @ y" "x' \<in> deriv a A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
279 |
by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
280 |
with `y \<in> B` show "w \<in> ?R" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
281 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
282 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
283 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
284 |
from concI[OF `[] \<in> A`, simplified] |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
285 |
have "B \<subseteq> A @@ B" .. |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
286 |
then have "deriv a B \<subseteq> deriv a (A @@ B)" by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
287 |
with deriv_conc_subset[of a A B] |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
288 |
show "?R \<subseteq> ?L" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
289 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
290 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
291 |
lemma wlog_no_eps: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
292 |
assumes PB: "\<And>B. [] \<notin> B \<Longrightarrow> P B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
293 |
assumes preserved: "\<And>A. P A \<Longrightarrow> P (insert [] A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
294 |
shows "P A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
295 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
296 |
let ?B = "A - {[]}" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
297 |
have "P ?B" by (rule PB) auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
298 |
thus "P A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
299 |
proof cases |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
300 |
assume "[] \<in> A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
301 |
then have "(insert [] ?B) = A" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
302 |
with preserved[OF `P ?B`] |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
303 |
show ?thesis by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
304 |
qed auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
305 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
306 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
307 |
lemma deriv_insert_eps[simp]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
308 |
"deriv a (insert [] A) = deriv a A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
309 |
by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
310 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
311 |
lemma deriv_star[simp]: "deriv a (star A) = deriv a A @@ star A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
312 |
proof (induct A rule: wlog_no_eps) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
313 |
fix B :: "'a list set" assume "[] \<notin> B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
314 |
thus "deriv a (star B) = deriv a B @@ star B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
315 |
by (subst star_unfold_left) (simp add: deriv_conc1) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
316 |
qed auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
317 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
318 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
319 |
subsection{* @{const bisimilar} *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
320 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
321 |
lemma equal_if_bisimilar: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
322 |
assumes "bisimilar K L" shows "K = L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
323 |
proof (rule set_eqI) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
324 |
fix w |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
325 |
from `bisimilar K L` show "w \<in> K \<longleftrightarrow> w \<in> L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
326 |
proof (induct w arbitrary: K L) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
327 |
case Nil thus ?case by (auto elim: bisimilar.cases) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
328 |
next |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
329 |
case (Cons a w K L) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
330 |
from `bisimilar K L` have "bisimilar (deriv a K) (deriv a L)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
331 |
by (auto elim: bisimilar.cases) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
332 |
then have "w \<in> deriv a K \<longleftrightarrow> w \<in> deriv a L" by (rule Cons(1)) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
333 |
thus ?case by (auto simp: deriv_def) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
334 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
335 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
336 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
337 |
lemma language_coinduct: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
338 |
fixes R (infixl "\<sim>" 50) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
339 |
assumes "K \<sim> L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
340 |
assumes "\<And>K L. K \<sim> L \<Longrightarrow> ([] \<in> K \<longleftrightarrow> [] \<in> L)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
341 |
assumes "\<And>K L x. K \<sim> L \<Longrightarrow> deriv x K \<sim> deriv x L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
342 |
shows "K = L" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
343 |
apply (rule equal_if_bisimilar) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
344 |
apply (rule bisimilar.coinduct[of R, OF `K \<sim> L`]) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
345 |
apply (auto simp: assms) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
346 |
done |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
347 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
348 |
end |