Myhill.thy
author zhang
Sun, 30 Jan 2011 12:22:07 +0000
changeset 49 59936c012add
parent 48 61d9684a557a
child 55 d71424eb5d0c
permissions -rw-r--r--
Illustration added together with renewed explainations for case STAR.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     1
theory Myhill
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
     2
  imports Myhill_1
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     3
begin
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     4
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
     5
section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     6
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
     7
subsection {* The scheme*}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     8
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     9
text {* 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    10
  The following convenient notation @{text "x \<approx>Lang y"} means:
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    11
  string @{text "x"} and @{text "y"} are equivalent with respect to 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    12
  language @{text "Lang"}.
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    13
  *}
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    14
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    15
definition
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
    16
  str_eq :: "string \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> bool" ("_ \<approx>_ _")
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    17
where
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    18
  "x \<approx>Lang y \<equiv> (x, y) \<in> (\<approx>Lang)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    19
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
    20
text {*
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    21
  The basic idea to show the finiteness of the partition induced by relation @{text "\<approx>Lang"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    22
  is to attach a tag @{text "tag(x)"} to every string @{text "x"}, the set of tags are carfully 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    23
  choosen, so that the range of tagging function @{text "tag"} (denoted @{text "range(tag)"}) is finite. 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    24
  If strings with the same tag are equivlent with respect @{text "\<approx>Lang"},
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    25
  i.e. @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"} (this property is named `injectivity' in the following), 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    26
  then it can be proved that: the partition given rise by @{text "(\<approx>Lang)"} is finite. 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    27
  
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    28
  There are two arguments for this. The first goes as the following:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    29
  \begin{enumerate}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    30
    \item First, the tagging function @{text "tag"} induces an equivalent relation @{text "(=tag=)"} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    31
          (defiintion of @{text "f_eq_rel"} and lemma @{text "equiv_f_eq_rel"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    32
    \item It is shown that: if the range of @{text "tag"} is finite, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    33
           the partition given rise by @{text "(=tag=)"} is finite (lemma @{text "finite_eq_f_rel"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    34
    \item It is proved that if equivalent relation @{text "R1"} is more refined than @{text "R2"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    35
           (expressed as @{text "R1 \<subseteq> R2"}),
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    36
           and the partition induced by @{text "R1"} is finite, then the partition induced by @{text "R2"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    37
           is finite as well (lemma @{text "refined_partition_finite"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    38
    \item The injectivity assumption @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"} implies that
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    39
            @{text "(=tag=)"} is more refined than @{text "(\<approx>Lang)"}.
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    40
    \item Combining the points above, we have: the partition induced by language @{text "Lang"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    41
          is finite (lemma @{text "tag_finite_imageD"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    42
  \end{enumerate}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    43
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 35
diff changeset
    44
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    45
definition 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    46
   f_eq_rel ("=_=")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    47
where
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    48
   "(=f=) = {(x, y) | x y. f x = f y}"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    49
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    50
lemma equiv_f_eq_rel:"equiv UNIV (=f=)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    51
  by (auto simp:equiv_def f_eq_rel_def refl_on_def sym_def trans_def)
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    52
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    53
lemma finite_range_image: "finite (range f) \<Longrightarrow> finite (f ` A)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    54
  by (rule_tac B = "{y. \<exists>x. y = f x}" in finite_subset, auto simp:image_def)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    55
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    56
lemma finite_eq_f_rel:
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    57
  assumes rng_fnt: "finite (range tag)"
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    58
  shows "finite (UNIV // (=tag=))"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    59
proof -
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    60
  let "?f" =  "op ` tag" and ?A = "(UNIV // (=tag=))"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    61
  show ?thesis
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    62
  proof (rule_tac f = "?f" and A = ?A in finite_imageD) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    63
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    64
      The finiteness of @{text "f"}-image is a simple consequence of assumption @{text "rng_fnt"}:
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    65
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    66
    show "finite (?f ` ?A)" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    67
    proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    68
      have "\<forall> X. ?f X \<in> (Pow (range tag))" by (auto simp:image_def Pow_def)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    69
      moreover from rng_fnt have "finite (Pow (range tag))" by simp
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    70
      ultimately have "finite (range ?f)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    71
        by (auto simp only:image_def intro:finite_subset)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    72
      from finite_range_image [OF this] show ?thesis .
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    73
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    74
  next
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    75
    -- {* 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    76
      The injectivity of @{text "f"}-image is a consequence of the definition of @{text "(=tag=)"}:
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    77
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    78
    show "inj_on ?f ?A" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    79
    proof-
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    80
      { fix X Y
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    81
        assume X_in: "X \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    82
          and  Y_in: "Y \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    83
          and  tag_eq: "?f X = ?f Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    84
        have "X = Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    85
        proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    86
          from X_in Y_in tag_eq 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    87
          obtain x y 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    88
            where x_in: "x \<in> X" and y_in: "y \<in> Y" and eq_tg: "tag x = tag y"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    89
            unfolding quotient_def Image_def str_eq_rel_def 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    90
                                   str_eq_def image_def f_eq_rel_def
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    91
            apply simp by blast
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    92
          with X_in Y_in show ?thesis 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    93
            by (auto simp:quotient_def str_eq_rel_def str_eq_def f_eq_rel_def) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    94
        qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    95
      } thus ?thesis unfolding inj_on_def by auto
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    96
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    97
  qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    98
qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    99
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   100
lemma finite_image_finite: "\<lbrakk>\<forall> x \<in> A. f x \<in> B; finite B\<rbrakk> \<Longrightarrow> finite (f ` A)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   101
  by (rule finite_subset [of _ B], auto)
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   102
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   103
lemma refined_partition_finite:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   104
  fixes R1 R2 A
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   105
  assumes fnt: "finite (A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   106
  and refined: "R1 \<subseteq> R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   107
  and eq1: "equiv A R1" and eq2: "equiv A R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   108
  shows "finite (A // R2)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   109
proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   110
  let ?f = "\<lambda> X. {R1 `` {x} | x. x \<in> X}" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   111
    and ?A = "(A // R2)" and ?B = "(A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   112
  show ?thesis
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   113
  proof(rule_tac f = ?f and A = ?A in finite_imageD)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   114
    show "finite (?f ` ?A)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   115
    proof(rule finite_subset [of _ "Pow ?B"])
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   116
      from fnt show "finite (Pow (A // R1))" by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   117
    next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   118
      from eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   119
      show " ?f ` A // R2 \<subseteq> Pow ?B"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   120
        apply (unfold image_def Pow_def quotient_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   121
        by (rule_tac x = xb in bexI, simp, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   122
                 unfold equiv_def sym_def refl_on_def, blast)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   123
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   124
  next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   125
    show "inj_on ?f ?A"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   126
    proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   127
      { fix X Y
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   128
        assume X_in: "X \<in> ?A" and Y_in: "Y \<in> ?A" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   129
          and eq_f: "?f X = ?f Y" (is "?L = ?R")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   130
        have "X = Y" using X_in
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   131
        proof(rule quotientE)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   132
          fix x
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   133
          assume "X = R2 `` {x}" and "x \<in> A" with eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   134
          have x_in: "x \<in> X" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   135
            by (unfold equiv_def quotient_def refl_on_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   136
          with eq_f have "R1 `` {x} \<in> ?R" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   137
          then obtain y where 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   138
            y_in: "y \<in> Y" and eq_r: "R1 `` {x} = R1 ``{y}" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   139
          have "(x, y) \<in> R1"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   140
          proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   141
            from x_in X_in y_in Y_in eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   142
            have "x \<in> A" and "y \<in> A" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   143
              by (unfold equiv_def quotient_def refl_on_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   144
            from eq_equiv_class_iff [OF eq1 this] and eq_r
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   145
            show ?thesis by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   146
          qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   147
          with refined have xy_r2: "(x, y) \<in> R2" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   148
          from quotient_eqI [OF eq2 X_in Y_in x_in y_in this]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   149
          show ?thesis .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   150
        qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   151
      } thus ?thesis by (auto simp:inj_on_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   152
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   153
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   154
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   155
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   156
lemma equiv_lang_eq: "equiv UNIV (\<approx>Lang)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   157
  apply (unfold equiv_def str_eq_rel_def sym_def refl_on_def trans_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   158
  by blast
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   159
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   160
lemma tag_finite_imageD:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   161
  fixes tag
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   162
  assumes rng_fnt: "finite (range tag)" 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   163
  -- {* Suppose the rang of tagging fucntion @{text "tag"} is finite. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   164
  and same_tag_eqvt: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<approx>Lang n"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   165
  -- {* And strings with same tag are equivalent *}
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   166
  shows "finite (UNIV // (\<approx>Lang))"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   167
proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   168
  let ?R1 = "(=tag=)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   169
  show ?thesis
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   170
  proof(rule_tac refined_partition_finite [of _ ?R1])
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   171
    from finite_eq_f_rel [OF rng_fnt]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   172
     show "finite (UNIV // =tag=)" . 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   173
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   174
     from same_tag_eqvt
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   175
     show "(=tag=) \<subseteq> (\<approx>Lang)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   176
       by (auto simp:f_eq_rel_def str_eq_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   177
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   178
     from equiv_f_eq_rel
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   179
     show "equiv UNIV (=tag=)" by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   180
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   181
     from equiv_lang_eq
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   182
     show "equiv UNIV (\<approx>Lang)" by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   183
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   184
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   185
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   186
text {*
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   187
  A more concise, but less intelligible argument for @{text "tag_finite_imageD"} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   188
  is given as the following. The basic idea is still using standard library 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   189
  lemma @{thm [source] "finite_imageD"}:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   190
  \[
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   191
  @{thm "finite_imageD" [no_vars]}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   192
  \]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   193
  which says: if the image of injective function @{text "f"} over set @{text "A"} is 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   194
  finite, then @{text "A"} must be finte, as we did in the lemmas above.
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   195
  *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   196
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   197
lemma 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   198
  fixes tag
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   199
  assumes rng_fnt: "finite (range tag)" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   200
  -- {* Suppose the rang of tagging fucntion @{text "tag"} is finite. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   201
  and same_tag_eqvt: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<approx>Lang n"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   202
  -- {* And strings with same tag are equivalent *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   203
  shows "finite (UNIV // (\<approx>Lang))"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   204
  -- {* Then the partition generated by @{text "(\<approx>Lang)"} is finite. *}
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   205
proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   206
  -- {* The particular @{text "f"} and @{text "A"} used in @{thm [source] "finite_imageD"} are:*}
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   207
  let "?f" =  "op ` tag" and ?A = "(UNIV // \<approx>Lang)"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   208
  show ?thesis
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   209
  proof (rule_tac f = "?f" and A = ?A in finite_imageD) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   210
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   211
      The finiteness of @{text "f"}-image is a simple consequence of assumption @{text "rng_fnt"}:
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   212
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   213
    show "finite (?f ` ?A)" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   214
    proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   215
      have "\<forall> X. ?f X \<in> (Pow (range tag))" by (auto simp:image_def Pow_def)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   216
      moreover from rng_fnt have "finite (Pow (range tag))" by simp
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   217
      ultimately have "finite (range ?f)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   218
        by (auto simp only:image_def intro:finite_subset)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   219
      from finite_range_image [OF this] show ?thesis .
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   220
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   221
  next
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   222
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   223
      The injectivity of @{text "f"} is the consequence of assumption @{text "same_tag_eqvt"}:
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   224
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   225
    show "inj_on ?f ?A" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   226
    proof-
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   227
      { fix X Y
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   228
        assume X_in: "X \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   229
          and  Y_in: "Y \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   230
          and  tag_eq: "?f X = ?f Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   231
        have "X = Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   232
        proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   233
          from X_in Y_in tag_eq 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   234
          obtain x y where x_in: "x \<in> X" and y_in: "y \<in> Y" and eq_tg: "tag x = tag y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   235
            unfolding quotient_def Image_def str_eq_rel_def str_eq_def image_def
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   236
            apply simp by blast 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   237
          from same_tag_eqvt [OF eq_tg] have "x \<approx>Lang y" .
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   238
          with X_in Y_in x_in y_in
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   239
          show ?thesis by (auto simp:quotient_def str_eq_rel_def str_eq_def) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   240
        qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   241
      } thus ?thesis unfolding inj_on_def by auto
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   242
    qed
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   243
  qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   244
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   245
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   246
subsection {* The proof*}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   247
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   248
subsubsection {* The case for @{const "NULL"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   249
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   250
lemma quot_null_eq:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   251
  shows "(UNIV // \<approx>{}) = ({UNIV}::lang set)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   252
  unfolding quotient_def Image_def str_eq_rel_def by auto
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   253
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   254
lemma quot_null_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   255
  shows "finite ((UNIV // \<approx>{})::lang set)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   256
unfolding quot_null_eq by simp
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   257
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   258
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   259
subsubsection {* The case for @{const "EMPTY"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   260
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   261
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   262
lemma quot_empty_subset:
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   263
  "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   264
proof
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   265
  fix x
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   266
  assume "x \<in> UNIV // \<approx>{[]}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   267
  then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[]}}" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   268
    unfolding quotient_def Image_def by blast
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   269
  show "x \<in> {{[]}, UNIV - {[]}}"
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   270
  proof (cases "y = []")
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   271
    case True with h
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   272
    have "x = {[]}" by (auto simp: str_eq_rel_def)
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   273
    thus ?thesis by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   274
  next
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   275
    case False with h
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   276
    have "x = UNIV - {[]}" by (auto simp: str_eq_rel_def)
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   277
    thus ?thesis by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   278
  qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   279
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   280
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   281
lemma quot_empty_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   282
  shows "finite (UNIV // (\<approx>{[]}))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   283
by (rule finite_subset[OF quot_empty_subset]) (simp)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   284
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   285
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   286
subsubsection {* The case for @{const "CHAR"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   287
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   288
lemma quot_char_subset:
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   289
  "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   290
proof 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   291
  fix x 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   292
  assume "x \<in> UNIV // \<approx>{[c]}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   293
  then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[c]}}" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   294
    unfolding quotient_def Image_def by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   295
  show "x \<in> {{[]},{[c]}, UNIV - {[], [c]}}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   296
  proof -
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   297
    { assume "y = []" hence "x = {[]}" using h 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   298
        by (auto simp:str_eq_rel_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   299
    } moreover {
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   300
      assume "y = [c]" hence "x = {[c]}" using h 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   301
        by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   302
    } moreover {
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   303
      assume "y \<noteq> []" and "y \<noteq> [c]"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   304
      hence "\<forall> z. (y @ z) \<noteq> [c]" by (case_tac y, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   305
      moreover have "\<And> p. (p \<noteq> [] \<and> p \<noteq> [c]) = (\<forall> q. p @ q \<noteq> [c])" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   306
        by (case_tac p, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   307
      ultimately have "x = UNIV - {[],[c]}" using h
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   308
        by (auto simp add:str_eq_rel_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   309
    } ultimately show ?thesis by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   310
  qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   311
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   312
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   313
lemma quot_char_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   314
  shows "finite (UNIV // (\<approx>{[c]}))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   315
by (rule finite_subset[OF quot_char_subset]) (simp)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   316
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   317
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   318
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   319
subsubsection {* The case for @{text "SEQ"}*}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   320
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   321
definition 
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   322
  tag_str_SEQ :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang set)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   323
where
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   324
  "tag_str_SEQ L1 L2 = 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   325
     (\<lambda>x. (\<approx>L1 `` {x}, {(\<approx>L2 `` {x - xa}) | xa.  xa \<le> x \<and> xa \<in> L1}))"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   326
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   327
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   328
lemma append_seq_elim:
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   329
  assumes "x @ y \<in> L\<^isub>1 ;; L\<^isub>2"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   330
  shows "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2) \<or> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   331
          (\<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   332
proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   333
  from assms obtain s\<^isub>1 s\<^isub>2 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   334
    where "x @ y = s\<^isub>1 @ s\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   335
    and in_seq: "s\<^isub>1 \<in> L\<^isub>1 \<and> s\<^isub>2 \<in> L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   336
    by (auto simp:Seq_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   337
  hence "(x \<le> s\<^isub>1 \<and> (s\<^isub>1 - x) @ s\<^isub>2 = y) \<or> (s\<^isub>1 \<le> x \<and> (x - s\<^isub>1) @ y = s\<^isub>2)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   338
    using app_eq_dest by auto
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   339
  moreover have "\<lbrakk>x \<le> s\<^isub>1; (s\<^isub>1 - x) @ s\<^isub>2 = y\<rbrakk> \<Longrightarrow> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   340
                       \<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   341
    using in_seq by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   342
  moreover have "\<lbrakk>s\<^isub>1 \<le> x; (x - s\<^isub>1) @ y = s\<^isub>2\<rbrakk> \<Longrightarrow> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   343
                    \<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   344
    using in_seq by (rule_tac x = s\<^isub>1 in exI, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   345
  ultimately show ?thesis by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   346
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   347
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   348
lemma tag_str_SEQ_injI:
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   349
  "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \<Longrightarrow> m \<approx>(L\<^isub>1 ;; L\<^isub>2) n"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   350
proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   351
  { fix x y z
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   352
    assume xz_in_seq: "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   353
    and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   354
    have"y @ z \<in> L\<^isub>1 ;; L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   355
    proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   356
      have "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   357
               (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   358
        using xz_in_seq append_seq_elim by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   359
      moreover {
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   360
        fix xa
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   361
        assume h1: "xa \<le> x" and h2: "xa \<in> L\<^isub>1" and h3: "(x - xa) @ z \<in> L\<^isub>2"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   362
        obtain ya where "ya \<le> y" and "ya \<in> L\<^isub>1" and "(y - ya) @ z \<in> L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   363
        proof -
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   364
          have "\<exists> ya.  ya \<le> y \<and> ya \<in> L\<^isub>1 \<and> (x - xa) \<approx>L\<^isub>2 (y - ya)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   365
          proof -
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   366
            have "{\<approx>L\<^isub>2 `` {x - xa} |xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   367
                  {\<approx>L\<^isub>2 `` {y - xa} |xa. xa \<le> y \<and> xa \<in> L\<^isub>1}" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   368
                          (is "?Left = ?Right") 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   369
              using h1 tag_xy by (auto simp:tag_str_SEQ_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   370
            moreover have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Left" using h1 h2 by auto
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   371
            ultimately have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Right" by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   372
            thus ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   373
          qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   374
          with prems show ?thesis by (auto simp:str_eq_rel_def str_eq_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   375
        qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   376
        hence "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" by (erule_tac prefixE, auto simp:Seq_def)          
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   377
      } moreover {
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   378
        fix za
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   379
        assume h1: "za \<le> z" and h2: "(x @ za) \<in> L\<^isub>1" and h3: "z - za \<in> L\<^isub>2"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   380
        hence "y @ za \<in> L\<^isub>1"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   381
        proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   382
          have "\<approx>L\<^isub>1 `` {x} = \<approx>L\<^isub>1 `` {y}" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   383
            using h1 tag_xy by (auto simp:tag_str_SEQ_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   384
          with h2 show ?thesis 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   385
            by (auto simp:Image_def str_eq_rel_def str_eq_def) 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   386
        qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   387
        with h1 h3 have "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   388
          by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   389
      }
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   390
      ultimately show ?thesis by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   391
    qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   392
  } thus "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \<Longrightarrow> m \<approx>(L\<^isub>1 ;; L\<^isub>2) n" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   393
    by (auto simp add: str_eq_def str_eq_rel_def)
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   394
qed 
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   395
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   396
lemma quot_seq_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   397
  fixes L1 L2::"lang"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   398
  assumes fin1: "finite (UNIV // \<approx>L1)" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   399
  and     fin2: "finite (UNIV // \<approx>L2)" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   400
  shows "finite (UNIV // \<approx>(L1 ;; L2))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   401
proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   402
  show "\<And>x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \<Longrightarrow> x \<approx>(L1 ;; L2) y"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   403
    by (rule tag_str_SEQ_injI)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   404
next
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   405
  have *: "finite ((UNIV // \<approx>L1) \<times> (Pow (UNIV // \<approx>L2)))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   406
    using fin1 fin2 by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   407
  show "finite (range (tag_str_SEQ L1 L2))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   408
    unfolding tag_str_SEQ_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   409
    apply(rule finite_subset[OF _ *])
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   410
    unfolding quotient_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   411
    by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   412
qed
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   413
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   414
subsubsection {* The case for @{const ALT} *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   415
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   416
definition 
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   417
  tag_str_ALT :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   418
where
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   419
  "tag_str_ALT L1 L2 = (\<lambda>x. (\<approx>L1 `` {x}, \<approx>L2 `` {x}))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   420
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   421
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   422
lemma quot_union_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   423
  fixes L1 L2::"lang"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   424
  assumes finite1: "finite (UNIV // \<approx>L1)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   425
  and     finite2: "finite (UNIV // \<approx>L2)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   426
  shows "finite (UNIV // \<approx>(L1 \<union> L2))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   427
proof (rule_tac tag = "tag_str_ALT L1 L2" in tag_finite_imageD)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   428
  show "\<And>x y. tag_str_ALT L1 L2 x = tag_str_ALT L1 L2 y \<Longrightarrow> x \<approx>(L1 \<union> L2) y"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   429
    unfolding tag_str_ALT_def 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   430
    unfolding str_eq_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   431
    unfolding Image_def 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   432
    unfolding str_eq_rel_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   433
    by auto
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   434
next
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   435
  have *: "finite ((UNIV // \<approx>L1) \<times> (UNIV // \<approx>L2))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   436
    using finite1 finite2 by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   437
  show "finite (range (tag_str_ALT L1 L2))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   438
    unfolding tag_str_ALT_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   439
    apply(rule finite_subset[OF _ *])
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   440
    unfolding quotient_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   441
    by auto
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   442
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   443
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   444
subsubsection {* The case for @{const "STAR"} *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   445
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   446
text {* 
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   447
  This turned out to be the trickiest case. The essential goal is 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   448
  to proved @{text "y @ z \<in>  L\<^isub>1*"} under the assumptions that @{text "x @ z \<in>  L\<^isub>1*"}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   449
  and that @{text "x"} and @{text "y"} have the same tag. The reasoning goes as the following:
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   450
  \begin{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   451
    \item Since @{text "x @ z \<in>  L\<^isub>1*"} holds, a prefix @{text "xa"} of @{text "x"} can be found
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   452
          such that @{text "xa \<in> L\<^isub>1*"} and @{text "(x - xa)@z \<in> L\<^isub>1*"}, as shown in Fig. \ref{first_split}(a).
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   453
          Such a prefix always exists, @{text "xa = []"}, for example, is one. 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   454
    \item There could be many but fintie many of such @{text "xa"}, from which we can find the longest
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   455
          and name it @{text "xa_max"}, as shown in Fig. \ref{max_split}(b).
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   456
    \item The next step is to split @{text "z"} into @{text "za"} and @{text "zb"} such that
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   457
           @{text "(x - xa_max) @ za \<in> L\<^isub>1"} and @{text "zb \<in> L\<^isub>1*"}  as shown in Fig. \ref{last_split}(d).
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   458
          Such a split always exists because:
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   459
          \begin{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   460
            \item Because @{text "(x - x_max) @ z \<in> L\<^isub>1*"}, it can always be split into prefix @{text "a"}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   461
              and suffix @{text "b"}, such that @{text "a \<in> L\<^isub>1"} and @{text "b \<in> L\<^isub>1*"},
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   462
              as shown in Fig. \ref{ab_split}(c).
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   463
            \item But the prefix @{text "a"} CANNOT be shorter than @{text "x - xa_max"}, otherwise
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   464
                   @{text "xa_max"} is not the max in it's kind.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   465
            \item Now, @{text "za"} is just @{text "a - (x - xa_max)"} and @{text "zb"} is just @{text "b"}.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   466
          \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   467
    \item  \label{tansfer_step} 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   468
         By the assumption that @{text "x"} and @{text "y"} have the same tag, the structure on @{text "x @ z"}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   469
          can be transferred to @{text "y @ z"} as shown in Fig. \ref{trans_split}(e). The detailed steps are:
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   470
          \begin{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   471
            \item A @{text "y"}-prefix @{text "ya"} corresponding to @{text "xa"} can be found, 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   472
                  which satisfies conditions: @{text "ya \<in> L\<^isub>1*"} and @{text "(y - ya)@za \<in> L\<^isub>1"}.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   473
            \item Since we already know @{text "zb \<in> L\<^isub>1*"}, we get @{text "(y - ya)@za@zb \<in> L\<^isub>1*"},
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   474
                  and this is just @{text "(y - ya)@z \<in> L\<^isub>1*"}.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   475
            \item With fact @{text "ya \<in> L\<^isub>1*"}, we finally get @{text "y@z \<in> L\<^isub>1*"}.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   476
          \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   477
  \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   478
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   479
  The formal proof of lemma @{text "tag_str_STAR_injI"} faithfully follows this informal argument 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   480
  while the tagging function @{text "tag_str_STAR"} is defined to make the transfer in step
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   481
  \ref{transfer_step}4 feasible.
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   482
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   483
   
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   484
\begin{figure}[h!]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   485
\centering
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   486
\subfigure[First split]{\label{first_split}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   487
\scalebox{0.9}{
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   488
\begin{tikzpicture}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   489
    \node[draw,minimum height=3.8ex] (xa) {$\hspace{2em}xa\hspace{2em}$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   490
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{5em}x - xa\hspace{5em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   491
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (z) { $\hspace{21em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   492
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   493
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   494
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   495
               node[midway, above=0.5em]{$x$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   496
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   497
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   498
           (z.north west) -- ($(z.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   499
               node[midway, above=0.5em]{$z$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   500
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   501
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   502
           ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   503
               node[midway, above=0.8em]{$x @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   504
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   505
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   506
           ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   507
               node[midway, below=0.5em]{$(x - xa) @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   508
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   509
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   510
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   511
               node[midway, below=0.5em]{$xa \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   512
\end{tikzpicture}}}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   513
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   514
\subfigure[Max split]{\label{max_split}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   515
\scalebox{0.9}{
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   516
\begin{tikzpicture}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   517
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{4em}xa\_max\hspace{4em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   518
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{0.5em}x - xa\_max\hspace{0.5em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   519
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (z) { $\hspace{21em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   520
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   521
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   522
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   523
               node[midway, above=0.5em]{$x$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   524
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   525
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   526
           (z.north west) -- ($(z.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   527
               node[midway, above=0.5em]{$z$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   528
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   529
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   530
           ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   531
               node[midway, above=0.8em]{$x @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   532
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   533
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   534
           ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   535
               node[midway, below=0.5em]{$(x - xa\_max) @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   536
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   537
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   538
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   539
               node[midway, below=0.5em]{$xa \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   540
\end{tikzpicture}}}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   541
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   542
\subfigure[Max split with $a$ and $b$]{\label{ab_split}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   543
\scalebox{0.9}{
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   544
\begin{tikzpicture}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   545
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{4em}xa\_max\hspace{4em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   546
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{0.5em}x - xa\_max\hspace{0.5em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   547
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (z) { $\hspace{21em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   548
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   549
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   550
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   551
               node[midway, above=0.5em]{$x$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   552
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   553
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   554
           (z.north west) -- ($(z.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   555
               node[midway, above=0.5em]{$z$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   556
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   557
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   558
           ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   559
               node[midway, above=0.8em]{$x @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   560
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   561
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   562
           ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   563
               node[midway, below=0.5em]{$(x - xa\_max) @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   564
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   565
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   566
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   567
               node[midway, below=0.5em]{$xa \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   568
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   569
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   570
           ($(xxa.south east)+(6em,-5ex)$) -- ($(xxa.south west)+(0em,-5ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   571
               node[midway, below=0.5em]{$a \in L_1$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   572
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   573
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   574
           ($(z.south east)+(0em,-5ex)$) -- ($(xxa.south east)+(6em,-5ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   575
               node[midway, below=0.5em]{$b \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   576
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   577
\end{tikzpicture}}}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   578
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   579
\subfigure[Last split]{\label{last_split}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   580
\scalebox{0.9}{
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   581
\begin{tikzpicture}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   582
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{4em}xa\_max\hspace{4em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   583
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{0.5em}x - xa\_max\hspace{0.5em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   584
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (za) { $\hspace{2em}za\hspace{2em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   585
    \node[draw,minimum height=3.8ex, right=-0.03em of za] (zb) { $\hspace{7em}zb\hspace{7em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   586
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   587
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   588
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   589
               node[midway, above=0.5em]{$x$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   590
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   591
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   592
           (za.north west) -- ($(zb.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   593
               node[midway, above=0.5em]{$z$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   594
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   595
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   596
           ($(xa.north west)+(0em,3ex)$) -- ($(zb.north east)+(0em,3ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   597
               node[midway, above=0.8em]{$x @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   598
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   599
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   600
           ($(za.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   601
               node[midway, below=0.5em]{$(x - xa\_max) @ za \in L_1$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   602
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   603
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   604
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   605
               node[midway, below=0.5em]{$xa\_max \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   606
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   607
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   608
           ($(zb.south east)+(0em,0ex)$) -- ($(zb.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   609
               node[midway, below=0.5em]{$zb \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   610
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   611
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   612
           ($(zb.south east)+(0em,-4ex)$) -- ($(xxa.south west)+(0em,-4ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   613
               node[midway, below=0.5em]{$(x - xa\_max)@z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   614
\end{tikzpicture}}}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   615
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   616
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   617
\subfigure[Transferring to $y$]{\label{trans_split}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   618
\scalebox{0.9}{
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   619
\begin{tikzpicture}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   620
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{5em}ya\hspace{5em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   621
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{2em}y - ya\hspace{2em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   622
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (za) { $\hspace{2em}za\hspace{2em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   623
    \node[draw,minimum height=3.8ex, right=-0.03em of za] (zb) { $\hspace{7em}zb\hspace{7em}$ };
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   624
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   625
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   626
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   627
               node[midway, above=0.5em]{$y$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   628
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   629
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   630
           (za.north west) -- ($(zb.north east)+(0em,0em)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   631
               node[midway, above=0.5em]{$z$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   632
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   633
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   634
           ($(xa.north west)+(0em,3ex)$) -- ($(zb.north east)+(0em,3ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   635
               node[midway, above=0.8em]{$y @ z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   636
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   637
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   638
           ($(za.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   639
               node[midway, below=0.5em]{$(y - ya) @ za \in L_1$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   640
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   641
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   642
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   643
               node[midway, below=0.5em]{$ya \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   644
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   645
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   646
           ($(zb.south east)+(0em,0ex)$) -- ($(zb.south west)+(0em,0ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   647
               node[midway, below=0.5em]{$zb \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   648
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   649
    \draw[decoration={brace,transform={yscale=3}},decorate]
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   650
           ($(zb.south east)+(0em,-4ex)$) -- ($(xxa.south west)+(0em,-4ex)$)
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   651
               node[midway, below=0.5em]{$(y - ya)@z \in L_1*$};
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   652
\end{tikzpicture}}}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   653
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   654
\caption{The case for $STAR$}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   655
\end{figure}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   656
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   657
*} 
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   658
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   659
definition 
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   660
  tag_str_STAR :: "lang \<Rightarrow> string \<Rightarrow> lang set"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   661
where
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   662
  "tag_str_STAR L1 = (\<lambda>x. {\<approx>L1 `` {x - xa} | xa. xa < x \<and> xa \<in> L1\<star>})"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   663
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   664
text {* A technical lemma. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   665
lemma finite_set_has_max: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   666
           (\<exists> max \<in> A. \<forall> a \<in> A. f a <= (f max :: nat))"
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   667
proof (induct rule:finite.induct)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   668
  case emptyI thus ?case by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   669
next
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   670
  case (insertI A a)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   671
  show ?case
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   672
  proof (cases "A = {}")
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   673
    case True thus ?thesis by (rule_tac x = a in bexI, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   674
  next
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   675
    case False
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   676
    with prems obtain max 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   677
      where h1: "max \<in> A" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   678
      and h2: "\<forall>a\<in>A. f a \<le> f max" by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   679
    show ?thesis
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   680
    proof (cases "f a \<le> f max")
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   681
      assume "f a \<le> f max"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   682
      with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   683
    next
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   684
      assume "\<not> (f a \<le> f max)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   685
      thus ?thesis using h2 by (rule_tac x = a in bexI, auto)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   686
    qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   687
  qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   688
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   689
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   690
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   691
text {* Technical lemma. *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   692
lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   693
apply (induct x rule:rev_induct, simp)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   694
apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \<union> {xs}")
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   695
by (auto simp:strict_prefix_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   696
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   697
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   698
lemma tag_str_STAR_injI:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   699
  fixes v w
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   700
  assumes eq_tag: "tag_str_STAR L\<^isub>1 v = tag_str_STAR L\<^isub>1 w"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   701
  shows "(v::string) \<approx>(L\<^isub>1\<star>) w"
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   702
proof-
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   703
    -- {* 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   704
    \begin{minipage}{0.9\textwidth}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   705
    According to the definition of @{text "\<approx>Lang"}, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   706
    proving @{text "v \<approx>(L\<^isub>1\<star>) w"} amounts to
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   707
    showing: for any string @{text "u"},
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   708
    if @{text "v @ u \<in> (L\<^isub>1\<star>)"} then @{text "w @ u \<in> (L\<^isub>1\<star>)"} and vice versa.
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   709
    The reasoning pattern for both directions are the same, as derived
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   710
    in the following:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   711
    \end{minipage}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   712
    *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   713
  { fix x y z
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   714
    assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   715
      and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   716
    have "y @ z \<in> L\<^isub>1\<star>"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   717
    proof(cases "x = []")
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   718
      -- {* 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   719
        The degenerated case when @{text "x"} is a null string is easy to prove:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   720
        *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   721
      case True
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   722
      with tag_xy have "y = []" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   723
        by (auto simp:tag_str_STAR_def strict_prefix_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   724
      thus ?thesis using xz_in_star True by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   725
    next
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   726
        -- {*
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   727
        \begin{minipage}{0.9\textwidth}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   728
        The case when @{text "x"} is not null, and
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   729
        @{text "x @ z"} is in @{text "L\<^isub>1\<star>"}, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   730
        \end{minipage}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   731
        *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   732
      case False
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   733
      -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   734
        \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   735
        Since @{text "x @ z \<in> L\<^isub>1\<star>"}, @{text "x"} can always be splited
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   736
        by a prefix @{text "xa"} together with its suffix @{text "x - xa"}, such
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   737
        that both @{text "xa"} and @{text "(x - xa) @ z"} are in @{text "L\<^isub>1\<star>"},
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   738
        and there could be many such splittings.Therefore, the following set @{text "?S"} 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   739
        is nonempty, and finite as well:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   740
        \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   741
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   742
      let ?S = "{xa. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   743
      have "finite ?S"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   744
        by (rule_tac B = "{xa. xa < x}" in finite_subset, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   745
          auto simp:finite_strict_prefix_set)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   746
      moreover have "?S \<noteq> {}" using False xz_in_star
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   747
        by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   748
      -- {* Since @{text "?S"} is finite, we can always single out the longest
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   749
          and name it @{text "xa_max"}:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   750
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   751
      ultimately have "\<exists> xa_max \<in> ?S. \<forall> xa \<in> ?S. length xa \<le> length xa_max" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   752
        using finite_set_has_max by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   753
      then obtain xa_max 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   754
        where h1: "xa_max < x" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   755
        and h2: "xa_max \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   756
        and h3: "(x - xa_max) @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   757
        and h4:"\<forall> xa < x. xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>  
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   758
                                     \<longrightarrow> length xa \<le> length xa_max"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   759
        by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   760
      -- {*
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   761
          \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   762
          By the equality of tags, the counterpart of @{text "xa_max"} among 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   763
          @{text "y"}-prefixes, named @{text "ya"}, can be found:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   764
          \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   765
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   766
      obtain ya 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   767
        where h5: "ya < y" and h6: "ya \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   768
        and eq_xya: "(x - xa_max) \<approx>L\<^isub>1 (y - ya)"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   769
      proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   770
        from tag_xy have "{\<approx>L\<^isub>1 `` {x - xa} |xa. xa < x \<and> xa \<in> L\<^isub>1\<star>} = 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   771
          {\<approx>L\<^isub>1 `` {y - xa} |xa. xa < y \<and> xa \<in> L\<^isub>1\<star>}" (is "?left = ?right")
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   772
          by (auto simp:tag_str_STAR_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   773
        moreover have "\<approx>L\<^isub>1 `` {x - xa_max} \<in> ?left" using h1 h2 by auto
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   774
        ultimately have "\<approx>L\<^isub>1 `` {x - xa_max} \<in> ?right" by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   775
        with prems show ?thesis apply 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   776
          (simp add:Image_def str_eq_rel_def str_eq_def) by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   777
      qed 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   778
      -- {*
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   779
          \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   780
          If the following proposition can be proved, then the @{text "?thesis"}:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   781
          @{text "y @ z \<in> L\<^isub>1\<star>"} is just a simple consequence.
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   782
          \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   783
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   784
      have "(y - ya) @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   785
      proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   786
        -- {* The idea is to split the suffix @{text "z"} into @{text "za"} and @{text "zb"}, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   787
          such that: *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   788
        obtain za zb where eq_zab: "z = za @ zb" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   789
          and l_za: "(y - ya)@za \<in> L\<^isub>1" and ls_zb: "zb \<in> L\<^isub>1\<star>"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   790
        proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   791
          -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   792
            \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   793
            Since @{text  "(x - xa_max) @ z"} is in @{text "L\<^isub>1\<star>"}, it can be split into
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   794
            @{text "a"} and @{text "b"} such that:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   795
            \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   796
            *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   797
          from h1 have "(x - xa_max) @ z \<noteq> []" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   798
            by (auto simp:strict_prefix_def elim:prefixE)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   799
          from star_decom [OF h3 this]
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   800
          obtain a b where a_in: "a \<in> L\<^isub>1" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   801
            and a_neq: "a \<noteq> []" and b_in: "b \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   802
            and ab_max: "(x - xa_max) @ z = a @ b" by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   803
          -- {* Now the candiates for @{text "za"} and @{text "zb"} are found:*}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   804
          let ?za = "a - (x - xa_max)" and ?zb = "b"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   805
          have pfx: "(x - xa_max) \<le> a" (is "?P1") 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   806
            and eq_z: "z = ?za @ ?zb" (is "?P2")
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   807
          proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   808
            -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   809
              \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   810
              Since @{text "(x - xa_max) @ z = a @ b"}, the string @{text "(x - xa_max) @ z"}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   811
              could be splited in two ways:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   812
              \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   813
              *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   814
            have "((x - xa_max) \<le> a \<and> (a - (x - xa_max)) @ b = z) \<or> 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   815
              (a < (x - xa_max) \<and> ((x - xa_max) - a) @ z = b)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   816
              using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   817
            moreover {
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   818
              -- {* However, the undsired way can be refuted by absurdity: *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   819
              assume np: "a < (x - xa_max)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   820
                and b_eqs: "((x - xa_max) - a) @ z = b"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   821
              have "False"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   822
              proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   823
                let ?xa_max' = "xa_max @ a"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   824
                have "?xa_max' < x" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   825
                  using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   826
                moreover have "?xa_max' \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   827
                  using a_in h2 by (simp add:star_intro3) 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   828
                moreover have "(x - ?xa_max') @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   829
                  using b_eqs b_in np h1 by (simp add:diff_diff_appd)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   830
                moreover have "\<not> (length ?xa_max' \<le> length xa_max)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   831
                  using a_neq by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   832
                ultimately show ?thesis using h4 by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   833
              qed }
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   834
            -- {* Now it can be shown that the splitting goes the way we desired. *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   835
            ultimately show ?P1 and ?P2 by auto
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   836
          qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   837
          hence "(x - xa_max)@?za \<in> L\<^isub>1" using a_in by (auto elim:prefixE)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   838
          -- {* Now candidates @{text "?za"} and @{text "?zb"} have all the requred properteis. *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   839
          with eq_xya have "(y - ya) @ ?za \<in> L\<^isub>1" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   840
            by (auto simp:str_eq_def str_eq_rel_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   841
           with eq_z and b_in prems
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   842
          show ?thesis by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   843
        qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   844
        -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   845
            \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   846
            From the properties of @{text "za"} and @{text "zb"} such obtained,
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   847
              @{text "?thesis"} can be shown easily. 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   848
            \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   849
            *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   850
        from step [OF l_za ls_zb]
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   851
        have "((y - ya) @ za) @ zb \<in> L\<^isub>1\<star>" .
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   852
        with eq_zab show ?thesis by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   853
      qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   854
      with h5 h6 show ?thesis 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   855
        by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   856
    qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   857
  } 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   858
  -- {* By instantiating the reasoning pattern just derived for both directions:*} 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   859
  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   860
  -- {* The thesis is proved as a trival consequence: *} 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   861
    show  ?thesis by (unfold str_eq_def str_eq_rel_def, blast)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   862
qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   863
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   864
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   865
lemma -- {* The oringal version with a poor readability*}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   866
  fixes v w
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   867
  assumes eq_tag: "tag_str_STAR L\<^isub>1 v = tag_str_STAR L\<^isub>1 w"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   868
  shows "(v::string) \<approx>(L\<^isub>1\<star>) w"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   869
proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   870
    -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   871
    \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   872
    According to the definition of @{text "\<approx>Lang"}, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   873
    proving @{text "v \<approx>(L\<^isub>1\<star>) w"} amounts to
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   874
    showing: for any string @{text "u"},
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   875
    if @{text "v @ u \<in> (L\<^isub>1\<star>)"} then @{text "w @ u \<in> (L\<^isub>1\<star>)"} and vice versa.
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   876
    The reasoning pattern for both directions are the same, as derived
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   877
    in the following:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   878
    \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   879
    *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   880
  { fix x y z
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   881
    assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   882
      and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   883
    have "y @ z \<in> L\<^isub>1\<star>"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   884
    proof(cases "x = []")
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   885
      -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   886
        The degenerated case when @{text "x"} is a null string is easy to prove:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   887
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   888
      case True
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   889
      with tag_xy have "y = []" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   890
        by (auto simp:tag_str_STAR_def strict_prefix_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   891
      thus ?thesis using xz_in_star True by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   892
    next
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   893
        -- {*
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   894
        \begin{minipage}{0.9\textwidth}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   895
        The case when @{text "x"} is not null, and
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   896
        @{text "x @ z"} is in @{text "L\<^isub>1\<star>"}, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   897
        \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   898
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   899
      case False
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   900
      obtain x_max 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   901
        where h1: "x_max < x" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   902
        and h2: "x_max \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   903
        and h3: "(x - x_max) @ z \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   904
        and h4:"\<forall> xa < x. xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   905
                                     \<longrightarrow> length xa \<le> length x_max"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   906
      proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   907
        let ?S = "{xa. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   908
        have "finite ?S"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   909
          by (rule_tac B = "{xa. xa < x}" in finite_subset, 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   910
                                auto simp:finite_strict_prefix_set)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   911
        moreover have "?S \<noteq> {}" using False xz_in_star
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   912
          by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   913
        ultimately have "\<exists> max \<in> ?S. \<forall> a \<in> ?S. length a \<le> length max" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   914
          using finite_set_has_max by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   915
        with prems show ?thesis by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   916
      qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   917
      obtain ya 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   918
        where h5: "ya < y" and h6: "ya \<in> L\<^isub>1\<star>" and h7: "(x - x_max) \<approx>L\<^isub>1 (y - ya)"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   919
      proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   920
        from tag_xy have "{\<approx>L\<^isub>1 `` {x - xa} |xa. xa < x \<and> xa \<in> L\<^isub>1\<star>} = 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   921
          {\<approx>L\<^isub>1 `` {y - xa} |xa. xa < y \<and> xa \<in> L\<^isub>1\<star>}" (is "?left = ?right")
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   922
          by (auto simp:tag_str_STAR_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   923
        moreover have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?left" using h1 h2 by auto
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   924
        ultimately have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?right" by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   925
        with prems show ?thesis apply 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   926
          (simp add:Image_def str_eq_rel_def str_eq_def) by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   927
      qed      
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   928
      have "(y - ya) @ z \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   929
      proof-
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   930
        from h3 h1 obtain a b where a_in: "a \<in> L\<^isub>1" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   931
          and a_neq: "a \<noteq> []" and b_in: "b \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   932
          and ab_max: "(x - x_max) @ z = a @ b" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   933
          by (drule_tac star_decom, auto simp:strict_prefix_def elim:prefixE)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   934
        have "(x - x_max) \<le> a \<and> (a - (x - x_max)) @ b = z" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   935
        proof -
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   936
          have "((x - x_max) \<le> a \<and> (a - (x - x_max)) @ b = z) \<or> 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   937
                            (a < (x - x_max) \<and> ((x - x_max) - a) @ z = b)" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   938
            using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   939
          moreover { 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   940
            assume np: "a < (x - x_max)" and b_eqs: " ((x - x_max) - a) @ z = b"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   941
            have "False"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   942
            proof -
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   943
              let ?x_max' = "x_max @ a"
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   944
              have "?x_max' < x" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   945
                using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   946
              moreover have "?x_max' \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   947
                using a_in h2 by (simp add:star_intro3) 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   948
              moreover have "(x - ?x_max') @ z \<in> L\<^isub>1\<star>" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   949
                using b_eqs b_in np h1 by (simp add:diff_diff_appd)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   950
              moreover have "\<not> (length ?x_max' \<le> length x_max)" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   951
                using a_neq by simp
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   952
              ultimately show ?thesis using h4 by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   953
            qed 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   954
          } ultimately show ?thesis by blast
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   955
        qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   956
        then obtain za where z_decom: "z = za @ b" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   957
          and x_za: "(x - x_max) @ za \<in> L\<^isub>1" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   958
          using a_in by (auto elim:prefixE)        
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   959
        from x_za h7 have "(y - ya) @ za \<in> L\<^isub>1" 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   960
          by (auto simp:str_eq_def str_eq_rel_def)
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   961
        with z_decom b_in show ?thesis by (auto dest!:step[of "(y - ya) @ za"])
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   962
      qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   963
      with h5 h6 show ?thesis 
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   964
        by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   965
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   966
  } 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   967
  -- {* By instantiating the reasoning pattern just derived for both directions:*} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   968
  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   969
  -- {* The thesis is proved as a trival consequence: *} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   970
    show  ?thesis by (unfold str_eq_def str_eq_rel_def, blast)
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   971
qed
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   972
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   973
lemma quot_star_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   974
  fixes L1::"lang"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   975
  assumes finite1: "finite (UNIV // \<approx>L1)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   976
  shows "finite (UNIV // \<approx>(L1\<star>))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   977
proof (rule_tac tag = "tag_str_STAR L1" in tag_finite_imageD)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   978
  show "\<And>x y. tag_str_STAR L1 x = tag_str_STAR L1 y \<Longrightarrow> x \<approx>(L1\<star>) y"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   979
    by (rule tag_str_STAR_injI)
40
50d00d7dc413 tuned a bit more the last STAR-proof
wu
parents: 39
diff changeset
   980
next
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   981
  have *: "finite (Pow (UNIV // \<approx>L1))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   982
    using finite1 by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   983
  show "finite (range (tag_str_STAR L1))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   984
    unfolding tag_str_STAR_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   985
    apply(rule finite_subset[OF _ *])
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   986
    unfolding quotient_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   987
    by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   988
qed
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   989
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   990
subsubsection{* The conclusion *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   991
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   992
lemma rexp_imp_finite:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   993
  fixes r::"rexp"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   994
  shows "finite (UNIV // \<approx>(L r))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   995
by (induct r) (auto)
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   996
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   997
end
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   998
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   999
(*
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1000
lemma refined_quotient_union_eq:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1001
  assumes refined: "R1 \<subseteq> R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1002
  and eq1: "equiv A R1" and eq2: "equiv A R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1003
  and y_in: "y \<in> A"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1004
  shows "\<Union>{R1 `` {x} | x. x \<in> (R2 `` {y})} = R2 `` {y}"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1005
proof
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1006
  show "\<Union>{R1 `` {x} |x. x \<in> R2 `` {y}} \<subseteq> R2 `` {y}" (is "?L \<subseteq> ?R")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1007
  proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1008
    { fix z
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1009
      assume zl: "z \<in> ?L" and nzr: "z \<notin> ?R"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1010
      have "False"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1011
      proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1012
        from zl and eq1 eq2 and y_in 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1013
        obtain x where xy2: "(x, y) \<in> R2" and zx1: "(z, x) \<in> R1"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1014
          by (simp only:equiv_def sym_def, blast)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1015
        have "(z, y) \<in> R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1016
        proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1017
          from zx1 and refined have "(z, x) \<in> R2" by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1018
          moreover from xy2 have "(x, y) \<in> R2" .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1019
          ultimately show ?thesis using eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1020
            by (simp only:equiv_def, unfold trans_def, blast)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1021
        qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1022
        with nzr eq2 show ?thesis by (auto simp:equiv_def sym_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1023
      qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1024
    } thus ?thesis by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1025
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1026
next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1027
  show "R2 `` {y} \<subseteq> \<Union>{R1 `` {x} |x. x \<in> R2 `` {y}}" (is "?L \<subseteq> ?R")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1028
  proof
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1029
    fix x
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1030
    assume x_in: "x \<in> ?L"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1031
    with eq1 eq2 have "x \<in> R1 `` {x}" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1032
      by (unfold equiv_def refl_on_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1033
    with x_in show "x \<in> ?R" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1034
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1035
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1036
*)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1037
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1038
(*
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1039
lemma refined_partition_finite:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1040
  fixes R1 R2 A
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1041
  assumes fnt: "finite (A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1042
  and refined: "R1 \<subseteq> R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1043
  and eq1: "equiv A R1" and eq2: "equiv A R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1044
  shows "finite (A // R2)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1045
proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1046
  let ?f = "\<lambda> X. {R1 `` {x} | x. x \<in> X}" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1047
    and ?A = "(A // R2)" and ?B = "(A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1048
  show ?thesis
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1049
  proof(rule_tac f = ?f and A = ?A in finite_imageD)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1050
    show "finite (?f ` ?A)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1051
    proof(rule finite_subset [of _ "Pow ?B"])
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1052
      from fnt show "finite (Pow (A // R1))" by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1053
    next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1054
      from eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1055
      show " ?f ` A // R2 \<subseteq> Pow ?B"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1056
        apply (unfold image_def Pow_def quotient_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1057
        by (rule_tac x = xb in bexI, simp, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1058
                 unfold equiv_def sym_def refl_on_def, blast)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1059
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1060
  next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1061
    show "inj_on ?f ?A"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1062
    proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1063
      { fix X Y
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1064
        assume X_in: "X \<in> ?A" and Y_in: "Y \<in> ?A" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1065
          and eq_f: "?f X = ?f Y" (is "?L = ?R")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1066
        hence "X = Y"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1067
        proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1068
          from X_in eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1069
          obtain x 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1070
            where x_in: "x \<in> A" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1071
            and eq_x: "X = R2 `` {x}" (is "X = ?X")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1072
            by (unfold quotient_def equiv_def refl_on_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1073
          from Y_in eq2 obtain y 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1074
            where y_in: "y \<in> A"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1075
            and eq_y: "Y = R2 `` {y}" (is "Y = ?Y")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1076
            by (unfold quotient_def equiv_def refl_on_def, auto)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1077
          have "?X = ?Y"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1078
          proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1079
            from eq_f have "\<Union> ?L = \<Union> ?R" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1080
            moreover have "\<Union> ?L = ?X"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1081
            proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1082
              from eq_x have "\<Union> ?L =  \<Union>{R1 `` {x} |x. x \<in> ?X}" by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1083
              also from refined_quotient_union_eq [OF refined eq1 eq2 x_in]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1084
              have "\<dots> = ?X" .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1085
              finally show ?thesis .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1086
            qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1087
            moreover have "\<Union> ?R = ?Y"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1088
            proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1089
              from eq_y have "\<Union> ?R =  \<Union>{R1 `` {y} |y. y \<in> ?Y}" by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1090
              also from refined_quotient_union_eq [OF refined eq1 eq2 y_in]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1091
              have "\<dots> = ?Y" .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1092
              finally show ?thesis .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1093
            qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1094
            ultimately show ?thesis by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1095
          qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1096
          with eq_x eq_y show ?thesis by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1097
        qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1098
      } thus ?thesis by (auto simp:inj_on_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1099
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1100
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1101
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
  1102
*)