Closure.thy
author urbanc
Wed, 18 May 2011 19:54:43 +0000
changeset 162 e93760534354
parent 5 074d9a4b2bc9
child 166 7743d2ad71d1
permissions -rw-r--r--
added directory for journal version; took uptodate version of the theory files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     1
theory Closure
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     2
imports Myhill_2
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
     3
begin
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
     4
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     5
section {* Closure properties of regular languages *}
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
     6
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     7
abbreviation
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     8
  regular :: "lang \<Rightarrow> bool"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
     9
where
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    10
  "regular A \<equiv> \<exists>r::rexp. A = L r"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    11
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    12
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    13
lemma closure_union[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    14
  assumes "regular A" "regular B" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    15
  shows "regular (A \<union> B)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    16
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    17
  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    18
  then have "A \<union> B = L (ALT r1 r2)" by simp
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    19
  then show "regular (A \<union> B)" by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    20
qed
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    21
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    22
lemma closure_seq[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    23
  assumes "regular A" "regular B" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    24
  shows "regular (A ;; B)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    25
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    26
  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    27
  then have "A ;; B = L (SEQ r1 r2)" by simp
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    28
  then show "regular (A ;; B)" by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    29
qed
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    30
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    31
lemma closure_star[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    32
  assumes "regular A"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    33
  shows "regular (A\<star>)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    34
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    35
  from assms obtain r::rexp where "L r = A" by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    36
  then have "A\<star> = L (STAR r)" by simp
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    37
  then show "regular (A\<star>)" by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    38
qed
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    39
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    40
lemma closure_complement[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    41
  assumes "regular A"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    42
  shows "regular (- A)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    43
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    44
  from assms have "finite (UNIV // \<approx>A)" by (simp add: Myhill_Nerode)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    45
  then have "finite (UNIV // \<approx>(-A))" by (simp add: str_eq_rel_def)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    46
  then show "regular (- A)" by (simp add: Myhill_Nerode)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    47
qed
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    48
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    49
lemma closure_difference[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    50
  assumes "regular A" "regular B" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    51
  shows "regular (A - B)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    52
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    53
  have "A - B = - (- A \<union> B)" by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    54
  moreover
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    55
  have "regular (- (- A \<union> B))" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    56
    using assms by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    57
  ultimately show "regular (A - B)" by simp
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    58
qed
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    59
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    60
lemma closure_intersection[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    61
  assumes "regular A" "regular B" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    62
  shows "regular (A \<inter> B)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    63
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    64
  have "A \<inter> B = - (- A \<union> - B)" by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    65
  moreover
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    66
  have "regular (- (- A \<union> - B))" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    67
    using assms by blast
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    68
  ultimately show "regular (A \<inter> B)" by simp
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    69
qed
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    70
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    71
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    72
text {* closure under string reversal *}
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    73
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    74
fun
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    75
  Rev :: "rexp \<Rightarrow> rexp"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    76
where
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    77
  "Rev NULL = NULL"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    78
| "Rev EMPTY = EMPTY"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    79
| "Rev (CHAR c) = CHAR c"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    80
| "Rev (ALT r1 r2) = ALT (Rev r1) (Rev r2)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    81
| "Rev (SEQ r1 r2) = SEQ (Rev r2) (Rev r1)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    82
| "Rev (STAR r) = STAR (Rev r)"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    83
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    84
lemma rev_Seq:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    85
  "(rev ` A) ;; (rev ` B) = rev ` (B ;; A)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    86
unfolding Seq_def image_def
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    87
apply(auto)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    88
apply(rule_tac x="xb @ xa" in exI)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    89
apply(auto)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    90
done
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    91
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    92
lemma rev_Star1:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    93
  assumes a: "s \<in> (rev ` A)\<star>"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    94
  shows "s \<in> rev ` (A\<star>)"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
    95
using a
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    96
proof(induct rule: star_induct)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    97
  case (step s1 s2)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    98
  have inj: "inj (rev::string \<Rightarrow> string)" unfolding inj_on_def by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
    99
  have "s1 \<in> rev ` A" "s2 \<in> rev ` (A\<star>)" by fact+
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   100
  then obtain x1 x2 where "x1 \<in> A" "x2 \<in> A\<star>" and eqs: "s1 = rev x1" "s2 = rev x2" by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   101
  then have "x1 \<in> A\<star>" "x2 \<in> A\<star>" by (auto intro: star_intro2)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   102
  then have "x2 @ x1 \<in> A\<star>" by (auto intro: star_intro1)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   103
  then have "rev (x2 @ x1) \<in> rev ` A\<star>" using inj by (simp only: inj_image_mem_iff)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   104
  then show "s1 @ s2 \<in>  rev ` A\<star>" using eqs by simp
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   105
qed (auto)
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
   106
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   107
lemma rev_Star2:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   108
  assumes a: "s \<in> A\<star>"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   109
  shows "rev s \<in> (rev ` A)\<star>"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   110
using a
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   111
proof(induct rule: star_induct)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   112
  case (step s1 s2)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   113
  have inj: "inj (rev::string \<Rightarrow> string)" unfolding inj_on_def by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   114
  have "s1 \<in> A"by fact
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   115
  then have "rev s1 \<in> rev ` A" using inj by (simp only: inj_image_mem_iff)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   116
  then have "rev s1 \<in> (rev ` A)\<star>" by (auto intro: star_intro2)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   117
  moreover
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   118
  have "rev s2 \<in> (rev ` A)\<star>" by fact
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   119
  ultimately show "rev (s1 @ s2) \<in>  (rev ` A)\<star>" by (auto intro: star_intro1)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   120
qed (auto)
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
   121
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   122
lemma rev_Star:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   123
  "(rev ` A)\<star> = rev ` (A\<star>)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   124
using rev_Star1 rev_Star2 by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   125
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   126
lemma rev_lang:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   127
  "L (Rev r) = rev ` (L r)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   128
by (induct r) (simp_all add: rev_Star rev_Seq image_Un)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   129
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   130
lemma closure_reversal[intro]:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   131
  assumes "regular A"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   132
  shows "regular (rev ` A)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   133
proof -
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   134
  from assms obtain r::rexp where "L r = A" by auto
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   135
  then have "L (Rev r) = rev ` A" by (simp add: rev_lang)
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   136
  then show "regular (rev` A)" by blast
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
   137
qed
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
   138
  
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents:
diff changeset
   139
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 5
diff changeset
   140
end