Derivatives.thy
author urbanc
Mon, 25 Jul 2011 13:33:38 +0000
changeset 170 b1258b7d2789
child 174 2b414a8a7132
permissions -rw-r--r--
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     1
theory Derivatives
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     2
imports Myhill_2
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     3
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     4
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     5
section {* Left-Quotients and Derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     6
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     7
subsection {* Left-Quotients *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     8
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     9
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    10
  Delta :: "'a lang \<Rightarrow> 'a lang"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    11
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    12
  "Delta A = (if [] \<in> A then {[]} else {})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    13
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    14
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    15
  Der :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    16
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    17
  "Der c A \<equiv> {s. [c] @ s \<in> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    18
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    19
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    20
  Ders :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    21
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    22
  "Ders s A \<equiv> {s'. s @ s' \<in> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    23
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    24
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    25
  Ders_set :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    26
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    27
  "Ders_set A B \<equiv> {s' | s s'. s @ s' \<in> B \<and> s \<in> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    28
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    29
lemma Ders_set_Ders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    30
  shows "Ders_set A B = (\<Union>s \<in> A. Ders s B)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    31
unfolding Ders_set_def Ders_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    32
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    33
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    34
lemma Der_zero [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    35
  shows "Der c {} = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    36
unfolding Der_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    37
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    38
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    39
lemma Der_one [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    40
  shows "Der c {[]} = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    41
unfolding Der_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    42
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    43
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    44
lemma Der_atom [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    45
  shows "Der c {[d]} = (if c = d then {[]} else {})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    46
unfolding Der_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    47
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    48
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    49
lemma Der_union [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    50
  shows "Der c (A \<union> B) = Der c A \<union> Der c B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    51
unfolding Der_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    52
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    53
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    54
lemma Der_conc [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    55
  shows "Der c (A \<cdot> B) = (Der c A) \<cdot> B \<union> (Delta A \<cdot> Der c B)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    56
unfolding Der_def Delta_def conc_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    57
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
    58
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    59
lemma Der_star [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    60
  shows "Der c (A\<star>) = (Der c A) \<cdot> A\<star>"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    61
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    62
  have incl: "Delta A \<cdot> Der c (A\<star>) \<subseteq> (Der c A) \<cdot> A\<star>"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    63
    unfolding Der_def Delta_def 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    64
    apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    65
    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
    66
    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
    67
    done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    68
    
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    69
  have "Der c (A\<star>) = Der c ({[]} \<union> A \<cdot> A\<star>)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    70
    by (simp only: star_cases[symmetric])
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    71
  also have "... = Der c (A \<cdot> A\<star>)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    72
    by (simp only: Der_union Der_one) (simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    73
  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    74
    by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    75
  also have "... =  (Der c A) \<cdot> A\<star>"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    76
    using incl by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    77
  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    78
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    79
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    80
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    81
lemma Ders_singleton:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    82
  shows "Ders [c] A = Der c A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    83
unfolding Der_def Ders_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    84
by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    85
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    86
lemma Ders_append:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    87
  shows "Ders (s1 @ s2) A = Ders s2 (Ders s1 A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    88
unfolding Ders_def by simp 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    89
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    90
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    91
text {* Relating the Myhill-Nerode relation with left-quotients. *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    92
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    93
lemma MN_Rel_Ders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    94
  shows "x \<approx>A y \<longleftrightarrow> Ders x A = Ders y A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    95
unfolding Ders_def str_eq_def str_eq_rel_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    96
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    97
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    98
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    99
subsection {* Brozowsky's derivatives of regular expressions *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   100
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   101
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   102
  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
   103
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   104
  "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
   105
| "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
   106
| "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
   107
| "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
   108
| "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
   109
| "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
   110
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   111
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   112
  der :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   113
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   114
  "der c (Zero) = Zero"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   115
| "der c (One) = Zero"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   116
| "der c (Atom c') = (if c = c' then One else Zero)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   117
| "der c (Plus r1 r2) = Plus (der c r1) (der c r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   118
| "der c (Times r1 r2) = Plus (Times (der c r1) r2) (if nullable r1 then der c r2 else Zero)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   119
| "der c (Star r) = Times (der c r) (Star r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   120
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   121
function 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   122
  ders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   123
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   124
  "ders [] r = r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   125
| "ders (s @ [c]) r = der c (ders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   126
by (auto) (metis rev_cases)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   127
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   128
termination
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   129
  by (relation "measure (length o fst)") (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   130
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   131
lemma Delta_nullable:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   132
  shows "Delta (lang r) = (if nullable r then {[]} else {})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   133
unfolding Delta_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   134
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
   135
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   136
lemma Der_der:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   137
  shows "Der c (lang r) = lang (der c r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   138
by (induct r) (simp_all add: Delta_nullable)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   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
lemma Ders_ders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   141
  shows "Ders s (lang r) = lang (ders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   142
apply(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
   143
apply(simp add: Ders_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   144
apply(simp only: ders.simps)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   145
apply(simp only: Ders_append)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   146
apply(simp only: Ders_singleton)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   147
apply(simp only: Der_der)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   148
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   149
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   150
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   151
subsection {* Antimirov's Partial Derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   152
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   153
abbreviation
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   154
  "Times_set rs r \<equiv> {Times r' r | r'. r' \<in> rs}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   155
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   156
fun
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   157
  pder :: "'a \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   158
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   159
  "pder c Zero = {Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   160
| "pder c One = {Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   161
| "pder c (Atom c') = (if c = c' then {One} else {Zero})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   162
| "pder c (Plus r1 r2) = (pder c r1) \<union> (pder c r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   163
| "pder c (Times r1 r2) = Times_set (pder c r1) r2 \<union> (if nullable r1 then pder c r2 else {})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   164
| "pder c (Star r) = Times_set (pder c r) (Star r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   165
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   166
abbreviation
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   167
  "pder_set c rs \<equiv> \<Union>r \<in> rs. pder c r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   168
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   169
function 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   170
  pders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   171
where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   172
  "pders [] r = {r}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   173
| "pders (s @ [c]) r = pder_set c (pders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   174
by (auto) (metis rev_cases)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   175
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   176
termination
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   177
  by (relation "measure (length o fst)") (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   178
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   179
abbreviation
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   180
  "pders_set A r \<equiv> \<Union>s \<in> A. pders s r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   181
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   182
lemma pders_append:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   183
  "pders (s1 @ s2) r = \<Union> (pders s2) ` (pders s1 r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   184
apply(induct s2 arbitrary: s1 r 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
   185
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   186
apply(subst append_assoc[symmetric])
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   187
apply(simp only: pders.simps)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   188
apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   189
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   190
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   191
lemma pders_singleton:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   192
  "pders [c] r = pder c r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   193
apply(subst append_Nil[symmetric])
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   194
apply(simp only: pders.simps)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   195
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   196
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   197
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   198
lemma pders_set_lang:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   199
  shows "(\<Union> (lang ` pder_set c rs)) = (\<Union>r \<in> rs. (\<Union>lang ` (pder c r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   200
unfolding image_def 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   201
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   202
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   203
lemma pders_Zero [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   204
  shows "pders s Zero = {Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   205
by (induct s rule: rev_induct) (simp_all)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   206
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   207
lemma pders_One [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   208
  shows "pders s One = (if s = [] then {One} else {Zero})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   209
by (induct s rule: rev_induct) (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   210
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   211
lemma pders_Atom [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   212
  shows "pders s (Atom c) = (if s = [] then {Atom c} else (if s = [c] then {One} else {Zero}))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   213
by (induct s rule: rev_induct) (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   214
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   215
lemma pders_Plus [simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   216
  shows "pders s (Plus r1 r2) = (if s = [] then {Plus r1 r2} else (pders s r1) \<union> (pders s r2))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   217
by (induct s rule: rev_induct) (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   218
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   219
text {* Non-empty suffixes of a string *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   220
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   221
definition
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   222
  "Suf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   223
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   224
lemma Suf:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   225
  shows "Suf (s @ [c]) = (Suf s) \<cdot> {[c]} \<union> {[c]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   226
unfolding Suf_def conc_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   227
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
   228
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   229
lemma Suf_Union:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   230
  shows "(\<Union>v \<in> Suf s \<cdot> {[c]}. P v) = (\<Union>v \<in> Suf s. P (v @ [c]))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   231
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
   232
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   233
lemma pders_Times:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   234
  shows "pders s (Times r1 r2) \<subseteq> Times_set (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   235
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
   236
  case (snoc c s)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   237
  have ih: "pders s (Times r1 r2) \<subseteq> Times_set (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2)" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   238
    by fact
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   239
  have "pders (s @ [c]) (Times r1 r2) = pder_set c (pders s (Times r1 r2))" by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   240
  also have "\<dots> \<subseteq> pder_set c (Times_set (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   241
    using ih by (auto) (blast)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   242
  also have "\<dots> = pder_set c (Times_set (pders s r1) r2) \<union> pder_set c (\<Union>v \<in> Suf s. pders v r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   243
    by (simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   244
  also have "\<dots> = pder_set c (Times_set (pders s r1) r2) \<union> (\<Union>v \<in> Suf s. pder_set c (pders v r2))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   245
    by (simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   246
  also have "\<dots> \<subseteq> pder_set c (Times_set (pders s r1) r2) \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   247
    by (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   248
  also have "\<dots> \<subseteq> Times_set (pder_set c (pders s r1)) r2 \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   249
    by (auto simp add: if_splits) (blast)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   250
  also have "\<dots> = Times_set (pders (s @ [c]) r1) r2 \<union> (\<Union>v \<in> Suf (s @ [c]). pders v r2)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   251
    apply(subst (2) pders.simps)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   252
    apply(simp only: Suf)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   253
    apply(simp add: Suf_Union pders_singleton)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   254
    apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   255
    done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   256
  finally show ?case .
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   257
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
   258
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   259
lemma pders_Star:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   260
  assumes a: "s \<noteq> []"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   261
  shows "pders s (Star r) \<subseteq> (\<Union>v \<in> Suf s. Times_set (pders v r) (Star r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   262
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
   263
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
   264
  case (snoc c s)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   265
  have ih: "s \<noteq> [] \<Longrightarrow> pders s (Star r) \<subseteq> (\<Union>v\<in>Suf s. Times_set (pders v r) (Star r))" by fact
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   266
  { assume asm: "s \<noteq> []"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   267
    have "pders (s @ [c]) (Star r) = pder_set c (pders s (Star r))" by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   268
    also have "\<dots> \<subseteq> (pder_set c (\<Union>v\<in>Suf s. Times_set (pders v r) (Star r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   269
      using ih[OF asm] 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
   270
    also have "\<dots> = (\<Union>v\<in>Suf s. pder_set c (Times_set (pders v r) (Star r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   271
      by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   272
    also have "\<dots> \<subseteq> (\<Union>v\<in>Suf s. (Times_set (pder_set c (pders v r)) (Star r) \<union> pder c (Star r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   273
      by (auto 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
   274
    also have "\<dots> = (\<Union>v\<in>Suf s. (Times_set (pder_set c (pders v r)) (Star r))) \<union> pder c (Star r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   275
      using asm by (auto simp add: Suf_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   276
    also have "\<dots> = (\<Union>v\<in>Suf s. (Times_set (pders (v @ [c]) r) (Star r))) \<union> (Times_set (pder c r) (Star r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   277
      by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   278
    also have "\<dots> = (\<Union>v\<in>Suf (s @ [c]). (Times_set (pders v r) (Star r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   279
      apply(simp only: Suf)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   280
      apply(simp add: Suf_Union pders_singleton)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   281
      apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   282
      done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   283
    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
   284
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   285
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   286
  { 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
   287
    then 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
   288
      apply(simp add: pders_singleton Suf_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   289
      apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   290
      apply(rule_tac x="[c]" in exI)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   291
      apply(simp add: pders_singleton)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   292
      done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   293
  }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   294
  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
   295
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
   296
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   297
abbreviation 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   298
  "UNIV1 \<equiv> UNIV - {[]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   299
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   300
lemma pders_set_Zero:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   301
  shows "pders_set UNIV1 Zero = {Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   302
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   303
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   304
lemma pders_set_One:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   305
  shows "pders_set UNIV1 One = {Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   306
by (auto 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
   307
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   308
lemma pders_set_Atom:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   309
  shows "pders_set UNIV1 (Atom c) \<subseteq> {One, Zero}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   310
by (auto 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
   311
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   312
lemma pders_set_Plus:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   313
  shows "pders_set UNIV1 (Plus r1 r2) = pders_set UNIV1 r1 \<union> pders_set UNIV1 r2"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   314
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   315
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   316
lemma pders_set_Times_aux:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   317
  assumes a: "s \<in> UNIV1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   318
  shows "pders_set (Suf s) r2 \<subseteq> pders_set UNIV1 r2"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   319
using a by (auto simp add: Suf_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   320
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   321
lemma pders_set_Times:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   322
  shows "pders_set UNIV1 (Times r1 r2) \<subseteq> Times_set (pders_set UNIV1 r1) r2 \<union> pders_set UNIV1 r2"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   323
apply(rule UN_least)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   324
apply(rule subset_trans)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   325
apply(rule pders_Times)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   326
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   327
apply(rule conjI) 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   328
apply(auto)[1]
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   329
apply(rule subset_trans)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   330
apply(rule pders_set_Times_aux)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   331
apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   332
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   333
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   334
lemma pders_set_Star:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   335
  shows "pders_set UNIV1 (Star r) \<subseteq> Times_set (pders_set UNIV1 r) (Star r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   336
apply(rule UN_least)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   337
apply(rule subset_trans)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   338
apply(rule pders_Star)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   339
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   340
apply(simp add: Suf_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   341
apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   342
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   343
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   344
lemma finite_Times_set:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   345
  assumes a: "finite A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   346
  shows "finite (Times_set A r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   347
using a by (auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   348
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   349
lemma finite_pders_set_UNIV1:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   350
  shows "finite (pders_set UNIV1 r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   351
apply(induct r)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   352
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   353
apply(simp only: pders_set_One)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   354
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   355
apply(rule finite_subset)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   356
apply(rule pders_set_Atom)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   357
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   358
apply(simp only: pders_set_Plus)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   359
apply(simp)
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(rule finite_subset)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   361
apply(rule pders_set_Times)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   362
apply(simp only: finite_Times_set finite_Un)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   363
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   364
apply(rule finite_subset)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   365
apply(rule pders_set_Star)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   366
apply(simp only: finite_Times_set)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   367
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   368
    
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   369
lemma pders_set_UNIV_UNIV1:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   370
  shows "pders_set UNIV r = pders [] r \<union> pders_set UNIV1 r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   371
apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   372
apply(rule_tac x="[]" in exI)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   373
apply(simp)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   374
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   375
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   376
lemma finite_pders_set_UNIV:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   377
  shows "finite (pders_set UNIV r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   378
unfolding pders_set_UNIV_UNIV1
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   379
by (simp add: finite_pders_set_UNIV1)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   380
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   381
lemma finite_pders_set:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   382
  shows "finite (pders_set A r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   383
apply(rule rev_finite_subset)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   384
apply(rule_tac r="r" in finite_pders_set_UNIV)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   385
apply(auto)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   386
done
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   387
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   388
lemma finite_pders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   389
  shows "finite (pders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   390
using finite_pders_set[where A="{s}" and r="r"]
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   391
by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   392
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   393
lemma finite_pders2:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   394
  shows "finite {pders s r | s. s \<in> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   395
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   396
  have "{pders s r | s. s \<in> A} \<subseteq> Pow (pders_set A r)" by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   397
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   398
  have "finite (Pow (pders_set A r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   399
    using finite_pders_set by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   400
  ultimately 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   401
  show "finite {pders s r | s. s \<in> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   402
    by(rule finite_subset)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   403
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   404
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   405
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   406
subsection {* Relating left-quotients and partial derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   407
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   408
lemma Der_pder:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   409
  shows "Der c (lang r) = \<Union> lang ` (pder c r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   410
by (induct r) (auto simp add: Delta_nullable conc_UNION_distrib)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   411
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   412
lemma Ders_pders:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   413
  shows "Ders s (lang r) = \<Union> lang ` (pders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   414
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
   415
  case (snoc c s)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   416
  have ih: "Ders s (lang r) = \<Union> lang ` (pders s r)" by fact
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   417
  have "Ders (s @ [c]) (lang r) = Ders [c] (Ders s (lang r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   418
    by (simp add: Ders_append)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   419
  also have "\<dots> = Der c (\<Union> lang ` (pders s r))" using ih
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   420
    by (simp add: Ders_singleton)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   421
  also have "\<dots> = (\<Union>r\<in>pders s r. Der c (lang r))" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   422
    unfolding Der_def image_def by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   423
  also have "\<dots> = (\<Union>r\<in>pders s r. (\<Union> lang `  (pder c r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   424
    by (simp add: Der_pder)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   425
  also have "\<dots> = (\<Union>lang ` (pder_set c (pders s r)))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   426
    by (simp add: pders_set_lang)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   427
  also have "\<dots> = (\<Union>lang ` (pders (s @ [c]) r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   428
    by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   429
  finally show "Ders (s @ [c]) (lang r) = \<Union> lang ` pders (s @ [c]) r" .
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   430
qed (simp add: Ders_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   431
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   432
lemma Ders_set_pders_set:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   433
  shows "Ders_set A (lang r) = (\<Union> lang ` (pders_set A r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   434
by (simp add: Ders_set_Ders Ders_pders)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   435
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   436
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   437
subsection {* Relating derivatives and partial derivatives *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   438
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   439
lemma
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   440
  shows "(\<Union> lang ` (pder c r)) = lang (der c r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   441
unfolding Der_der[symmetric] Der_pder by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   442
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   443
lemma
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   444
  shows "(\<Union> lang ` (pders s r)) = lang (ders s r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   445
unfolding Ders_ders[symmetric] Ders_pders by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   446
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   447
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   448
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   449
subsection {*
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   450
  The second direction of the Myhill-Nerode theorem using
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   451
  partial derivatives.
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   452
*}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   453
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   454
lemma Myhill_Nerode3:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   455
  fixes r::"'a rexp"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   456
  shows "finite (UNIV // \<approx>(lang r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   457
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   458
  have "finite (UNIV // =(\<lambda>x. pders x r)=)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   459
  proof - 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   460
    have "range (\<lambda>x. pders x r) = {pders s r | s. s \<in> UNIV}" by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   461
    moreover 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   462
    have "finite {pders s r | s. s \<in> UNIV}" by (rule finite_pders2)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   463
    ultimately
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   464
    have "finite (range (\<lambda>x. pders x r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   465
      by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   466
    then show "finite (UNIV // =(\<lambda>x. pders x r)=)" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   467
      by (rule finite_eq_tag_rel)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   468
  qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   469
  moreover 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   470
  have "=(\<lambda>x. pders x r)= \<subseteq> \<approx>(lang r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   471
    unfolding tag_eq_rel_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   472
    unfolding str_eq_def2
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   473
    unfolding MN_Rel_Ders
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   474
    unfolding Ders_pders
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   475
    by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   476
  moreover 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   477
  have "equiv UNIV =(\<lambda>x. pders x r)="
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   478
    unfolding equiv_def refl_on_def sym_def trans_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   479
    unfolding tag_eq_rel_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   480
    by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   481
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   482
  have "equiv UNIV (\<approx>(lang r))"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   483
    unfolding equiv_def refl_on_def sym_def trans_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   484
    unfolding str_eq_rel_def
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   485
    by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   486
  ultimately show "finite (UNIV // \<approx>(lang r))" 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   487
    by (rule refined_partition_finite)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   488
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   489
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   490
end