Derivatives.thy
author urbanc
Thu, 11 Aug 2011 23:11:39 +0000
changeset 193 2a5ac68db24b
parent 191 f6a603be52d6
child 195 5bbe63876f84
permissions -rw-r--r--
finished section about derivatives and closure properties
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
theory Derivatives
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     2
imports Myhill_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
begin
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
section {* Left-Quotients and Derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     6
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     7
subsection {* Left-Quotients *}
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
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    10
  Delta :: "'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
    11
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
  "Delta A = (if [] \<in> A 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
    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
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    15
  Der :: "'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
    16
where
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    17
  "Der c A \<equiv> {s'. [c] @ s' \<in> A}"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    18
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    19
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    20
  Ders :: "'a list \<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
    21
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
  "Ders s A \<equiv> {s'. s @ s' \<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
    23
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    24
abbreviation 
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    25
  "Derss s A \<equiv> \<Union> (Ders s) ` A"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    26
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    27
lemma Der_simps [simp]:
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    28
  shows "Der c {} = {}"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    29
  and   "Der c {[]} = {}"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    30
  and   "Der c {[d]} = (if c = d then {[]} else {})"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    31
  and   "Der c (A \<union> B) = Der c A \<union> Der c B"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    32
unfolding Der_def by auto
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    33
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    34
lemma Der_conc [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    35
  shows "Der c (A \<cdot> B) = (Der c A) \<cdot> B \<union> (Delta A \<cdot> Der c B)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    36
unfolding Der_def Delta_def 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
    37
by (auto simp add: Cons_eq_append_conv)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    38
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    39
lemma Der_star [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    40
  shows "Der c (A\<star>) = (Der c A) \<cdot> A\<star>"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    41
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    42
  have incl: "Delta A \<cdot> Der c (A\<star>) \<subseteq> (Der c A) \<cdot> A\<star>"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    43
    unfolding Der_def Delta_def conc_def 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    44
    apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    45
    apply(drule star_decom)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    46
    apply(auto simp add: Cons_eq_append_conv)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    47
    done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    48
    
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
    49
  have "Der c (A\<star>) = Der c (A \<cdot> A\<star> \<union> {[]})"
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
    50
    by (simp only: star_unfold_left[symmetric])
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    51
  also have "... = Der c (A \<cdot> A\<star>)"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    52
    by (simp only: Der_simps) (simp)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    53
  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    54
    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
    55
  also have "... =  (Der c A) \<cdot> A\<star>"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    56
    using incl 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
    57
  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    58
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    59
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    60
lemma Ders_simps [simp]:
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    61
  shows "Ders [] A = A"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    62
  and   "Ders (c # s) A = Ders s (Der c A)"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    63
  and   "Ders (s1 @ s2) A = Ders s2 (Ders s1 A)"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    64
unfolding Ders_def Der_def by auto
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    65
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    66
subsection {* Brozowsky's derivatives of regular expressions *}
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
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    69
  nullable :: "'a rexp \<Rightarrow> bool"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    70
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    71
  "nullable (Zero) = False"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    72
| "nullable (One) = True"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    73
| "nullable (Atom c) = False"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    74
| "nullable (Plus r1 r2) = (nullable r1 \<or> nullable r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    75
| "nullable (Times r1 r2) = (nullable r1 \<and> nullable r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    76
| "nullable (Star r) = True"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    77
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    78
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    79
  der :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    80
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    81
  "der c (Zero) = Zero"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    82
| "der c (One) = Zero"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    83
| "der c (Atom c') = (if c = c' then One else Zero)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    84
| "der c (Plus r1 r2) = Plus (der c r1) (der c r2)"
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 170
diff changeset
    85
| "der c (Times r1 r2) = 
2b414a8a7132 more on the section about derivatives
urbanc
parents: 170
diff changeset
    86
    (if nullable r1 then Plus (Times (der c r1) r2) (der c r2) else Times (der c r1) r2)"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    87
| "der c (Star r) = Times (der c r) (Star r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    88
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    89
fun 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    90
  ders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    91
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    92
  "ders [] r = r"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    93
| "ders (c # s) r = ders s (der c 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
    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
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    96
lemma Delta_nullable:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    97
  shows "Delta (lang r) = (if nullable r 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
    98
unfolding Delta_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    99
by (induct r) (auto simp add: conc_def split: if_splits)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   100
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   101
lemma Der_der:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   102
  shows "Der c (lang r) = lang (der c r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   103
by (induct r) (simp_all add: Delta_nullable)
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 Ders_ders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   106
  shows "Ders s (lang r) = lang (ders s r)"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   107
by (induct s arbitrary: r) (simp_all add: Der_der)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   108
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   109
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   110
subsection {* Antimirov's Partial Derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   111
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   112
abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   113
  "Timess rs r \<equiv> {Times r' r | r'. r' \<in> rs}"
170
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
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   116
  pder :: "'a \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   117
where
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   118
  "pder c Zero = {}"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   119
| "pder c One = {}"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   120
| "pder c (Atom c') = (if c = c' then {One} else {})"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   121
| "pder c (Plus r1 r2) = (pder c r1) \<union> (pder c r2)"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   122
| "pder c (Times r1 r2) = 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   123
    (if nullable r1 then Timess (pder c r1) r2 \<union> pder c r2 else Timess (pder c r1) r2)"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   124
| "pder c (Star r) = Timess (pder c r) (Star 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
   125
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   126
abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   127
  "pder_set c rs \<equiv> \<Union> pder c ` rs"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   128
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   129
fun
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   130
  pders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   131
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   132
  "pders [] r = {r}"
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
   133
| "pders (c # s) r = \<Union> (pders s) ` (pder c 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
   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
abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   136
  "pderss s A \<equiv> \<Union> (pders s) ` A"
170
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 pders_append:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   139
  "pders (s1 @ s2) r = \<Union> (pders s2) ` (pders s1 r)"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   140
by (induct s1 arbitrary: r) (simp_all)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   141
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   142
lemma pders_snoc:
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   143
  shows "pders (s @ [c]) r = pder_set c (pders s r)"
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   144
by (simp add: pders_append)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   145
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   146
lemma pders_simps [simp]:
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   147
  shows "pders s Zero = (if s= [] then {Zero} else {})"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   148
  and   "pders s One = (if s = [] then {One} else {})"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   149
  and   "pders s (Plus r1 r2) = (if s = [] then {Plus r1 r2} else (pders s r1) \<union> (pders s r2))"
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   150
by (induct s) (simp_all)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   151
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   152
lemma pders_Atom [intro]:
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   153
  shows "pders s (Atom c) \<subseteq> {Atom c, One}"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   154
by (induct s) (simp_all)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   155
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   156
subsection {* Relating left-quotients and partial derivatives *}
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   157
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   158
lemma Der_pder:
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   159
  shows "Der c (lang r) = \<Union> lang ` (pder c r)"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   160
by (induct r) (auto simp add: Delta_nullable conc_UNION_distrib)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   161
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   162
lemma Ders_pders:
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   163
  shows "Ders s (lang r) = \<Union> lang ` (pders s r)"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   164
proof (induct s arbitrary: r)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   165
  case (Cons c s)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   166
  have ih: "\<And>r. Ders s (lang r) = \<Union> lang ` (pders s r)" by fact
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   167
  have "Ders (c # s) (lang r) = Ders s (Der c (lang r))" by simp
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   168
  also have "\<dots> = Ders s (\<Union> lang ` (pder c r))" by (simp add: Der_pder)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   169
  also have "\<dots> = Derss s (lang ` (pder c r))"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   170
    by (auto simp add:  Ders_def)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   171
  also have "\<dots> = \<Union> lang ` (pderss s (pder c r))"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   172
    using ih by auto
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   173
  also have "\<dots> = \<Union> lang ` (pders (c # s) r)" by simp
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   174
  finally show "Ders (c # s) (lang r) = \<Union> lang ` pders (c # s) r" .
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   175
qed (simp add: Ders_def)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   176
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   177
subsection {* Relating derivatives and partial derivatives *}
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   178
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   179
lemma der_pder:
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   180
  shows "(\<Union> lang ` (pder c r)) = lang (der c r)"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   181
unfolding Der_der[symmetric] Der_pder by simp
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   182
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   183
lemma ders_pders:
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   184
  shows "(\<Union> lang ` (pders s r)) = lang (ders s r)"
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   185
unfolding Ders_ders[symmetric] Ders_pders by simp
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   186
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   188
subsection {* There are only finitely many partial derivatives for a language *}
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   189
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   190
definition
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   191
  "pders_lang A r \<equiv> \<Union>x \<in> A. pders x 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
   192
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   193
lemma pders_lang_subsetI [intro]:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   194
  assumes "\<And>s. s \<in> A \<Longrightarrow> pders s r \<subseteq> C"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   195
  shows "pders_lang A r \<subseteq> C"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   196
using assms unfolding pders_lang_def by (rule UN_least)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   197
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   198
lemma pders_lang_union:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   199
  shows "pders_lang (A \<union> B) r = (pders_lang A r \<union> pders_lang B r)"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   200
by (simp add: pders_lang_def)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   201
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   202
lemma pders_lang_subset:
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   203
  shows "A \<subseteq> B \<Longrightarrow> pders_lang A r \<subseteq> pders_lang B r"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   204
by (auto simp add: pders_lang_def)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   205
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   206
definition
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   207
  "UNIV1 \<equiv> UNIV - {[]}"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   208
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   209
lemma pders_lang_Zero [simp]:
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   210
  shows "pders_lang UNIV1 Zero = {}"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   211
unfolding UNIV1_def pders_lang_def by auto
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   212
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   213
lemma pders_lang_One [simp]:
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   214
  shows "pders_lang UNIV1 One = {}"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   215
unfolding UNIV1_def pders_lang_def by (auto split: if_splits)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   216
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   217
lemma pders_lang_Atom [simp]:
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   218
  shows "pders_lang UNIV1 (Atom c) = {One}"
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   219
unfolding UNIV1_def pders_lang_def 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   220
apply(auto)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   221
apply(frule rev_subsetD)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   222
apply(rule pders_Atom)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   223
apply(simp)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   224
apply(case_tac xa)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   225
apply(auto split: if_splits)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   226
done
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   227
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   228
lemma pders_lang_Plus [simp]:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   229
  shows "pders_lang UNIV1 (Plus r1 r2) = pders_lang UNIV1 r1 \<union> pders_lang UNIV1 r2"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   230
unfolding UNIV1_def pders_lang_def by auto
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   231
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   232
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   233
text {* Non-empty suffixes of a string (needed for teh cases of @{const Times} and @{const Star} *}
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   234
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   235
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   236
  "Suf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   237
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   238
lemma Suf_snoc:
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   239
  shows "Suf (s @ [c]) = (Suf s) \<cdot> {[c]} \<union> {[c]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   240
unfolding Suf_def 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
   241
by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   242
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   243
lemma Suf_Union:
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   244
  shows "(\<Union>v \<in> Suf s \<cdot> {[c]}. f v) = (\<Union>v \<in> Suf s. f (v @ [c]))"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   245
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
   246
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   247
lemma pders_lang_snoc:
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   248
  shows "pders_lang (Suf s \<cdot> {[c]}) r = (pder_set c (pders_lang (Suf s) r))"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   249
unfolding pders_lang_def
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   250
by (simp add: Suf_Union pders_snoc)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   251
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   252
lemma pders_Times:
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   253
  shows "pders s (Times r1 r2) \<subseteq> Timess (pders s r1) r2 \<union> (pders_lang (Suf s) r2)"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   254
proof (induct s rule: rev_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   255
  case (snoc c s)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   256
  have ih: "pders s (Times r1 r2) \<subseteq> Timess (pders s r1) r2 \<union> (pders_lang (Suf s) r2)" 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   257
    by fact
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   258
  have "pders (s @ [c]) (Times r1 r2) = pder_set c (pders s (Times r1 r2))" 
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   259
    by (simp add: pders_snoc)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   260
  also have "\<dots> \<subseteq> pder_set c (Timess (pders s r1) r2 \<union> (pders_lang (Suf s) r2))"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   261
    using ih by (auto) (blast)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   262
  also have "\<dots> = pder_set c (Timess (pders s r1) r2) \<union> pder_set c (pders_lang (Suf s) r2)"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   263
    by (simp)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   264
  also have "\<dots> = pder_set c (Timess (pders s r1) r2) \<union> pders_lang (Suf s \<cdot> {[c]}) r2"
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   265
    by (simp add: pders_lang_snoc)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   266
  also have "\<dots> \<subseteq> pder_set c (Timess (pders s r1) r2) \<union> pder c r2 \<union> pders_lang (Suf s \<cdot> {[c]}) r2"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   267
    by auto
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   268
  also have "\<dots> \<subseteq> Timess (pder_set c (pders s r1)) r2 \<union> pder c r2 \<union> pders_lang (Suf s \<cdot> {[c]}) r2"
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   269
    by (auto simp add: if_splits) (blast)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   270
  also have "\<dots> = Timess (pders (s @ [c]) r1) r2 \<union> pder c r2 \<union> pders_lang (Suf s \<cdot> {[c]}) r2"
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   271
    by (simp add: pders_snoc)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   272
  also have "\<dots> \<subseteq> Timess (pders (s @ [c]) r1) r2 \<union> pders_lang (Suf (s @ [c])) r2"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   273
    unfolding pders_lang_def by (auto simp add: Suf_snoc)  
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   274
  finally show ?case .
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   275
qed (simp) 
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   276
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   277
lemma pders_lang_Times_aux1:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   278
  assumes a: "s \<in> UNIV1"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   279
  shows "pders_lang (Suf s) r \<subseteq> pders_lang UNIV1 r"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   280
using a unfolding UNIV1_def Suf_def pders_lang_def by auto
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   281
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   282
lemma pders_lang_Times_aux2:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   283
  assumes a: "s \<in> UNIV1"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   284
  shows "Timess (pders s r1) r2 \<subseteq> Timess (pders_lang UNIV1 r1) r2"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   285
using a unfolding pders_lang_def by auto
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   286
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   287
lemma pders_lang_Times [intro]:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   288
  shows "pders_lang UNIV1 (Times r1 r2) \<subseteq> Timess (pders_lang UNIV1 r1) r2 \<union> pders_lang UNIV1 r2"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   289
apply(rule pders_lang_subsetI)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   290
apply(rule subset_trans)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   291
apply(rule pders_Times)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   292
using pders_lang_Times_aux1 pders_lang_Times_aux2
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   293
apply(blast)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   294
done
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   295
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   296
lemma pders_Star:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   297
  assumes a: "s \<noteq> []"
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   298
  shows "pders s (Star r) \<subseteq> Timess (pders_lang (Suf s) r) (Star 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
   299
using a
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   300
proof (induct s rule: rev_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   301
  case (snoc c s)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   302
  have ih: "s \<noteq> [] \<Longrightarrow> pders s (Star r) \<subseteq> Timess (pders_lang (Suf s) r) (Star r)" by fact
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   303
  { assume asm: "s \<noteq> []"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   304
    have "pders (s @ [c]) (Star r) = pder_set c (pders s (Star r))" by (simp add: pders_snoc)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   305
    also have "\<dots> \<subseteq> pder_set c (Timess (pders_lang (Suf s) r) (Star r))"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   306
      using ih[OF asm] by (auto) (blast)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   307
    also have "\<dots> \<subseteq> Timess (pder_set c (pders_lang (Suf s) r)) (Star r) \<union> pder c (Star r)"
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   308
      by (auto split: if_splits) (blast)+
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   309
    also have "\<dots> \<subseteq> Timess (pders_lang (Suf (s @ [c])) r) (Star r) \<union> (Timess (pder c r) (Star r))"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   310
      by (simp only: Suf_snoc pders_lang_snoc pders_lang_union)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   311
         (auto simp add: pders_lang_def)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   312
    also have "\<dots> = Timess (pders_lang (Suf (s @ [c])) r) (Star r)"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   313
      by (auto simp add: Suf_snoc Suf_Union pders_snoc pders_lang_def)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   314
    finally have ?case .
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   315
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   316
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   317
  { assume asm: "s = []"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   318
    then have ?case
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   319
      apply (auto simp add: pders_lang_def pders_snoc Suf_def)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   320
      apply(rule_tac x = "[c]" in exI)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   321
      apply(auto)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   322
      done
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   323
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   324
  ultimately show ?case 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
   325
qed (simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   326
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   327
lemma pders_lang_Star [intro]:
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   328
  shows "pders_lang UNIV1 (Star r) \<subseteq> Timess (pders_lang UNIV1 r) (Star r)"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   329
apply(rule pders_lang_subsetI)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   330
apply(rule subset_trans)
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   331
apply(rule pders_Star)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   332
apply(simp add: UNIV1_def)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   333
apply(simp add: UNIV1_def Suf_def)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   334
apply(auto simp add: pders_lang_def)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   335
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   336
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   337
lemma finite_Timess [simp]:
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   338
  assumes a: "finite A"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   339
  shows "finite (Timess A r)"
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   340
using a by auto
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   341
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   342
lemma finite_pders_lang_UNIV1:
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   343
  shows "finite (pders_lang UNIV1 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
   344
apply(induct r)
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   345
apply(simp_all add: 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   346
  finite_subset[OF pders_lang_Times]
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   347
  finite_subset[OF pders_lang_Star])
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   348
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   349
    
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   350
lemma pders_lang_UNIV:
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   351
  shows "pders_lang UNIV r = pders [] r \<union> pders_lang UNIV1 r"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   352
unfolding UNIV1_def pders_lang_def
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   353
by blast
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   354
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   355
lemma finite_pders_lang_UNIV:
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   356
  shows "finite (pders_lang UNIV r)"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   357
unfolding pders_lang_UNIV
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   358
by (simp add: finite_pders_lang_UNIV1)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   359
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   360
lemma finite_pders_lang:
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   361
  shows "finite (pders_lang A r)"
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   362
apply(rule rev_finite_subset[OF finite_pders_lang_UNIV])
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   363
apply(rule pders_lang_subset)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   364
apply(simp)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   365
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   366
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   367
text {* Relating the Myhill-Nerode relation with left-quotients. *}
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   368
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   369
lemma MN_Rel_Ders:
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   370
  shows "x \<approx>A y \<longleftrightarrow> Ders x A = Ders y A"
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   371
unfolding Ders_def str_eq_def
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   372
by auto
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   373
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   374
subsection {*
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   375
  The second direction of the Myhill-Nerode theorem using
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   376
  partial derivatives.
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   377
*}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   378
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   379
lemma Myhill_Nerode3:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   380
  fixes r::"'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   381
  shows "finite (UNIV // \<approx>(lang r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   382
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   383
  have "finite (UNIV // =(\<lambda>x. pders x r)=)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   384
  proof - 
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   385
    have "range (\<lambda>x. pders x r) \<subseteq> Pow (pders_lang UNIV r)"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   386
      unfolding pders_lang_def by auto
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   387
    moreover 
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   388
    have "finite (Pow (pders_lang UNIV r))" by (simp add: finite_pders_lang)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   389
    ultimately
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   390
    have "finite (range (\<lambda>x. pders x r))"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   391
      by (simp add: finite_subset)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   392
    then show "finite (UNIV // =(\<lambda>x. pders x r)=)" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   393
      by (rule finite_eq_tag_rel)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   394
  qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   395
  moreover 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   396
  have "=(\<lambda>x. pders x r)= \<subseteq> \<approx>(lang r)"
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   397
    unfolding tag_eq_def
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   398
    by (auto simp add: MN_Rel_Ders Ders_pders)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   399
  moreover 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   400
  have "equiv UNIV =(\<lambda>x. pders x r)="
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   401
  and  "equiv UNIV (\<approx>(lang 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
   402
    unfolding equiv_def refl_on_def sym_def trans_def
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   403
    unfolding tag_eq_def str_eq_def
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   404
    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
   405
  ultimately show "finite (UNIV // \<approx>(lang r))" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   406
    by (rule refined_partition_finite)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   407
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   408
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   409
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   410
end