AFP-Submission/Derivatives.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 24 May 2016 11:36:21 +0100
changeset 191 6bb15b8e6301
permissions -rw-r--r--
added files that were submitted to afp
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
191
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
section "Derivatives of regular expressions"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
(* Author: Christian Urban *)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
theory Derivatives
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
imports Regular_Exp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
text{* This theory is based on work by Brozowski \cite{Brzozowski64} and Antimirov \cite{Antimirov95}. *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
subsection {* Brzozowski's derivatives of regular expressions *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
primrec
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
  deriv :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
  "deriv c (Zero) = Zero"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
| "deriv c (One) = Zero"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
| "deriv c (Atom c') = (if c = c' then One else Zero)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
| "deriv c (Plus r1 r2) = Plus (deriv c r1) (deriv c r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
| "deriv c (Times r1 r2) = 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
    (if nullable r1 then Plus (Times (deriv c r1) r2) (deriv c r2) else Times (deriv c r1) r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
| "deriv c (Star r) = Times (deriv c r) (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
primrec 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
  derivs :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  "derivs [] r = r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
| "derivs (c # s) r = derivs s (deriv c r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
lemma atoms_deriv_subset: "atoms (deriv x r) \<subseteq> atoms r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
by (induction r) (auto)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
lemma atoms_derivs_subset: "atoms (derivs w r) \<subseteq> atoms r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
by (induction w arbitrary: r) (auto dest: atoms_deriv_subset[THEN subsetD])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
lemma lang_deriv: "lang (deriv c r) = Deriv c (lang r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
by (induct r) (simp_all add: nullable_iff)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
lemma lang_derivs: "lang (derivs s r) = Derivs s (lang r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
by (induct s arbitrary: r) (simp_all add: lang_deriv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
text {* A regular expression matcher: *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
definition matcher :: "'a rexp \<Rightarrow> 'a list \<Rightarrow> bool" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
"matcher r s = nullable (derivs s r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
lemma matcher_correctness: "matcher r s \<longleftrightarrow> s \<in> lang r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
by (induct s arbitrary: r)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
   (simp_all add: nullable_iff lang_deriv matcher_def Deriv_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
subsection {* Antimirov's partial derivatives *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
abbreviation
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
  "Timess rs r \<equiv> (\<Union>r' \<in> rs. {Times r' r})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
primrec
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
  pderiv :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
  "pderiv c Zero = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
| "pderiv c One = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
| "pderiv c (Atom c') = (if c = c' then {One} else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
| "pderiv c (Plus r1 r2) = (pderiv c r1) \<union> (pderiv c r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
| "pderiv c (Times r1 r2) = 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
    (if nullable r1 then Timess (pderiv c r1) r2 \<union> pderiv c r2 else Timess (pderiv c r1) r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
| "pderiv c (Star r) = Timess (pderiv c r) (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
primrec
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
  pderivs :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
  "pderivs [] r = {r}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
| "pderivs (c # s) r = \<Union> (pderivs s ` pderiv c r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
abbreviation
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
 pderiv_set :: "'a \<Rightarrow> 'a rexp set \<Rightarrow> 'a rexp set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
  "pderiv_set c rs \<equiv> \<Union> (pderiv c ` rs)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
abbreviation
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
  pderivs_set :: "'a list \<Rightarrow> 'a rexp set \<Rightarrow> 'a rexp set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
  "pderivs_set s rs \<equiv> \<Union> (pderivs s ` rs)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
lemma pderivs_append:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
  "pderivs (s1 @ s2) r = \<Union> (pderivs s2 ` pderivs s1 r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
by (induct s1 arbitrary: r) (simp_all)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
lemma pderivs_snoc:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
  shows "pderivs (s @ [c]) r = pderiv_set c (pderivs s r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
by (simp add: pderivs_append)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
lemma pderivs_simps [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
  shows "pderivs s Zero = (if s = [] then {Zero} else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
  and   "pderivs s One = (if s = [] then {One} else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
  and   "pderivs s (Plus r1 r2) = (if s = [] then {Plus r1 r2} else (pderivs s r1) \<union> (pderivs s r2))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
by (induct s) (simp_all)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
lemma pderivs_Atom:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
  shows "pderivs s (Atom c) \<subseteq> {Atom c, One}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
by (induct s) (simp_all)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
subsection {* Relating left-quotients and partial derivatives *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
lemma Deriv_pderiv:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
  shows "Deriv c (lang r) = \<Union> (lang ` pderiv c r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
by (induct r) (auto simp add: nullable_iff conc_UNION_distrib)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
lemma Derivs_pderivs:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  shows "Derivs s (lang r) = \<Union> (lang ` pderivs s r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
proof (induct s arbitrary: r)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
  case (Cons c s)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
  have ih: "\<And>r. Derivs s (lang r) = \<Union> (lang ` pderivs s r)" by fact
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
  have "Derivs (c # s) (lang r) = Derivs s (Deriv c (lang r))" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
  also have "\<dots> = Derivs s (\<Union> (lang ` pderiv c r))" by (simp add: Deriv_pderiv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
  also have "\<dots> = Derivss s (lang ` (pderiv c r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
    by (auto simp add:  Derivs_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
  also have "\<dots> = \<Union> (lang ` (pderivs_set s (pderiv c r)))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
    using ih by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
  also have "\<dots> = \<Union> (lang ` (pderivs (c # s) r))" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
  finally show "Derivs (c # s) (lang r) = \<Union> (lang ` pderivs (c # s) r)" .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
qed (simp add: Derivs_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
subsection {* Relating derivatives and partial derivatives *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
lemma deriv_pderiv:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
  shows "\<Union> (lang ` (pderiv c r)) = lang (deriv c r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
unfolding lang_deriv Deriv_pderiv by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
lemma derivs_pderivs:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
  shows "\<Union> (lang ` (pderivs s r)) = lang (derivs s r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
unfolding lang_derivs Derivs_pderivs by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
subsection {* Finiteness property of partial derivatives *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
definition
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
  pderivs_lang :: "'a lang \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
  "pderivs_lang A r \<equiv> \<Union>x \<in> A. pderivs x r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
lemma pderivs_lang_subsetI:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
  assumes "\<And>s. s \<in> A \<Longrightarrow> pderivs s r \<subseteq> C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
  shows "pderivs_lang A r \<subseteq> C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
using assms unfolding pderivs_lang_def by (rule UN_least)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
lemma pderivs_lang_union:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
  shows "pderivs_lang (A \<union> B) r = (pderivs_lang A r \<union> pderivs_lang B r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
by (simp add: pderivs_lang_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
lemma pderivs_lang_subset:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
  shows "A \<subseteq> B \<Longrightarrow> pderivs_lang A r \<subseteq> pderivs_lang B r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
by (auto simp add: pderivs_lang_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
definition
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
  "UNIV1 \<equiv> UNIV - {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
lemma pderivs_lang_Zero [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
  shows "pderivs_lang UNIV1 Zero = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
unfolding UNIV1_def pderivs_lang_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
lemma pderivs_lang_One [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
  shows "pderivs_lang UNIV1 One = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
unfolding UNIV1_def pderivs_lang_def by (auto split: if_splits)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
lemma pderivs_lang_Atom [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  shows "pderivs_lang UNIV1 (Atom c) = {One}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
unfolding UNIV1_def pderivs_lang_def 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
apply(auto)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
apply(frule rev_subsetD)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
apply(rule pderivs_Atom)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
apply(simp)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
apply(case_tac xa)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
apply(auto split: if_splits)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
done
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
lemma pderivs_lang_Plus [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
  shows "pderivs_lang UNIV1 (Plus r1 r2) = pderivs_lang UNIV1 r1 \<union> pderivs_lang UNIV1 r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
unfolding UNIV1_def pderivs_lang_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
text {* Non-empty suffixes of a string (needed for the cases of @{const Times} and @{const Star} below) *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
definition
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
  "PSuf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
lemma PSuf_snoc:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
  shows "PSuf (s @ [c]) = (PSuf s) @@ {[c]} \<union> {[c]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
unfolding PSuf_def conc_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
lemma PSuf_Union:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
  shows "(\<Union>v \<in> PSuf s @@ {[c]}. f v) = (\<Union>v \<in> PSuf s. f (v @ [c]))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
by (auto simp add: conc_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
lemma pderivs_lang_snoc:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
  shows "pderivs_lang (PSuf s @@ {[c]}) r = (pderiv_set c (pderivs_lang (PSuf s) r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
unfolding pderivs_lang_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
by (simp add: PSuf_Union pderivs_snoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
lemma pderivs_Times:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
  shows "pderivs s (Times r1 r2) \<subseteq> Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf s) r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
proof (induct s rule: rev_induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
  case (snoc c s)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  have ih: "pderivs s (Times r1 r2) \<subseteq> Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf s) r2)" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
    by fact
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
  have "pderivs (s @ [c]) (Times r1 r2) = pderiv_set c (pderivs s (Times r1 r2))" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
    by (simp add: pderivs_snoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
  also have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs s r1) r2 \<union> (pderivs_lang (PSuf s) r2))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
    using ih by fast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
  also have "\<dots> = pderiv_set c (Timess (pderivs s r1) r2) \<union> pderiv_set c (pderivs_lang (PSuf s) r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
    by (simp)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
  also have "\<dots> = pderiv_set c (Timess (pderivs s r1) r2) \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
    by (simp add: pderivs_lang_snoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
  also 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
  have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs s r1) r2) \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
    by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
  also 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
  have "\<dots> \<subseteq> Timess (pderiv_set c (pderivs s r1)) r2 \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
    by (auto simp add: if_splits)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
  also have "\<dots> = Timess (pderivs (s @ [c]) r1) r2 \<union> pderiv c r2 \<union> pderivs_lang (PSuf s @@ {[c]}) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
    by (simp add: pderivs_snoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
  also have "\<dots> \<subseteq> Timess (pderivs (s @ [c]) r1) r2 \<union> pderivs_lang (PSuf (s @ [c])) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
    unfolding pderivs_lang_def by (auto simp add: PSuf_snoc)  
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
  finally show ?case .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
qed (simp) 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
lemma pderivs_lang_Times_aux1:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
  assumes a: "s \<in> UNIV1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
  shows "pderivs_lang (PSuf s) r \<subseteq> pderivs_lang UNIV1 r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
using a unfolding UNIV1_def PSuf_def pderivs_lang_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
lemma pderivs_lang_Times_aux2:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
  assumes a: "s \<in> UNIV1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
  shows "Timess (pderivs s r1) r2 \<subseteq> Timess (pderivs_lang UNIV1 r1) r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
using a unfolding pderivs_lang_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
lemma pderivs_lang_Times:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
  shows "pderivs_lang UNIV1 (Times r1 r2) \<subseteq> Timess (pderivs_lang UNIV1 r1) r2 \<union> pderivs_lang UNIV1 r2"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
apply(rule pderivs_lang_subsetI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
apply(rule subset_trans)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
apply(rule pderivs_Times)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
using pderivs_lang_Times_aux1 pderivs_lang_Times_aux2
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
apply(blast)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
done
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
lemma pderivs_Star:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
  assumes a: "s \<noteq> []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
  shows "pderivs s (Star r) \<subseteq> Timess (pderivs_lang (PSuf s) r) (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
using a
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
proof (induct s rule: rev_induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  case (snoc c s)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
  have ih: "s \<noteq> [] \<Longrightarrow> pderivs s (Star r) \<subseteq> Timess (pderivs_lang (PSuf s) r) (Star r)" by fact
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
  { assume asm: "s \<noteq> []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
    have "pderivs (s @ [c]) (Star r) = pderiv_set c (pderivs s (Star r))" by (simp add: pderivs_snoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
    also have "\<dots> \<subseteq> pderiv_set c (Timess (pderivs_lang (PSuf s) r) (Star r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
      using ih[OF asm] by fast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
    also have "\<dots> \<subseteq> Timess (pderiv_set c (pderivs_lang (PSuf s) r)) (Star r) \<union> pderiv c (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
      by (auto split: if_splits)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
    also have "\<dots> \<subseteq> Timess (pderivs_lang (PSuf (s @ [c])) r) (Star r) \<union> (Timess (pderiv c r) (Star r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
      by (simp only: PSuf_snoc pderivs_lang_snoc pderivs_lang_union)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
         (auto simp add: pderivs_lang_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
    also have "\<dots> = Timess (pderivs_lang (PSuf (s @ [c])) r) (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
      by (auto simp add: PSuf_snoc PSuf_Union pderivs_snoc pderivs_lang_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
    finally have ?case .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
  }
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
  moreover
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
  { assume asm: "s = []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
    then have ?case by (auto simp add: pderivs_lang_def pderivs_snoc PSuf_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
  }
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
  ultimately show ?case by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
qed (simp)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
lemma pderivs_lang_Star:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
  shows "pderivs_lang UNIV1 (Star r) \<subseteq> Timess (pderivs_lang UNIV1 r) (Star r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
apply(rule pderivs_lang_subsetI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
apply(rule subset_trans)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
apply(rule pderivs_Star)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
apply(simp add: UNIV1_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
apply(simp add: UNIV1_def PSuf_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
apply(auto simp add: pderivs_lang_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
done
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
lemma finite_Timess [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
  assumes a: "finite A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
  shows "finite (Timess A r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
using a by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
lemma finite_pderivs_lang_UNIV1:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
  shows "finite (pderivs_lang UNIV1 r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
apply(induct r)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
apply(simp_all add: 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
  finite_subset[OF pderivs_lang_Times]
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
  finite_subset[OF pderivs_lang_Star])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
done
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
    
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
lemma pderivs_lang_UNIV:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  shows "pderivs_lang UNIV r = pderivs [] r \<union> pderivs_lang UNIV1 r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
unfolding UNIV1_def pderivs_lang_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
lemma finite_pderivs_lang_UNIV:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
  shows "finite (pderivs_lang UNIV r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
unfolding pderivs_lang_UNIV
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
by (simp add: finite_pderivs_lang_UNIV1)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
lemma finite_pderivs_lang:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
  shows "finite (pderivs_lang A r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
by (metis finite_pderivs_lang_UNIV pderivs_lang_subset rev_finite_subset subset_UNIV)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
text{* The following relationship between the alphabetic width of regular expressions
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
(called @{text awidth} below) and the number of partial derivatives was proved
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
by Antimirov~\cite{Antimirov95} and formalized by Max Haslbeck. *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
fun awidth :: "'a rexp \<Rightarrow> nat" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
"awidth Zero = 0" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
"awidth One = 0" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
"awidth (Atom a) = 1" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
"awidth (Plus r1 r2) = awidth r1 + awidth r2" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
"awidth (Times r1 r2) = awidth r1 + awidth r2" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
"awidth (Star r1) = awidth r1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
lemma card_Timess_pderivs_lang_le:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
  "card (Timess (pderivs_lang A r) s) \<le> card (pderivs_lang A r)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
by (metis card_image_le finite_pderivs_lang image_eq_UN)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
lemma card_pderivs_lang_UNIV1_le_awidth: "card (pderivs_lang UNIV1 r) \<le> awidth r"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
proof (induction r)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
  case (Plus r1 r2)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
  have "card (pderivs_lang UNIV1 (Plus r1 r2)) = card (pderivs_lang UNIV1 r1 \<union> pderivs_lang UNIV1 r2)" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
  also have "\<dots> \<le> card (pderivs_lang UNIV1 r1) + card (pderivs_lang UNIV1 r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
    by(simp add: card_Un_le)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  also have "\<dots> \<le> awidth (Plus r1 r2)" using Plus.IH by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
  finally show ?case .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  case (Times r1 r2)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
  have "card (pderivs_lang UNIV1 (Times r1 r2)) \<le> card (Timess (pderivs_lang UNIV1 r1) r2 \<union> pderivs_lang UNIV1 r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
    by (simp add: card_mono finite_pderivs_lang pderivs_lang_Times)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
  also have "\<dots> \<le> card (Timess (pderivs_lang UNIV1 r1) r2) + card (pderivs_lang UNIV1 r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
    by (simp add: card_Un_le)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
  also have "\<dots> \<le> card (pderivs_lang UNIV1 r1) + card (pderivs_lang UNIV1 r2)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
    by (simp add: card_Timess_pderivs_lang_le)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
  also have "\<dots> \<le> awidth (Times r1 r2)" using Times.IH by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
  finally show ?case .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
  case (Star r)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
  have "card (pderivs_lang UNIV1 (Star r)) \<le> card (Timess (pderivs_lang UNIV1 r) (Star r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
    by (simp add: card_mono finite_pderivs_lang pderivs_lang_Star)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
  also have "\<dots> \<le> card (pderivs_lang UNIV1 r)" by (rule card_Timess_pderivs_lang_le)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
  also have "\<dots> \<le> awidth (Star r)" by (simp add: Star.IH)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
  finally show ?case .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
qed (auto)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
text{* Antimirov's Theorem 3.4: *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
theorem card_pderivs_lang_UNIV_le_awidth: "card (pderivs_lang UNIV r) \<le> awidth r + 1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
proof -
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
  have "card (insert r (pderivs_lang UNIV1 r)) \<le> Suc (card (pderivs_lang UNIV1 r))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
    by(auto simp: card_insert_if[OF finite_pderivs_lang_UNIV1])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
  also have "\<dots> \<le> Suc (awidth r)" by(simp add: card_pderivs_lang_UNIV1_le_awidth)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
  finally show ?thesis by(simp add: pderivs_lang_UNIV)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
qed 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
text{* Antimirov's Corollary 3.5: *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
corollary card_pderivs_lang_le_awidth: "card (pderivs_lang A r) \<le> awidth r + 1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
by(rule order_trans[OF
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
  card_mono[OF finite_pderivs_lang_UNIV pderivs_lang_subset[OF subset_UNIV]]
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
  card_pderivs_lang_UNIV_le_awidth])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
end