Derivatives.thy
author urbanc
Wed, 01 Feb 2012 17:40:00 +0000
changeset 275 22b6bd498419
parent 246 161128ccb65a
child 379 8c4b6fb43ebe
permissions -rw-r--r--
more on intro
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
241
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
     1
header "Derivatives of regular expressions"
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
     2
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
     3
(* Author: Christian Urban *)
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
     4
170
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 Derivatives
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
     6
imports Regular_Exp
170
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
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
     9
text{* This theory is based on work by Brozowski \cite{Brzozowski64} and Antimirov \cite{Antimirov95}. *}
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    10
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    11
subsection {* Left-Quotients of languages *}
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    12
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    13
definition Deriv :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    14
where "Deriv x A = { xs. x#xs \<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
    15
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    16
definition Derivs :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    17
where "Derivs xs A = { ys. xs @ ys \<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
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
    19
abbreviation 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    20
  Derivss :: "'a list \<Rightarrow> 'a lang set \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    21
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    22
  "Derivss s As \<equiv> \<Union> (Derivs s) ` As"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    23
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    24
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    25
lemma Deriv_empty[simp]:   "Deriv a {} = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    26
  and Deriv_epsilon[simp]: "Deriv a {[]} = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    27
  and Deriv_char[simp]:    "Deriv a {[b]} = (if a = b then {[]} else {})"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    28
  and Deriv_union[simp]:   "Deriv a (A \<union> B) = Deriv a A \<union> Deriv a B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    29
by (auto simp: Deriv_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
    30
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    31
lemma Deriv_conc_subset:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    32
"Deriv a A @@ B \<subseteq> Deriv a (A @@ B)" (is "?L \<subseteq> ?R")
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    33
proof 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    34
  fix w assume "w \<in> ?L"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    35
  then obtain u v where "w = u @ v" "a # u \<in> A" "v \<in> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    36
    by (auto simp: Deriv_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    37
  then have "a # w \<in> A @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    38
    by (auto intro: concI[of "a # u", simplified])
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    39
  thus "w \<in> ?R" by (auto simp: Deriv_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    40
qed
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    41
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    42
lemma Der_conc [simp]:
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    43
  shows "Deriv c (A @@ B) = (Deriv c A) @@ B \<union> (if [] \<in> A then Deriv c B else {})"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    44
unfolding Deriv_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
    45
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
    46
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    47
lemma Deriv_star [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    48
  shows "Deriv c (star A) = (Deriv c A) @@ star 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
    49
proof -
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    50
  have incl: "[] \<in> A \<Longrightarrow> Deriv c (star A) \<subseteq> (Deriv c A) @@ star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    51
    unfolding Deriv_def conc_def 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    52
    apply(auto simp add: Cons_eq_append_conv)
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
    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
    54
    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
    55
    done
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    56
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    57
  have "Deriv c (star A) = Deriv c (A @@ star A \<union> {[]})"
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
    58
    by (simp only: star_unfold_left[symmetric])
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    59
  also have "... = Deriv c (A @@ star A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    60
    by (simp only: Deriv_union) (simp)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    61
  also have "... =  (Deriv c A) @@ (star A) \<union> (if [] \<in> A then Deriv c (star A) 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
    62
    by simp
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    63
   also have "... =  (Deriv c A) @@ star 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
    64
    using incl by auto
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    65
  finally show "Deriv c (star A) = (Deriv c A) @@ star 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
    66
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    67
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    68
lemma Derivs_simps [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    69
  shows "Derivs [] A = A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    70
  and   "Derivs (c # s) A = Derivs s (Deriv c A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    71
  and   "Derivs (s1 @ s2) A = Derivs s2 (Derivs s1 A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    72
unfolding Derivs_def Deriv_def by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    73
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    74
241
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
    75
subsection {* Brozowski's derivatives of regular expressions *}
170
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
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    78
  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
    79
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    80
  "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
    81
| "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
    82
| "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
    83
| "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
    84
| "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
    85
| "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
    86
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    87
fun
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    88
  deriv :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    89
where
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    90
  "deriv c (Zero) = Zero"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    91
| "deriv c (One) = Zero"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    92
| "deriv c (Atom c') = (if c = c' then One else Zero)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    93
| "deriv c (Plus r1 r2) = Plus (deriv c r1) (deriv c r2)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    94
| "deriv c (Times r1 r2) = 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    95
    (if nullable r1 then Plus (Times (deriv c r1) r2) (deriv c r2) else Times (deriv c r1) r2)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    96
| "deriv c (Star r) = Times (deriv 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
    97
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
    98
fun 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
    99
  derivs :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   100
where
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   101
  "derivs [] r = r"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   102
| "derivs (c # s) r = derivs s (deriv 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
   103
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   104
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   105
lemma nullable_iff:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   106
  shows "nullable r \<longleftrightarrow> [] \<in> 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
   107
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
   108
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   109
lemma Deriv_deriv:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   110
  shows "Deriv c (lang r) = lang (deriv c r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   111
by (induct r) (simp_all add: nullable_iff)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   112
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   113
lemma Derivs_derivs:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   114
  shows "Derivs s (lang r) = lang (derivs s r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   115
by (induct s arbitrary: r) (simp_all add: Deriv_deriv)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   116
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   117
246
urbanc
parents: 241
diff changeset
   118
subsection {* Antimirov's partial derivatives *}
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   119
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   120
abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   121
  "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
   122
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   123
fun
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   124
  pderiv :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp set"
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
where
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   126
  "pderiv c Zero = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   127
| "pderiv c One = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   128
| "pderiv c (Atom c') = (if c = c' then {One} else {})"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   129
| "pderiv c (Plus r1 r2) = (pderiv c r1) \<union> (pderiv c r2)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   130
| "pderiv c (Times r1 r2) = 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   131
    (if nullable r1 then Timess (pderiv c r1) r2 \<union> pderiv c r2 else Timess (pderiv c r1) r2)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   132
| "pderiv c (Star r) = Timess (pderiv 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
   133
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   134
fun
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   135
  pderivs :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   136
where
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   137
  "pderivs [] r = {r}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   138
| "pderivs (c # s) r = \<Union> (pderivs s) ` (pderiv 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
   139
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   140
abbreviation
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   141
 pderiv_set :: "'a \<Rightarrow> 'a rexp set \<Rightarrow> 'a rexp set"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   142
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   143
  "pderiv_set c rs \<equiv> \<Union> pderiv 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
   144
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   145
abbreviation
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   146
  pderivs_set :: "'a list \<Rightarrow> 'a rexp set \<Rightarrow> 'a rexp set"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   147
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   148
  "pderivs_set s rs \<equiv> \<Union> (pderivs s) ` rs"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   149
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   150
lemma pderivs_append:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   151
  "pderivs (s1 @ s2) r = \<Union> (pderivs s2) ` (pderivs s1 r)"
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   152
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
   153
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   154
lemma pderivs_snoc:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   155
  shows "pderivs (s @ [c]) r = pderiv_set c (pderivs s r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   156
by (simp add: pderivs_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
   157
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   158
lemma pderivs_simps [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   159
  shows "pderivs s Zero = (if s = [] then {Zero} else {})"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   160
  and   "pderivs s One = (if s = [] then {One} else {})"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   161
  and   "pderivs s (Plus r1 r2) = (if s = [] then {Plus r1 r2} else (pderivs s r1) \<union> (pderivs s r2))"
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   162
by (induct s) (simp_all)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   163
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   164
lemma pderivs_Atom:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   165
  shows "pderivs s (Atom c) \<subseteq> {Atom c, One}"
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   166
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
   167
246
urbanc
parents: 241
diff changeset
   168
subsection {* Relating left-quotients and partial derivatives *}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   169
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   170
lemma Deriv_pderiv:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   171
  shows "Deriv c (lang r) = \<Union> lang ` (pderiv c r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   172
by (induct r) (auto simp add: nullable_iff 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
   173
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   174
lemma Derivs_pderivs:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   175
  shows "Derivs s (lang r) = \<Union> lang ` (pderivs s r)"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   176
proof (induct s arbitrary: r)
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   177
  case (Cons c s)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   178
  have ih: "\<And>r. Derivs s (lang r) = \<Union> lang ` (pderivs s r)" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   179
  have "Derivs (c # s) (lang r) = Derivs s (Deriv c (lang r))" by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   180
  also have "\<dots> = Derivs s (\<Union> lang ` (pderiv c r))" by (simp add: Deriv_pderiv)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   181
  also have "\<dots> = Derivss s (lang ` (pderiv c r))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   182
    by (auto simp add:  Derivs_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   183
  also have "\<dots> = \<Union> lang ` (pderivs_set s (pderiv c r))"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   184
    using ih by auto
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   185
  also have "\<dots> = \<Union> lang ` (pderivs (c # s) r)" by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   186
  finally show "Derivs (c # s) (lang r) = \<Union> lang ` pderivs (c # s) r" .
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   187
qed (simp add: Derivs_def)
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   188
246
urbanc
parents: 241
diff changeset
   189
subsection {* Relating derivatives and partial derivatives *}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   190
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   191
lemma deriv_pderiv:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   192
  shows "(\<Union> lang ` (pderiv c r)) = lang (deriv c r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   193
unfolding Deriv_deriv[symmetric] Deriv_pderiv by simp
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   194
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   195
lemma derivs_pderivs:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   196
  shows "(\<Union> lang ` (pderivs s r)) = lang (derivs s r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   197
unfolding Derivs_derivs[symmetric] Derivs_pderivs by simp
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   198
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   199
246
urbanc
parents: 241
diff changeset
   200
subsection {* Finiteness property of partial derivatives *}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   201
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   202
definition
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   203
  pderivs_lang :: "'a lang \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp set"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   204
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   205
  "pderivs_lang A r \<equiv> \<Union>x \<in> A. pderivs x r"
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   206
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   207
lemma pderivs_lang_subsetI:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   208
  assumes "\<And>s. s \<in> A \<Longrightarrow> pderivs s r \<subseteq> C"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   209
  shows "pderivs_lang A r \<subseteq> C"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   210
using assms unfolding pderivs_lang_def by (rule UN_least)
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   211
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   212
lemma pderivs_lang_union:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   213
  shows "pderivs_lang (A \<union> B) r = (pderivs_lang A r \<union> pderivs_lang B r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   214
by (simp add: pderivs_lang_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   215
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   216
lemma pderivs_lang_subset:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   217
  shows "A \<subseteq> B \<Longrightarrow> pderivs_lang A r \<subseteq> pderivs_lang B r"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   218
by (auto simp add: pderivs_lang_def)
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   219
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   220
definition
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   221
  "UNIV1 \<equiv> UNIV - {[]}"
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   222
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   223
lemma pderivs_lang_Zero [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   224
  shows "pderivs_lang UNIV1 Zero = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   225
unfolding UNIV1_def pderivs_lang_def by auto
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   226
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   227
lemma pderivs_lang_One [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   228
  shows "pderivs_lang UNIV1 One = {}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   229
unfolding UNIV1_def pderivs_lang_def by (auto split: if_splits)
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   230
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   231
lemma pderivs_lang_Atom [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   232
  shows "pderivs_lang UNIV1 (Atom c) = {One}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   233
unfolding UNIV1_def pderivs_lang_def 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   234
apply(auto)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   235
apply(frule rev_subsetD)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   236
apply(rule pderivs_Atom)
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   237
apply(simp)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   238
apply(case_tac xa)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   239
apply(auto split: if_splits)
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   240
done
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   241
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   242
lemma pderivs_lang_Plus [simp]:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   243
  shows "pderivs_lang UNIV1 (Plus r1 r2) = pderivs_lang UNIV1 r1 \<union> pderivs_lang UNIV1 r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   244
unfolding UNIV1_def pderivs_lang_def by auto
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   245
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   246
246
urbanc
parents: 241
diff changeset
   247
text {* Non-empty suffixes of a string (needed for the cases of @{const Times} and @{const Star} below) *}
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   248
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   249
definition
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   250
  "PSuf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   251
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   252
lemma PSuf_snoc:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   253
  shows "PSuf (s @ [c]) = (PSuf s) @@ {[c]} \<union> {[c]}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   254
unfolding PSuf_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
   255
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
   256
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   257
lemma PSuf_Union:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   258
  shows "(\<Union>v \<in> PSuf s @@ {[c]}. f v) = (\<Union>v \<in> PSuf 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
   259
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
   260
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   261
lemma pderivs_lang_snoc:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   262
  shows "pderivs_lang (PSuf s @@ {[c]}) r = (pderiv_set c (pderivs_lang (PSuf s) r))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   263
unfolding pderivs_lang_def
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   264
by (simp add: PSuf_Union pderivs_snoc)
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   265
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   266
lemma pderivs_Times:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   267
  shows "pderivs s (Times r1 r2) \<subseteq> Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf 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
   268
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
   269
  case (snoc c s)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   270
  have ih: "pderivs s (Times r1 r2) \<subseteq> Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf 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
   271
    by fact
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   272
  have "pderivs (s @ [c]) (Times r1 r2) = pderiv_set c (pderivs s (Times r1 r2))" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   273
    by (simp add: pderivs_snoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   274
  also have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf 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
   275
    using ih by (auto) (blast)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   276
  also have "\<dots> = pderiv_set c (Timess (pderivs s r1) r2) \<union> pderiv_set c (pderivs_lang (PSuf 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
   277
    by (simp)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   278
  also have "\<dots> = pderiv_set c (Timess (pderivs s r1) r2) \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   279
    by (simp add: pderivs_lang_snoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   280
  also 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   281
  have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs s r1) r2) \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   282
    by auto
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   283
  also 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   284
  have "\<dots> \<subseteq> Timess (pderiv_set c (pderivs s r1)) r2 \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   285
    by (auto simp add: if_splits) (blast)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   286
  also have "\<dots> = Timess (pderivs (s @ [c]) r1) r2 \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   287
    by (simp add: pderivs_snoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   288
  also have "\<dots> \<subseteq> Timess (pderivs (s @ [c]) r1) r2 \<union> pderivs_lang (PSuf (s @ [c])) r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   289
    unfolding pderivs_lang_def by (auto simp add: PSuf_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
   290
  finally show ?case .
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   291
qed (simp) 
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   292
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   293
lemma pderivs_lang_Times_aux1:
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   294
  assumes a: "s \<in> UNIV1"
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   295
  shows "pderivs_lang (PSuf s) r \<subseteq> pderivs_lang UNIV1 r"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   296
using a unfolding UNIV1_def PSuf_def pderivs_lang_def by auto
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   297
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   298
lemma pderivs_lang_Times_aux2:
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   299
  assumes a: "s \<in> UNIV1"
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   300
  shows "Timess (pderivs s r1) r2 \<subseteq> Timess (pderivs_lang UNIV1 r1) r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   301
using a unfolding pderivs_lang_def by auto
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   302
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   303
lemma pderivs_lang_Times:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   304
  shows "pderivs_lang UNIV1 (Times r1 r2) \<subseteq> Timess (pderivs_lang UNIV1 r1) r2 \<union> pderivs_lang UNIV1 r2"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   305
apply(rule pderivs_lang_subsetI)
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   306
apply(rule subset_trans)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   307
apply(rule pderivs_Times)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   308
using pderivs_lang_Times_aux1 pderivs_lang_Times_aux2
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   309
apply(blast)
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   310
done
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   311
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   312
lemma pderivs_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
   313
  assumes a: "s \<noteq> []"
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   314
  shows "pderivs s (Star r) \<subseteq> Timess (pderivs_lang (PSuf 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
   315
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
   316
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
   317
  case (snoc c s)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   318
  have ih: "s \<noteq> [] \<Longrightarrow> pderivs s (Star r) \<subseteq> Timess (pderivs_lang (PSuf 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
   319
  { assume asm: "s \<noteq> []"
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   320
    have "pderivs (s @ [c]) (Star r) = pderiv_set c (pderivs s (Star r))" by (simp add: pderivs_snoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   321
    also have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs_lang (PSuf s) r) (Star r))"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   322
      using ih[OF asm] by (auto) (blast)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   323
    also have "\<dots> \<subseteq> Timess (pderiv_set c (pderivs_lang (PSuf s) r)) (Star r) \<union> pderiv c (Star r)"
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   324
      by (auto split: if_splits) (blast)+
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   325
    also have "\<dots> \<subseteq> Timess (pderivs_lang (PSuf (s @ [c])) r) (Star r) \<union> (Timess (pderiv c r) (Star r))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   326
      by (simp only: PSuf_snoc pderivs_lang_snoc pderivs_lang_union)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   327
         (auto simp add: pderivs_lang_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   328
    also have "\<dots> = Timess (pderivs_lang (PSuf (s @ [c])) r) (Star r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   329
      by (auto simp add: PSuf_snoc PSuf_Union pderivs_snoc pderivs_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
   330
    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
   331
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   332
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   333
  { 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
   334
    then have ?case
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   335
      apply (auto simp add: pderivs_lang_def pderivs_snoc PSuf_def)
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   336
      apply(rule_tac x = "[c]" in exI)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   337
      apply(auto)
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   338
      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
   339
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   340
  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
   341
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
   342
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   343
lemma pderivs_lang_Star:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   344
  shows "pderivs_lang UNIV1 (Star r) \<subseteq> Timess (pderivs_lang UNIV1 r) (Star r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   345
apply(rule pderivs_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
   346
apply(rule subset_trans)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   347
apply(rule pderivs_Star)
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   348
apply(simp add: UNIV1_def)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   349
apply(simp add: UNIV1_def PSuf_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   350
apply(auto simp add: pderivs_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
   351
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   352
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   353
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
   354
  assumes a: "finite A"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 181
diff changeset
   355
  shows "finite (Timess A r)"
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   356
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
   357
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   358
lemma finite_pderivs_lang_UNIV1:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   359
  shows "finite (pderivs_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
   360
apply(induct r)
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 191
diff changeset
   361
apply(simp_all add: 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   362
  finite_subset[OF pderivs_lang_Times]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   363
  finite_subset[OF pderivs_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
   364
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   365
    
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   366
lemma pderivs_lang_UNIV:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   367
  shows "pderivs_lang UNIV r = pderivs [] r \<union> pderivs_lang UNIV1 r"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   368
unfolding UNIV1_def pderivs_lang_def
191
f6a603be52d6 slight polishing
urbanc
parents: 190
diff changeset
   369
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
   370
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   371
lemma finite_pderivs_lang_UNIV:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   372
  shows "finite (pderivs_lang UNIV r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   373
unfolding pderivs_lang_UNIV
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   374
by (simp add: finite_pderivs_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
   375
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   376
lemma finite_pderivs_lang:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   377
  shows "finite (pderivs_lang A r)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 195
diff changeset
   378
by (metis finite_pderivs_lang_UNIV pderivs_lang_subset rev_finite_subset subset_UNIV)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   379
179
edacc141060f small improvements
urbanc
parents: 174
diff changeset
   380
241
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   381
subsection {* A regular expression matcher based on Brozowski's derivatives *}
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   382
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   383
fun
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   384
  matcher :: "'a rexp \<Rightarrow> 'a list \<Rightarrow> bool"
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   385
where
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   386
  "matcher r s = nullable (derivs s r)"
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   387
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   388
lemma matcher_correctness:
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   389
  shows "matcher r s \<longleftrightarrow> s \<in> lang r"
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   390
by (induct s arbitrary: r)
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   391
   (simp_all add: nullable_iff Deriv_deriv[symmetric] Deriv_def)
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   392
68d48522ea9a tuning on the derivatives and closures theories
urbanc
parents: 203
diff changeset
   393
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   394
end