AFP-Submission/Regular_Set.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 10 Jun 2016 23:53:46 +0100
changeset 195 c2d36c3cf8ad
parent 191 6bb15b8e6301
permissions -rw-r--r--
run all posix tests
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
191
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
(*  Author: Tobias Nipkow, Alex Krauss, Christian Urban  *)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
section "Regular sets"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
theory Regular_Set
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
imports Main
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
type_synonym 'a lang = "'a list set"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
definition conc :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang" (infixr "@@" 75) where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
"A @@ B = {xs@ys | xs ys. xs:A & ys:B}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
text {* checks the code preprocessor for set comprehensions *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
export_code conc checking SML
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
overloading lang_pow == "compow :: nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
begin
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
  primrec lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
  "lang_pow 0 A = {[]}" |
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
end
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
text {* for code generation *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
definition lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  lang_pow_code_def [code_abbrev]: "lang_pow = compow"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
lemma [code]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  "lang_pow 0 A = {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
  by (simp_all add: lang_pow_code_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
hide_const (open) lang_pow
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
definition star :: "'a lang \<Rightarrow> 'a lang" where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
"star A = (\<Union>n. A ^^ n)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
subsection{* @{term "op @@"} *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
by (auto simp add: conc_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
lemma concE[elim]: 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
assumes "w \<in> A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
using assms by (auto simp: conc_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
lemma conc_mono: "A \<subseteq> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> A @@ B \<subseteq> C @@ D"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
by (auto simp: conc_def) 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
lemma conc_empty[simp]: shows "{} @@ A = {}" and "A @@ {} = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
lemma conc_epsilon[simp]: shows "{[]} @@ A = A" and "A @@ {[]} = A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
by (simp_all add:conc_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
lemma conc_assoc: "(A @@ B) @@ C = A @@ (B @@ C)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
lemma conc_Un_distrib:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
shows "A @@ (B \<union> C) = A @@ B \<union> A @@ C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
and   "(A \<union> B) @@ C = A @@ C \<union> B @@ C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
lemma conc_UNION_distrib:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
shows "A @@ UNION I M = UNION I (%i. A @@ M i)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
and   "UNION I M @@ A = UNION I (%i. M i @@ A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
lemma conc_subset_lists: "A \<subseteq> lists S \<Longrightarrow> B \<subseteq> lists S \<Longrightarrow> A @@ B \<subseteq> lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
by(fastforce simp: conc_def in_lists_conv_set)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
lemma Nil_in_conc[simp]: "[] \<in> A @@ B \<longleftrightarrow> [] \<in> A \<and> [] \<in> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
by (metis append_is_Nil_conv concE concI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
lemma concI_if_Nil1: "[] \<in> A \<Longrightarrow> xs : B \<Longrightarrow> xs \<in> A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
by (metis append_Nil concI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
lemma conc_Diff_if_Nil1: "[] \<in> A \<Longrightarrow> A @@ B = (A - {[]}) @@ B \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
by (fastforce elim: concI_if_Nil1)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
lemma concI_if_Nil2: "[] \<in> B \<Longrightarrow> xs : A \<Longrightarrow> xs \<in> A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
by (metis append_Nil2 concI)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
lemma conc_Diff_if_Nil2: "[] \<in> B \<Longrightarrow> A @@ B = A @@ (B - {[]}) \<union> A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
by (fastforce elim: concI_if_Nil2)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
lemma singleton_in_conc:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
  "[x] : A @@ B \<longleftrightarrow> [x] : A \<and> [] : B \<or> [] : A \<and> [x] : B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
by (fastforce simp: Cons_eq_append_conv append_eq_Cons_conv
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
       conc_Diff_if_Nil1 conc_Diff_if_Nil2)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
subsection{* @{term "A ^^ n"} *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
lemma lang_pow_add: "A ^^ (n + m) = A ^^ n @@ A ^^ m"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
by (induct n) (auto simp: conc_assoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
lemma lang_pow_empty: "{} ^^ n = (if n = 0 then {[]} else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
by (induct n) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
lemma lang_pow_empty_Suc[simp]: "({}::'a lang) ^^ Suc n = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
by (simp add: lang_pow_empty)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
lemma conc_pow_comm:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
  shows "A @@ (A ^^ n) = (A ^^ n) @@ A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
by (induct n) (simp_all add: conc_assoc[symmetric])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
lemma length_lang_pow_ub:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
  "ALL w : A. length w \<le> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<le> k*n"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
by(induct n arbitrary: w) (fastforce simp: conc_def)+
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
lemma length_lang_pow_lb:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
  "ALL w : A. length w \<ge> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<ge> k*n"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
by(induct n arbitrary: w) (fastforce simp: conc_def)+
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
lemma lang_pow_subset_lists: "A \<subseteq> lists S \<Longrightarrow> A ^^ n \<subseteq> lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
by(induction n)(auto simp: conc_subset_lists[OF assms])
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
subsection{* @{const star} *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
lemma star_subset_lists: "A \<subseteq> lists S \<Longrightarrow> star A \<subseteq> lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
unfolding star_def by(blast dest: lang_pow_subset_lists)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
lemma star_if_lang_pow[simp]: "w : A ^^ n \<Longrightarrow> w : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
by (auto simp: star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
lemma Nil_in_star[iff]: "[] : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
proof (rule star_if_lang_pow)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
  show "[] : A ^^ 0" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
lemma star_if_lang[simp]: assumes "w : A" shows "w : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
proof (rule star_if_lang_pow)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
  show "w : A ^^ 1" using `w : A` by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
lemma append_in_starI[simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
assumes "u : star A" and "v : star A" shows "u@v : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
proof -
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
  from `u : star A` obtain m where "u : A ^^ m" by (auto simp: star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
  moreover
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
  from `v : star A` obtain n where "v : A ^^ n" by (auto simp: star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
  ultimately have "u@v : A ^^ (m+n)" by (simp add: lang_pow_add)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
  thus ?thesis by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
lemma conc_star_star: "star A @@ star A = star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
by (auto simp: conc_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
lemma conc_star_comm:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
  shows "A @@ star A = star A @@ A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
unfolding star_def conc_pow_comm conc_UNION_distrib
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
lemma star_induct[consumes 1, case_names Nil append, induct set: star]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
assumes "w : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
  and "P []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
  and step: "!!u v. u : A \<Longrightarrow> v : star A \<Longrightarrow> P v \<Longrightarrow> P (u@v)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
shows "P w"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
proof -
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
  { fix n have "w : A ^^ n \<Longrightarrow> P w"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
    by (induct n arbitrary: w) (auto intro: `P []` step star_if_lang_pow) }
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  with `w : star A` show "P w" by (auto simp: star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
lemma star_empty[simp]: "star {} = {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
by (auto elim: star_induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
lemma star_epsilon[simp]: "star {[]} = {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
by (auto elim: star_induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
lemma star_idemp[simp]: "star (star A) = star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
by (auto elim: star_induct)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
lemma star_unfold_left: "star A = A @@ star A \<union> {[]}" (is "?L = ?R")
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
  show "?L \<subseteq> ?R" by (rule, erule star_induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
qed auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
lemma concat_in_star: "set ws \<subseteq> A \<Longrightarrow> concat ws : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
by (induct ws) simp_all
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
lemma in_star_iff_concat:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
  "w : star A = (EX ws. set ws \<subseteq> A & w = concat ws)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  (is "_ = (EX ws. ?R w ws)")
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
  assume "w : star A" thus "EX ws. ?R w ws"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
  proof induct
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
    case Nil have "?R [] []" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
    thus ?case ..
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
    case (append u v)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
    moreover
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
    then obtain ws where "set ws \<subseteq> A \<and> v = concat ws" by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
    ultimately have "?R (u@v) (u#ws)" by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
    thus ?case ..
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
  qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
  assume "EX us. ?R w us" thus "w : star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
  by (auto simp: concat_in_star)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
lemma star_conv_concat: "star A = {concat ws|ws. set ws \<subseteq> A}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
by (fastforce simp: in_star_iff_concat)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
lemma star_insert_eps[simp]: "star (insert [] A) = star(A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
proof-
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
  { fix us
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
    have "set us \<subseteq> insert [] A \<Longrightarrow> EX vs. concat us = concat vs \<and> set vs \<subseteq> A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
      (is "?P \<Longrightarrow> EX vs. ?Q vs")
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
    proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
      let ?vs = "filter (%u. u \<noteq> []) us"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
      show "?P \<Longrightarrow> ?Q ?vs" by (induct us) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
    qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
  } thus ?thesis by (auto simp: star_conv_concat)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
lemma star_unfold_left_Nil: "star A = (A - {[]}) @@ (star A) \<union> {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
by (metis insert_Diff_single star_insert_eps star_unfold_left)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
lemma star_Diff_Nil_fold: "(A - {[]}) @@ star A = star A - {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
proof -
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
  have "[] \<notin> (A - {[]}) @@ star A" by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
  thus ?thesis using star_unfold_left_Nil by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
lemma star_decom: 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
  assumes a: "x \<in> star A" "x \<noteq> []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
  shows "\<exists>a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> A \<and> b \<in> star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
using a by (induct rule: star_induct) (blast)+
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
subsection {* Left-Quotients of languages *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
definition Deriv :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
where "Deriv x A = { xs. x#xs \<in> A }"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
definition Derivs :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
where "Derivs xs A = { ys. xs @ ys \<in> A }"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
abbreviation 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
  Derivss :: "'a list \<Rightarrow> 'a lang set \<Rightarrow> 'a lang"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
  "Derivss s As \<equiv> \<Union> (Derivs s ` As)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
lemma Deriv_empty[simp]:   "Deriv a {} = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  and Deriv_epsilon[simp]: "Deriv a {[]} = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
  and Deriv_char[simp]:    "Deriv a {[b]} = (if a = b then {[]} else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
  and Deriv_union[simp]:   "Deriv a (A \<union> B) = Deriv a A \<union> Deriv a B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
  and Deriv_inter[simp]:   "Deriv a (A \<inter> B) = Deriv a A \<inter> Deriv a B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
  and Deriv_compl[simp]:   "Deriv a (-A) = - Deriv a A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
  and Deriv_Union[simp]:   "Deriv a (Union M) = Union(Deriv a ` M)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
  and Deriv_UN[simp]:      "Deriv a (UN x:I. S x) = (UN x:I. Deriv a (S x))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
by (auto simp: Deriv_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
lemma Der_conc [simp]: 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  shows "Deriv c (A @@ B) = (Deriv c A) @@ B \<union> (if [] \<in> A then Deriv c B else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
unfolding Deriv_def conc_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
by (auto simp add: Cons_eq_append_conv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
lemma Deriv_star [simp]: 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
  shows "Deriv c (star A) = (Deriv c A) @@ star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
proof -
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
  have "Deriv c (star A) = Deriv c ({[]} \<union> A @@ star A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
    by (metis star_unfold_left sup.commute)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
  also have "... = Deriv c (A @@ star A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
    unfolding Deriv_union by (simp)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
  also have "... = (Deriv c A) @@ (star A) \<union> (if [] \<in> A then Deriv c (star A) else {})"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
    by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
  also have "... =  (Deriv c A) @@ star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
    unfolding conc_def Deriv_def
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
    using star_decom by (force simp add: Cons_eq_append_conv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
  finally show "Deriv c (star A) = (Deriv c A) @@ star A" .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
lemma Deriv_diff[simp]:   
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
  shows "Deriv c (A - B) = Deriv c A - Deriv c B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
by(auto simp add: Deriv_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
lemma Deriv_lists[simp]: "c : S \<Longrightarrow> Deriv c (lists S) = lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
by(auto simp add: Deriv_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
lemma Derivs_simps [simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
  shows "Derivs [] A = A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
  and   "Derivs (c # s) A = Derivs s (Deriv c A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
  and   "Derivs (s1 @ s2) A = Derivs s2 (Derivs s1 A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
unfolding Derivs_def Deriv_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
lemma in_fold_Deriv: "v \<in> fold Deriv w L \<longleftrightarrow> w @ v \<in> L"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
  by (induct w arbitrary: L) (simp_all add: Deriv_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
lemma Derivs_alt_def: "Derivs w L = fold Deriv w L"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  by (induct w arbitrary: L) simp_all
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
subsection {* Shuffle product *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
fun shuffle where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
  "shuffle [] ys = {ys}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
| "shuffle xs [] = {xs}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
| "shuffle (x # xs) (y # ys) =
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
    {x # w | w . w \<in> shuffle xs (y # ys)} \<union>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
    {y # w | w . w \<in> shuffle (x # xs) ys}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
lemma shuffle_empty2[simp]: "shuffle xs [] = {xs}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
  by (cases xs) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
lemma Nil_in_shuffle[simp]: "[] \<in> shuffle xs ys \<longleftrightarrow> xs = [] \<and> ys = []"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
  by (induct xs ys rule: shuffle.induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
definition Shuffle (infixr "\<parallel>" 80) where
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  "Shuffle A B = \<Union>{shuffle xs ys | xs ys. xs \<in> A \<and> ys \<in> B}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
lemma shuffleE:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
  "zs \<in> shuffle xs ys \<Longrightarrow>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
    (zs = xs \<Longrightarrow> ys = [] \<Longrightarrow> P) \<Longrightarrow>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
    (zs = ys \<Longrightarrow> xs = [] \<Longrightarrow> P) \<Longrightarrow>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
    (\<And>x xs' z zs'. xs = x # xs' \<Longrightarrow> zs = z # zs' \<Longrightarrow> x = z \<Longrightarrow> zs' \<in> shuffle xs' ys \<Longrightarrow> P) \<Longrightarrow>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
    (\<And>y ys' z zs'. ys = y # ys' \<Longrightarrow> zs = z # zs' \<Longrightarrow> y = z \<Longrightarrow> zs' \<in> shuffle xs ys' \<Longrightarrow> P) \<Longrightarrow> P"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
  by (induct xs ys rule: shuffle.induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
lemma Cons_in_shuffle_iff:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
  "z # zs \<in> shuffle xs ys \<longleftrightarrow>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
    (xs \<noteq> [] \<and> hd xs = z \<and> zs \<in> shuffle (tl xs) ys \<or>
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
     ys \<noteq> [] \<and> hd ys = z \<and> zs \<in> shuffle xs (tl ys))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
  by (induct xs ys rule: shuffle.induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
lemma Deriv_Shuffle[simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  "Deriv a (A \<parallel> B) = Deriv a A \<parallel> B \<union> A \<parallel> Deriv a B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
  unfolding Shuffle_def Deriv_def by (fastforce simp: Cons_in_shuffle_iff neq_Nil_conv)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
lemma shuffle_subset_lists:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
  assumes "A \<subseteq> lists S" "B \<subseteq> lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
  shows "A \<parallel> B \<subseteq> lists S"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
unfolding Shuffle_def proof safe
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
  fix x and zs xs ys :: "'a list"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
  assume zs: "zs \<in> shuffle xs ys" "x \<in> set zs" and "xs \<in> A" "ys \<in> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
  with assms have "xs \<in> lists S" "ys \<in> lists S" by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
  with zs show "x \<in> S" by (induct xs ys arbitrary: zs rule: shuffle.induct) auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
lemma Nil_in_Shuffle[simp]: "[] \<in> A \<parallel> B \<longleftrightarrow> [] \<in> A \<and> [] \<in> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
  unfolding Shuffle_def by force
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
lemma shuffle_Un_distrib:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
shows "A \<parallel> (B \<union> C) = A \<parallel> B \<union> A \<parallel> C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
and   "A \<parallel> (B \<union> C) = A \<parallel> B \<union> A \<parallel> C"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
unfolding Shuffle_def by fast+
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
lemma shuffle_UNION_distrib:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
shows "A \<parallel> UNION I M = UNION I (%i. A \<parallel> M i)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
and   "UNION I M \<parallel> A = UNION I (%i. M i \<parallel> A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
unfolding Shuffle_def by fast+
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
lemma Shuffle_empty[simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
  "A \<parallel> {} = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
  "{} \<parallel> B = {}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
  unfolding Shuffle_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
lemma Shuffle_eps[simp]:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
  "A \<parallel> {[]} = A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
  "{[]} \<parallel> B = B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
  unfolding Shuffle_def by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
subsection {* Arden's Lemma *}
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   373
lemma arden_helper:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
  assumes eq: "X = A @@ X \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
  shows "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
proof (induct n)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
  case 0 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
  show "X = (A ^^ Suc 0) @@ X \<union> (\<Union>m\<le>0. (A ^^ m) @@ B)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
    using eq by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
  case (Suc n)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
  have ih: "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" by fact
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
  also have "\<dots> = (A ^^ Suc n) @@ (A @@ X \<union> B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" using eq by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> ((A ^^ Suc n) @@ B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
    by (simp add: conc_Un_distrib conc_assoc[symmetric] conc_pow_comm)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
    by (auto simp add: le_Suc_eq)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
  finally show "X = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)" .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
lemma Arden:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
  assumes "[] \<notin> A" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
  shows "X = A @@ X \<union> B \<longleftrightarrow> X = star A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
  assume eq: "X = A @@ X \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
  { fix w assume "w : X"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
    let ?n = "size w"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
      by (metis length_lang_pow_lb nat_mult_1)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
    hence "ALL u : A^^(?n+1)@@X. length u \<ge> ?n+1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
      by(auto simp only: conc_def length_append)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
    hence "w \<notin> A^^(?n+1)@@X" by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
    hence "w : star A @@ B" using `w : X` using arden_helper[OF eq, where n="?n"]
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
      by (auto simp add: star_def conc_UNION_distrib)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
  } moreover
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
  { fix w assume "w : star A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
    hence "EX n. w : A^^n @@ B" by(auto simp: conc_def star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
    hence "w : X" using arden_helper[OF eq] by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
  } ultimately show "X = star A @@ B" by blast 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
  assume eq: "X = star A @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
  have "star A = A @@ star A \<union> {[]}"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
    by (rule star_unfold_left)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
  then have "star A @@ B = (A @@ star A \<union> {[]}) @@ B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
    by metis
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
  also have "\<dots> = (A @@ star A) @@ B \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
    unfolding conc_Un_distrib by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
  also have "\<dots> = A @@ (star A @@ B) \<union> B" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
    by (simp only: conc_assoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
  finally show "X = A @@ X \<union> B" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
    using eq by blast 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
lemma reversed_arden_helper:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
  assumes eq: "X = X @@ A \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
  shows "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
proof (induct n)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
  case 0 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
  show "X = X @@ (A ^^ Suc 0) \<union> (\<Union>m\<le>0. B @@ (A ^^ m))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
    using eq by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
next
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
  case (Suc n)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
  have ih: "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" by fact
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
  also have "\<dots> = (X @@ A \<union> B) @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" using eq by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (B @@ (A ^^ Suc n)) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
    by (simp add: conc_Un_distrib conc_assoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
    by (auto simp add: le_Suc_eq)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
  finally show "X = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))" .
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   444
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
theorem reversed_Arden:
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
  assumes nemp: "[] \<notin> A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
  shows "X = X @@ A \<union> B \<longleftrightarrow> X = B @@ star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
proof
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
 assume eq: "X = X @@ A \<union> B"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
  { fix w assume "w : X"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
    let ?n = "size w"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
      by (metis length_lang_pow_lb nat_mult_1)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
    hence "ALL u : X @@ A^^(?n+1). length u \<ge> ?n+1"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
      by(auto simp only: conc_def length_append)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
    hence "w \<notin> X @@ A^^(?n+1)" by auto
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
    hence "w : B @@ star A" using `w : X` using reversed_arden_helper[OF eq, where n="?n"]
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
      by (auto simp add: star_def conc_UNION_distrib)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
  } moreover
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
  { fix w assume "w : B @@ star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
    hence "EX n. w : B @@ A^^n" by (auto simp: conc_def star_def)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
    hence "w : X" using reversed_arden_helper[OF eq] by blast
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
  } ultimately show "X = B @@ star A" by blast 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
next 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
  assume eq: "X = B @@ star A"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
  have "star A = {[]} \<union> star A @@ A" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   469
    unfolding conc_star_comm[symmetric]
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
    by(metis Un_commute star_unfold_left)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
  then have "B @@ star A = B @@ ({[]} \<union> star A @@ A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
    by metis
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
  also have "\<dots> = B \<union> B @@ (star A @@ A)"
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
    unfolding conc_Un_distrib by simp
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
  also have "\<dots> = B \<union> (B @@ star A) @@ A" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
    by (simp only: conc_assoc)
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
  finally show "X = X @@ A \<union> B" 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
    using eq by blast 
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
qed
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
6bb15b8e6301 added files that were submitted to afp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
end