Regular_Set.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 11 Jul 2013 16:46:05 +0100
changeset 384 60bcf13adb77
parent 379 8c4b6fb43ebe
permissions -rw-r--r--
comment by Chunhan
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
(*  Author: Tobias Nipkow, Alex Krauss  *)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     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
header "Regular sets"
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
theory Regular_Set
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     6
imports Main
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     7
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     8
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     9
type_synonym 'a lang = "'a list set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    10
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    11
definition conc :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang" (infixr "@@" 75) 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
"A @@ B = {xs@ys | xs ys. xs:A & ys:B}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    13
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    14
text {* checks the code preprocessor for set comprehensions *}
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    15
export_code conc checking SML
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    16
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    17
overloading lang_pow == "compow :: nat \<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
    18
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    19
  primrec lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    20
  "lang_pow 0 A = {[]}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    21
  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    22
end
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    23
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    24
text {* for code generation *}
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    25
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    26
definition lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    27
  lang_pow_code_def [code_abbrev]: "lang_pow = compow"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    28
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    29
lemma [code]:
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    30
  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    31
  "lang_pow 0 A = {[]}"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    32
  by (simp_all add: lang_pow_code_def)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    33
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    34
hide_const (open) lang_pow
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    35
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    36
definition star :: "'a lang \<Rightarrow> 'a lang" where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    37
"star A = (\<Union>n. A ^^ n)"
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
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    40
subsection{* @{term "op @@"} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    41
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    42
lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    43
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
    44
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    45
lemma concE[elim]: 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    46
assumes "w \<in> A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    47
obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    48
using assms by (auto simp: 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
    49
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    50
lemma conc_mono: "A \<subseteq> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> A @@ B \<subseteq> C @@ D"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    51
by (auto simp: 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
    52
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    53
lemma conc_empty[simp]: shows "{} @@ A = {}" and "A @@ {} = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    54
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    55
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    56
lemma conc_epsilon[simp]: shows "{[]} @@ A = A" and "A @@ {[]} = A"
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 (simp_all 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
    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 conc_assoc: "(A @@ B) @@ C = A @@ (B @@ C)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    60
by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    61
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    62
lemma conc_Un_distrib:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    63
shows "A @@ (B \<union> C) = A @@ B \<union> A @@ C"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    64
and   "(A \<union> B) @@ C = A @@ C \<union> B @@ C"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    65
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
    66
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    67
lemma 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
    68
shows "A @@ UNION I M = UNION I (%i. A @@ M i)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    69
and   "UNION I M @@ A = UNION I (%i. M i @@ A)"
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 auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    71
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    72
lemma conc_subset_lists: "A \<subseteq> lists S \<Longrightarrow> B \<subseteq> lists S \<Longrightarrow> A @@ B \<subseteq> lists S"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    73
by(fastforce simp: conc_def in_lists_conv_set)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    74
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    75
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    76
subsection{* @{term "A ^^ n"} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    77
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    78
lemma lang_pow_add: "A ^^ (n + m) = A ^^ n @@ A ^^ m"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    79
by (induct n) (auto simp: conc_assoc)
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 lang_pow_empty: "{} ^^ n = (if n = 0 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
    82
by (induct n) auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    83
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    84
lemma lang_pow_empty_Suc[simp]: "({}::'a lang) ^^ Suc n = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    85
by (simp add: lang_pow_empty)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    86
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    87
lemma conc_pow_comm:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    88
  shows "A @@ (A ^^ n) = (A ^^ n) @@ A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    89
by (induct n) (simp_all add: conc_assoc[symmetric])
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    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
lemma length_lang_pow_ub:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    92
  "ALL w : A. length w \<le> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<le> k*n"
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    93
by(induct n arbitrary: w) (fastforce simp: conc_def)+
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    94
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    95
lemma length_lang_pow_lb:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    96
  "ALL w : A. length w \<ge> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<ge> k*n"
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    97
by(induct n arbitrary: w) (fastforce simp: conc_def)+
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    98
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    99
lemma lang_pow_subset_lists: "A \<subseteq> lists S \<Longrightarrow> A ^^ n \<subseteq> lists S"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   100
by(induction n)(auto simp: conc_subset_lists[OF assms])
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   101
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   102
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   103
subsection{* @{const star} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   104
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   105
lemma star_subset_lists: "A \<subseteq> lists S \<Longrightarrow> star A \<subseteq> lists S"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   106
unfolding star_def by(blast dest: lang_pow_subset_lists)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   107
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   108
lemma star_if_lang_pow[simp]: "w : A ^^ n \<Longrightarrow> w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   109
by (auto simp: star_def)
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
lemma Nil_in_star[iff]: "[] : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   112
proof (rule star_if_lang_pow)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   113
  show "[] : A ^^ 0" 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
   114
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   115
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   116
lemma star_if_lang[simp]: assumes "w : A" shows "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   117
proof (rule star_if_lang_pow)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   118
  show "w : A ^^ 1" using `w : A` 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
   119
qed
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
lemma append_in_starI[simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   122
assumes "u : star A" and "v : star A" shows "u@v : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   123
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   124
  from `u : star A` obtain m where "u : A ^^ m" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   125
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   126
  from `v : star A` obtain n where "v : A ^^ n" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   127
  ultimately have "u@v : A ^^ (m+n)" by (simp add: lang_pow_add)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   128
  thus ?thesis 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
   129
qed
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 conc_star_star: "star A @@ star A = star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   132
by (auto simp: 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
   133
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   134
lemma conc_star_comm:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   135
  shows "A @@ star A = star A @@ A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   136
unfolding star_def conc_pow_comm conc_UNION_distrib
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   137
by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   138
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   139
lemma star_induct[consumes 1, case_names Nil append, induct 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
   140
assumes "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   141
  and "P []"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   142
  and step: "!!u v. u : A \<Longrightarrow> v : star A \<Longrightarrow> P v \<Longrightarrow> P (u@v)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   143
shows "P w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   144
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   145
  { fix n have "w : A ^^ n \<Longrightarrow> P w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   146
    by (induct n arbitrary: w) (auto intro: `P []` step star_if_lang_pow) }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   147
  with `w : star A` show "P w" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   148
qed
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
lemma star_empty[simp]: "star {} = {[]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   151
by (auto elim: star_induct)
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
lemma star_epsilon[simp]: "star {[]} = {[]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   154
by (auto elim: star_induct)
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
lemma star_idemp[simp]: "star (star A) = star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   157
by (auto elim: star_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   158
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   159
lemma star_unfold_left: "star A = A @@ star A \<union> {[]}" (is "?L = ?R")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   160
proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   161
  show "?L \<subseteq> ?R" by (rule, erule star_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
   162
qed auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   163
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   164
lemma concat_in_star: "set ws \<subseteq> A \<Longrightarrow> concat ws : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   165
by (induct ws) 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
   166
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   167
lemma in_star_iff_concat:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   168
  "w : star A = (EX ws. set ws \<subseteq> A & w = concat ws)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   169
  (is "_ = (EX ws. ?R w ws)")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   170
proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   171
  assume "w : star A" thus "EX ws. ?R w ws"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   172
  proof induct
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   173
    case Nil have "?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
   174
    thus ?case ..
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   175
  next
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   176
    case (append u v)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   177
    moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   178
    then obtain ws where "set ws \<subseteq> A \<and> v = concat ws" 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
   179
    ultimately have "?R (u@v) (u#ws)" 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
   180
    thus ?case ..
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   181
  qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   182
next
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   183
  assume "EX us. ?R w us" thus "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   184
  by (auto simp: concat_in_star)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   185
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   186
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   187
lemma star_conv_concat: "star A = {concat ws|ws. set ws \<subseteq> A}"
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   188
by (fastforce simp: in_star_iff_concat)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   189
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   190
lemma star_insert_eps[simp]: "star (insert [] A) = star(A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   191
proof-
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   192
  { fix us
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   193
    have "set us \<subseteq> insert [] A \<Longrightarrow> EX vs. concat us = concat vs \<and> set vs \<subseteq> A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   194
      (is "?P \<Longrightarrow> EX vs. ?Q vs")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   195
    proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   196
      let ?vs = "filter (%u. u \<noteq> []) us"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   197
      show "?P \<Longrightarrow> ?Q ?vs" by (induct us) auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   198
    qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   199
  } thus ?thesis by (auto simp: star_conv_concat)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   200
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   201
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   202
lemma star_decom: 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   203
  assumes a: "x \<in> star A" "x \<noteq> []"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   204
  shows "\<exists>a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> A \<and> b \<in> star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   205
using a by (induct rule: star_induct) (blast)+
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   206
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   207
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   208
subsection {* Left-Quotients of languages *}
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   209
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   210
definition Deriv :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   211
where "Deriv x A = { xs. x#xs \<in> A }"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   212
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   213
definition Derivs :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   214
where "Derivs xs A = { ys. xs @ ys \<in> A }"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   215
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   216
abbreviation 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   217
  Derivss :: "'a list \<Rightarrow> 'a lang set \<Rightarrow> 'a lang"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   218
where
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 372
diff changeset
   219
  "Derivss s As \<equiv> \<Union> (Derivs s ` As)"
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   220
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   221
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   222
lemma Deriv_empty[simp]:   "Deriv a {} = {}"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   223
  and Deriv_epsilon[simp]: "Deriv a {[]} = {}"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   224
  and Deriv_char[simp]:    "Deriv a {[b]} = (if a = b then {[]} else {})"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   225
  and Deriv_union[simp]:   "Deriv a (A \<union> B) = Deriv a A \<union> Deriv a B"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   226
  and Deriv_inter[simp]:   "Deriv a (A \<inter> B) = Deriv a A \<inter> Deriv a B"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   227
  and Deriv_compl[simp]:   "Deriv a (-A) = - Deriv a A"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   228
  and Deriv_Union[simp]:   "Deriv a (Union M) = Union(Deriv a ` M)"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   229
  and Deriv_UN[simp]:      "Deriv a (UN x:I. S x) = (UN x:I. Deriv a (S x))"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   230
by (auto simp: Deriv_def)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   231
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   232
lemma Der_conc [simp]: "Deriv c (A @@ B) = (Deriv c A) @@ B \<union> (if [] \<in> A then Deriv c B else {})"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   233
unfolding Deriv_def conc_def
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   234
by (auto simp add: Cons_eq_append_conv)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   235
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   236
lemma Deriv_star [simp]: "Deriv c (star A) = (Deriv c A) @@ star A"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   237
proof -
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   238
  have incl: "[] \<in> A \<Longrightarrow> Deriv c (star A) \<subseteq> (Deriv c A) @@ star A"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   239
    unfolding Deriv_def conc_def 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   240
    apply(auto simp add: Cons_eq_append_conv)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   241
    apply(drule star_decom)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   242
    apply(auto simp add: Cons_eq_append_conv)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   243
    done
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   244
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   245
  have "Deriv c (star A) = Deriv c (A @@ star A \<union> {[]})"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   246
    by (simp only: star_unfold_left[symmetric])
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   247
  also have "... = Deriv c (A @@ star A)"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   248
    by (simp only: Deriv_union) (simp)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   249
  also have "... =  (Deriv c A) @@ (star A) \<union> (if [] \<in> A then Deriv c (star A) else {})"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   250
    by simp
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   251
   also have "... =  (Deriv c A) @@ star A"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   252
    using incl by auto
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   253
  finally show "Deriv c (star A) = (Deriv c A) @@ star A" . 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   254
qed
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   255
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   256
lemma Deriv_diff[simp]: "Deriv c (A - B) = Deriv c A - Deriv c B"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   257
by(auto simp add: Deriv_def)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   258
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   259
lemma Deriv_lists[simp]: "c : S \<Longrightarrow> Deriv c (lists S) = lists S"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   260
by(auto simp add: Deriv_def)
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   261
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   262
lemma Derivs_simps [simp]:
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   263
  shows "Derivs [] A = A"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   264
  and   "Derivs (c # s) A = Derivs s (Deriv c A)"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   265
  and   "Derivs (s1 @ s2) A = Derivs s2 (Derivs s1 A)"
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   266
unfolding Derivs_def Deriv_def by auto
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   267
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
   268
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   269
subsection {* Arden's Lemma *}
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   270
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   271
lemma arden_helper:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   272
  assumes eq: "X = A @@ X \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   273
  shows "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   274
proof (induct n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   275
  case 0 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   276
  show "X = (A ^^ Suc 0) @@ X \<union> (\<Union>m\<le>0. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   277
    using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   278
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   279
  case (Suc n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   280
  have ih: "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   281
  also have "\<dots> = (A ^^ Suc n) @@ (A @@ X \<union> B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   282
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> ((A ^^ Suc n) @@ B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   283
    by (simp add: conc_Un_distrib conc_assoc[symmetric] conc_pow_comm)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   284
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   285
    by (auto simp add: le_Suc_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   286
  finally show "X = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)" .
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   287
qed
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   288
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   289
lemma Arden:
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   290
  assumes "[] \<notin> A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   291
  shows "X = A @@ X \<union> B \<longleftrightarrow> X = star A @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   292
proof
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   293
  assume eq: "X = A @@ X \<union> B"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   294
  { fix w assume "w : X"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   295
    let ?n = "size w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   296
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   297
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   298
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   299
      by (metis length_lang_pow_lb nat_mult_1)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   300
    hence "ALL u : A^^(?n+1)@@X. length u \<ge> ?n+1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   301
      by(auto simp only: conc_def length_append)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   302
    hence "w \<notin> A^^(?n+1)@@X" by auto
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   303
    hence "w : star A @@ B" using `w : X` using arden_helper[OF eq, where n="?n"]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   304
      by (auto simp add: star_def conc_UNION_distrib)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   305
  } moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   306
  { fix w assume "w : star A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   307
    hence "EX n. w : A^^n @@ B" by(auto simp: conc_def star_def)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   308
    hence "w : X" using arden_helper[OF eq] by blast
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   309
  } ultimately show "X = star A @@ B" by blast 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   310
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   311
  assume eq: "X = star A @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   312
  have "star A = A @@ star A \<union> {[]}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   313
    by (rule star_unfold_left)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   314
  then have "star A @@ B = (A @@ star A \<union> {[]}) @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   315
    by metis
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   316
  also have "\<dots> = (A @@ star A) @@ B \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   317
    unfolding conc_Un_distrib by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   318
  also have "\<dots> = A @@ (star A @@ B) \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   319
    by (simp only: conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   320
  finally show "X = A @@ X \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   321
    using eq by blast 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   322
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   323
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   324
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   325
lemma reversed_arden_helper:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   326
  assumes eq: "X = X @@ A \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   327
  shows "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   328
proof (induct n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   329
  case 0 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   330
  show "X = X @@ (A ^^ Suc 0) \<union> (\<Union>m\<le>0. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   331
    using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   332
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   333
  case (Suc n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   334
  have ih: "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   335
  also have "\<dots> = (X @@ A \<union> B) @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   336
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (B @@ (A ^^ Suc n)) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   337
    by (simp add: conc_Un_distrib conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   338
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   339
    by (auto simp add: le_Suc_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   340
  finally show "X = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))" .
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   341
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   342
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   343
theorem reversed_Arden:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   344
  assumes nemp: "[] \<notin> A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   345
  shows "X = X @@ A \<union> B \<longleftrightarrow> X = B @@ star A"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   346
proof
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   347
 assume eq: "X = X @@ A \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   348
  { fix w assume "w : X"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   349
    let ?n = "size w"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   350
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   351
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   352
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   353
      by (metis length_lang_pow_lb nat_mult_1)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   354
    hence "ALL u : X @@ A^^(?n+1). length u \<ge> ?n+1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   355
      by(auto simp only: conc_def length_append)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   356
    hence "w \<notin> X @@ A^^(?n+1)" by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   357
    hence "w : B @@ star A" using `w : X` using reversed_arden_helper[OF eq, where n="?n"]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   358
      by (auto simp add: star_def conc_UNION_distrib)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   359
  } moreover
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   360
  { fix w assume "w : B @@ star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   361
    hence "EX n. w : B @@ A^^n" by (auto simp: conc_def star_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   362
    hence "w : X" using reversed_arden_helper[OF eq] by blast
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   363
  } ultimately show "X = B @@ star A" by blast 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   364
next 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   365
  assume eq: "X = B @@ star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   366
  have "star A = {[]} \<union> star A @@ A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   367
    unfolding conc_star_comm[symmetric]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   368
    by(metis Un_commute star_unfold_left)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   369
  then have "B @@ star A = B @@ ({[]} \<union> star A @@ A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   370
    by metis
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   371
  also have "\<dots> = B \<union> B @@ (star A @@ A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   372
    unfolding conc_Un_distrib by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   373
  also have "\<dots> = B \<union> (B @@ star A) @@ A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   374
    by (simp only: conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   375
  finally show "X = X @@ A \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   376
    using eq by blast 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   377
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   378
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   379
end