Myhill.thy
author zhang
Wed, 02 Feb 2011 13:25:09 +0000
changeset 57 76ab7c09d575
parent 56 b3898315e687
child 62 d94209ad2880
permissions -rw-r--r--
Myhill.pdf modified
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 {*
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    21
  The main lemma (@{text "rexp_imp_finite"}) is proved by a structural induction over regular expressions.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    22
  While base cases (cases for @{const "NULL"}, @{const "EMPTY"}, @{const "CHAR"}) are quite straight forward,
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    23
  the inductive cases are rather involved. What we have when starting to prove these inductive caes is that
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    24
  the partitions induced by the componet language are finite. The basic idea to show the finiteness of the 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    25
  partition induced by the composite language is to attach a tag @{text "tag(x)"} to every string 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    26
  @{text "x"}. The tags are made of equivalent classes from the component partitions. Let @{text "tag"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    27
  be the tagging function and @{text "Lang"} be the composite language, it can be proved that
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    28
  if strings with the same tag are equivalent with respect to @{text "Lang"}, expressed as:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    29
  \[
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    30
  @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    31
  \]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    32
  then the partition induced by @{text "Lang"} must be finite. There are two arguments for this. 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    33
  The first goes as the following:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    34
  \begin{enumerate}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    35
    \item First, the tagging function @{text "tag"} induces an equivalent relation @{text "(=tag=)"} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    36
          (defiintion of @{text "f_eq_rel"} and lemma @{text "equiv_f_eq_rel"}).
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    37
    \item It is shown that: if the range of @{text "tag"} (denoted @{text "range(tag)"}) is finite, 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    38
           the partition given rise by @{text "(=tag=)"} is finite (lemma @{text "finite_eq_f_rel"}).
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    39
           Since tags are made from equivalent classes from component partitions, and the inductive
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    40
           hypothesis ensures the finiteness of these partitions, it is not difficult to prove
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
    41
           the finiteness of @{text "range(tag)"}.
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    42
    \item It is proved that if equivalent relation @{text "R1"} is more refined than @{text "R2"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    43
           (expressed as @{text "R1 \<subseteq> R2"}),
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    44
           and the partition induced by @{text "R1"} is finite, then the partition induced by @{text "R2"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    45
           is finite as well (lemma @{text "refined_partition_finite"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    46
    \item The injectivity assumption @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"} implies that
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    47
            @{text "(=tag=)"} is more refined than @{text "(\<approx>Lang)"}.
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    48
    \item Combining the points above, we have: the partition induced by language @{text "Lang"}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    49
          is finite (lemma @{text "tag_finite_imageD"}).
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    50
  \end{enumerate}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    51
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 35
diff changeset
    52
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    53
definition 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    54
   f_eq_rel ("=_=")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    55
where
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    56
   "(=f=) = {(x, y) | x y. f x = f y}"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    57
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    58
lemma equiv_f_eq_rel:"equiv UNIV (=f=)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    59
  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
    60
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    61
lemma finite_range_image: "finite (range f) \<Longrightarrow> finite (f ` A)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    62
  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
    63
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    64
lemma finite_eq_f_rel:
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    65
  assumes rng_fnt: "finite (range tag)"
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    66
  shows "finite (UNIV // (=tag=))"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    67
proof -
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    68
  let "?f" =  "op ` tag" and ?A = "(UNIV // (=tag=))"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    69
  show ?thesis
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    70
  proof (rule_tac f = "?f" and A = ?A in finite_imageD) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    71
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    72
      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
    73
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    74
    show "finite (?f ` ?A)" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    75
    proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    76
      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
    77
      moreover from rng_fnt have "finite (Pow (range tag))" by simp
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    78
      ultimately have "finite (range ?f)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    79
        by (auto simp only:image_def intro:finite_subset)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    80
      from finite_range_image [OF this] show ?thesis .
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    81
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    82
  next
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    83
    -- {* 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    84
      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
    85
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    86
    show "inj_on ?f ?A" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    87
    proof-
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    88
      { fix X Y
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    89
        assume X_in: "X \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    90
          and  Y_in: "Y \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    91
          and  tag_eq: "?f X = ?f Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    92
        have "X = Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    93
        proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    94
          from X_in Y_in tag_eq 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    95
          obtain x y 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    96
            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
    97
            unfolding quotient_def Image_def str_eq_rel_def 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
    98
                                   str_eq_def image_def f_eq_rel_def
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
    99
            apply simp by blast
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   100
          with X_in Y_in show ?thesis 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   101
            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
   102
        qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   103
      } thus ?thesis unfolding inj_on_def by auto
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   104
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   105
  qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   106
qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   107
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   108
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
   109
  by (rule finite_subset [of _ B], auto)
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   110
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   111
lemma refined_partition_finite:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   112
  fixes R1 R2 A
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   113
  assumes fnt: "finite (A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   114
  and refined: "R1 \<subseteq> R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   115
  and eq1: "equiv A R1" and eq2: "equiv A R2"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   116
  shows "finite (A // R2)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   117
proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   118
  let ?f = "\<lambda> X. {R1 `` {x} | x. x \<in> X}" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   119
    and ?A = "(A // R2)" and ?B = "(A // R1)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   120
  show ?thesis
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   121
  proof(rule_tac f = ?f and A = ?A in finite_imageD)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   122
    show "finite (?f ` ?A)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   123
    proof(rule finite_subset [of _ "Pow ?B"])
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   124
      from fnt show "finite (Pow (A // R1))" by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   125
    next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   126
      from eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   127
      show " ?f ` A // R2 \<subseteq> Pow ?B"
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   128
        unfolding image_def Pow_def quotient_def
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   129
        apply auto
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   130
        by (rule_tac x = xb in bexI, simp, 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   131
                 unfold equiv_def sym_def refl_on_def, blast)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   132
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   133
  next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   134
    show "inj_on ?f ?A"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   135
    proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   136
      { fix X Y
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   137
        assume X_in: "X \<in> ?A" and Y_in: "Y \<in> ?A" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   138
          and eq_f: "?f X = ?f Y" (is "?L = ?R")
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   139
        have "X = Y" using X_in
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   140
        proof(rule quotientE)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   141
          fix x
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   142
          assume "X = R2 `` {x}" and "x \<in> A" with eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   143
          have x_in: "x \<in> X" 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   144
            unfolding equiv_def quotient_def refl_on_def by auto
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   145
          with eq_f have "R1 `` {x} \<in> ?R" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   146
          then obtain y where 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   147
            y_in: "y \<in> Y" and eq_r: "R1 `` {x} = R1 ``{y}" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   148
          have "(x, y) \<in> R1"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   149
          proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   150
            from x_in X_in y_in Y_in eq2
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   151
            have "x \<in> A" and "y \<in> A" 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   152
              unfolding equiv_def quotient_def refl_on_def by auto
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   153
            from eq_equiv_class_iff [OF eq1 this] and eq_r
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   154
            show ?thesis by simp
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   155
          qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   156
          with refined have xy_r2: "(x, y) \<in> R2" by auto
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   157
          from quotient_eqI [OF eq2 X_in Y_in x_in y_in this]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   158
          show ?thesis .
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   159
        qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   160
      } thus ?thesis by (auto simp:inj_on_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   161
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   162
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   163
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   164
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   165
lemma equiv_lang_eq: "equiv UNIV (\<approx>Lang)"
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   166
  unfolding equiv_def str_eq_rel_def sym_def refl_on_def trans_def
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   167
  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
   168
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
   169
lemma tag_finite_imageD:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   170
  fixes tag
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   171
  assumes rng_fnt: "finite (range tag)" 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   172
  -- {* Suppose the rang of tagging fucntion @{text "tag"} is finite. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   173
  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
   174
  -- {* And strings with same tag are equivalent *}
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   175
  shows "finite (UNIV // (\<approx>Lang))"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   176
proof -
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   177
  let ?R1 = "(=tag=)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   178
  show ?thesis
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   179
  proof(rule_tac refined_partition_finite [of _ ?R1])
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   180
    from finite_eq_f_rel [OF rng_fnt]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   181
     show "finite (UNIV // =tag=)" . 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   182
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   183
     from same_tag_eqvt
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   184
     show "(=tag=) \<subseteq> (\<approx>Lang)"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   185
       by (auto simp:f_eq_rel_def str_eq_def)
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   186
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   187
     from equiv_f_eq_rel
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   188
     show "equiv UNIV (=tag=)" by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   189
   next
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   190
     from equiv_lang_eq
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   191
     show "equiv UNIV (\<approx>Lang)" by blast
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   192
  qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   193
qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   194
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   195
text {*
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   196
  A more concise, but less intelligible argument for @{text "tag_finite_imageD"} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   197
  is given as the following. The basic idea is still using standard library 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   198
  lemma @{thm [source] "finite_imageD"}:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   199
  \[
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   200
  @{thm "finite_imageD" [no_vars]}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   201
  \]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   202
  which says: if the image of injective function @{text "f"} over set @{text "A"} is 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   203
  finite, then @{text "A"} must be finte, as we did in the lemmas above.
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   204
  *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   205
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   206
lemma 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   207
  fixes tag
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   208
  assumes rng_fnt: "finite (range tag)" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   209
  -- {* Suppose the rang of tagging fucntion @{text "tag"} is finite. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   210
  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
   211
  -- {* And strings with same tag are equivalent *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   212
  shows "finite (UNIV // (\<approx>Lang))"
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   213
  -- {* Then the partition generated by @{text "(\<approx>Lang)"} is finite. *}
42
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
  -- {* The particular @{text "f"} and @{text "A"} used in @{thm [source] "finite_imageD"} are:*}
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   216
  let "?f" =  "op ` tag" and ?A = "(UNIV // \<approx>Lang)"
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   217
  show ?thesis
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   218
  proof (rule_tac f = "?f" and A = ?A in finite_imageD) 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   219
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   220
      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
   221
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   222
    show "finite (?f ` ?A)" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   223
    proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   224
      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
   225
      moreover from rng_fnt have "finite (Pow (range tag))" by simp
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   226
      ultimately have "finite (range ?f)"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   227
        by (auto simp only:image_def intro:finite_subset)
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   228
      from finite_range_image [OF this] show ?thesis .
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   229
    qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   230
  next
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   231
    -- {* 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   232
      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
   233
      *}
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   234
    show "inj_on ?f ?A" 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   235
    proof-
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   236
      { fix X Y
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   237
        assume X_in: "X \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   238
          and  Y_in: "Y \<in> ?A"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   239
          and  tag_eq: "?f X = ?f Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   240
        have "X = Y"
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   241
        proof -
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   242
          from X_in Y_in tag_eq 
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   243
          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
   244
            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
   245
            apply simp by blast 
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   246
          from same_tag_eqvt [OF eq_tg] have "x \<approx>Lang y" .
42
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   247
          with X_in Y_in x_in y_in
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   248
          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
   249
        qed
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   250
      } thus ?thesis unfolding inj_on_def by auto
f809cb54de4e Trying to solve the confict
zhang
parents: 40
diff changeset
   251
    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
   252
  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
   253
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
   254
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   255
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
   256
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   257
text {*
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   258
  Each case is given in a separate section, as well as the final main lemma. Detailed explainations accompanied by
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   259
  illustrations are given for non-trivial cases.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   260
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   261
  For ever inductive case, there are two tasks, the easier one is to show the range finiteness of
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   262
  of the tagging function based on the finiteness of component partitions, the
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   263
  difficult one is to show that strings with the same tag are equivalent with respect to the 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   264
  composite language. Suppose the composite language be @{text "Lang"}, tagging function be 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   265
  @{text "tag"}, it amounts to show:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   266
  \[
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   267
  @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   268
  \]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   269
  expanding the definition of @{text "\<approx>Lang"}, it amounts to show:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   270
  \[
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   271
  @{text "tag(x) = tag(y) \<Longrightarrow> (\<forall> z. x@z \<in> Lang \<longleftrightarrow> y@z \<in> Lang)"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   272
  \]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   273
  Because the assumed tag equlity @{text "tag(x) = tag(y)"} is symmetric,
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   274
  it is suffcient to show just one direction:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   275
  \[
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   276
  @{text "\<And> x y z. \<lbrakk>tag(x) = tag(y); x@z \<in> Lang\<rbrakk> \<Longrightarrow> y@z \<in> Lang"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   277
  \]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   278
  This is the pattern followed by every inductive case.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   279
  *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   280
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   281
subsubsection {* The base case for @{const "NULL"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   282
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   283
lemma quot_null_eq:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   284
  shows "(UNIV // \<approx>{}) = ({UNIV}::lang set)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   285
  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
   286
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   287
lemma quot_null_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   288
  shows "finite ((UNIV // \<approx>{})::lang set)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   289
unfolding quot_null_eq by simp
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   290
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   291
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   292
subsubsection {* The base case for @{const "EMPTY"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   293
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
   294
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
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
   296
  "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
   297
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
   298
  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
   299
  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
   300
  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
   301
    unfolding quotient_def Image_def by blast
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   302
  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
   303
  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
   304
    case True with h
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   305
    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
   306
    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
   307
  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
   308
    case False with h
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   309
    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
   310
    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
   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
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
   313
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   314
lemma quot_empty_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   315
  shows "finite (UNIV // (\<approx>{[]}))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   316
by (rule finite_subset[OF quot_empty_subset]) (simp)
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
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   319
subsubsection {* The base case for @{const "CHAR"} *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   320
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
   321
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
   322
  "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
   323
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
   324
  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
   325
  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
   326
  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
   327
    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
   328
  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
   329
  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
   330
    { 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
   331
        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
   332
    } 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
   333
      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
   334
        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
   335
    } 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
   336
      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
   337
      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
   338
      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
   339
        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
   340
      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
   341
        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
   342
    } 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
   343
  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
   344
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
   345
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   346
lemma quot_char_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   347
  shows "finite (UNIV // (\<approx>{[c]}))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   348
by (rule finite_subset[OF quot_char_subset]) (simp)
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   349
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   350
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   351
subsubsection {* The inductive 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
   352
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
definition 
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   354
  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
   355
where
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   356
  "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
   357
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   358
lemma quot_union_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   359
  fixes L1 L2::"lang"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   360
  assumes finite1: "finite (UNIV // \<approx>L1)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   361
  and     finite2: "finite (UNIV // \<approx>L2)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   362
  shows "finite (UNIV // \<approx>(L1 \<union> L2))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   363
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
   364
  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
   365
    unfolding tag_str_ALT_def 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   366
    unfolding str_eq_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   367
    unfolding Image_def 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   368
    unfolding str_eq_rel_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   369
    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
   370
next
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   371
  have *: "finite ((UNIV // \<approx>L1) \<times> (UNIV // \<approx>L2))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   372
    using finite1 finite2 by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   373
  show "finite (range (tag_str_ALT L1 L2))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   374
    unfolding tag_str_ALT_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   375
    apply(rule finite_subset[OF _ *])
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   376
    unfolding quotient_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   377
    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
   378
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
   379
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   380
subsubsection {* The inductive case for @{text "SEQ"}*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   381
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   382
text {*
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   383
  For case @{const "SEQ"}, the language @{text "L"} is @{text "L\<^isub>1 ;; L\<^isub>2"}.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   384
  Given @{text "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"}, according to the defintion of @{text " L\<^isub>1 ;; L\<^isub>2"},
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   385
  string @{text "x @ z"} can be splitted with the prefix in @{text "L\<^isub>1"} and suffix in @{text "L\<^isub>2"}.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   386
  The split point can either be in @{text "x"} (as shown in Fig. \ref{seq_first_split}),
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   387
  or in @{text "z"} (as shown in Fig. \ref{seq_snd_split}). Whichever way it goes, the structure
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   388
  on @{text "x @ z"} cn be transfered faithfully onto @{text "y @ z"} 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   389
  (as shown in Fig. \ref{seq_trans_first_split} and \ref{seq_trans_snd_split}) with the the help of the assumed 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   390
  tag equality. The following tag function @{text "tag_str_SEQ"} is such designed to facilitate
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   391
  such transfers and lemma @{text "tag_str_SEQ_injI"} formalizes the informal argument above. The details 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   392
  of structure transfer will be given their.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   393
\input{fig_seq}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   394
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   395
  *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   396
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   397
definition 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   398
  tag_str_SEQ :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang set)"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   399
where
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   400
  "tag_str_SEQ L1 L2 = 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   401
     (\<lambda>x. (\<approx>L1 `` {x}, {(\<approx>L2 `` {x - xa}) | xa.  xa \<le> x \<and> xa \<in> L1}))"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   402
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   403
text {* The following is a techical lemma which helps to split the @{text "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"} mentioned above.*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   404
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   405
lemma append_seq_elim:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   406
  assumes "x @ y \<in> L\<^isub>1 ;; L\<^isub>2"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   407
  shows "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2) \<or> 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   408
          (\<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2)"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   409
proof-
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   410
  from assms obtain s\<^isub>1 s\<^isub>2 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   411
    where eq_xys: "x @ y = s\<^isub>1 @ s\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   412
    and in_seq: "s\<^isub>1 \<in> L\<^isub>1 \<and> s\<^isub>2 \<in> L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   413
    by (auto simp:Seq_def)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   414
  from app_eq_dest [OF eq_xys]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   415
  have
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   416
    "(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)" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   417
               (is "?Split1 \<or> ?Split2") .
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   418
  moreover have "?Split1 \<Longrightarrow> \<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   419
    using in_seq by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   420
  moreover have "?Split2 \<Longrightarrow> \<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   421
    using in_seq by (rule_tac x = s\<^isub>1 in exI, auto)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   422
  ultimately show ?thesis by blast
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   423
qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   424
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   425
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   426
lemma tag_str_SEQ_injI:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   427
  fixes v w 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   428
  assumes eq_tag: "tag_str_SEQ L\<^isub>1 L\<^isub>2 v = tag_str_SEQ L\<^isub>1 L\<^isub>2 w" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   429
  shows "v \<approx>(L\<^isub>1 ;; L\<^isub>2) w"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   430
proof-
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   431
    -- {* As explained before, a pattern for just one direction needs to be dealt with:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   432
  { fix x y z
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   433
    assume xz_in_seq: "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   434
    and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   435
    have"y @ z \<in> L\<^isub>1 ;; L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   436
    proof-
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   437
      -- {* There are two ways to split @{text "x@z"}: *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   438
      from append_seq_elim [OF xz_in_seq]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   439
      have "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   440
               (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)" .
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   441
      -- {* It can be shown that @{text "?thesis"} holds in either case: *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   442
      moreover {
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   443
        -- {* The case for the first split:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   444
        fix xa
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   445
        assume h1: "xa \<le> x" and h2: "xa \<in> L\<^isub>1" and h3: "(x - xa) @ z \<in> L\<^isub>2"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   446
        -- {* The following subgoal implements the structure transfer:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   447
        obtain ya 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   448
          where "ya \<le> y" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   449
          and "ya \<in> L\<^isub>1" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   450
          and "(y - ya) @ z \<in> L\<^isub>2"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   451
        proof -
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   452
        -- {*
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   453
            \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   454
            By expanding the definition of 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   455
            @{thm [display] "tag_xy"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   456
            and extracting the second compoent, we get:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   457
            \end{minipage}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   458
            *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   459
          have "{\<approx>L\<^isub>2 `` {x - xa} |xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   460
                   {\<approx>L\<^isub>2 `` {y - ya} |ya. ya \<le> y \<and> ya \<in> L\<^isub>1}" (is "?Left = ?Right")
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   461
            using tag_xy unfolding tag_str_SEQ_def by simp
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   462
            -- {* Since @{thm "h1"} and @{thm "h2"} hold, it is not difficult to show: *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   463
          moreover have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Left" using h1 h2 by auto
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   464
            -- {* 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   465
            \begin{minipage}{0.7\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   466
            Through tag equality, equivalent class @{term "\<approx>L\<^isub>2 `` {x - xa}"} also 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   467
                  belongs to the @{text "?Right"}:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   468
            \end{minipage}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   469
            *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   470
          ultimately have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Right" by simp
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   471
            -- {* From this, the counterpart of @{text "xa"} in @{text "y"} is obtained:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   472
          then obtain ya 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   473
            where eq_xya: "\<approx>L\<^isub>2 `` {x - xa} = \<approx>L\<^isub>2 `` {y - ya}" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   474
            and pref_ya: "ya \<le> y" and ya_in: "ya \<in> L\<^isub>1"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   475
            by simp blast
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   476
          -- {* It can be proved that @{text "ya"} has the desired property:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   477
          have "(y - ya)@z \<in> L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   478
          proof -
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   479
            from eq_xya have "(x - xa)  \<approx>L\<^isub>2 (y - ya)" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   480
              unfolding Image_def str_eq_rel_def str_eq_def by auto
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   481
            with h3 show ?thesis unfolding str_eq_rel_def str_eq_def by simp
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   482
          qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   483
          -- {* Now, @{text "ya"} has all properties to be a qualified candidate:*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   484
          with pref_ya ya_in 
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   485
          show ?thesis using that by blast
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   486
        qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   487
          -- {* From the properties of @{text "ya"}, @{text "y @ z \<in> L\<^isub>1 ;; L\<^isub>2"} is derived easily.*}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   488
        hence "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" by (erule_tac prefixE, auto simp:Seq_def)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   489
      } moreover {
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   490
        -- {* The other case is even more simpler: *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   491
        fix za
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   492
        assume h1: "za \<le> z" and h2: "(x @ za) \<in> L\<^isub>1" and h3: "z - za \<in> L\<^isub>2"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   493
        have "y @ za \<in> L\<^isub>1"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   494
        proof-
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   495
          have "\<approx>L\<^isub>1 `` {x} = \<approx>L\<^isub>1 `` {y}" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   496
            using tag_xy unfolding tag_str_SEQ_def by simp
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   497
          with h2 show ?thesis
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   498
            unfolding Image_def str_eq_rel_def str_eq_def by auto
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   499
        qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   500
        with h1 h3 have "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   501
          by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   502
      }
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   503
      ultimately show ?thesis by blast
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   504
    qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   505
  } 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   506
  -- {* 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   507
      \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   508
      @{text "?thesis"} is proved by exploiting the symmetry of 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   509
      @{thm [source] "eq_tag"}:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   510
      \end{minipage}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   511
      *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   512
  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   513
    show ?thesis unfolding str_eq_def str_eq_rel_def by blast
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   514
qed 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   515
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   516
lemma quot_seq_finiteI [intro]:
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   517
  fixes L1 L2::"lang"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   518
  assumes fin1: "finite (UNIV // \<approx>L1)" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   519
  and     fin2: "finite (UNIV // \<approx>L2)" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   520
  shows "finite (UNIV // \<approx>(L1 ;; L2))"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   521
proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   522
  show "\<And>x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \<Longrightarrow> x \<approx>(L1 ;; L2) y"
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   523
    by (rule tag_str_SEQ_injI)
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   524
next
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   525
  have *: "finite ((UNIV // \<approx>L1) \<times> (Pow (UNIV // \<approx>L2)))" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   526
    using fin1 fin2 by auto
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   527
  show "finite (range (tag_str_SEQ L1 L2))" 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   528
    unfolding tag_str_SEQ_def
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   529
    apply(rule finite_subset[OF _ *])
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   530
    unfolding quotient_def
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   531
    by auto
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   532
qed
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   533
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   534
subsubsection {* The inductive 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
   535
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
   536
text {* 
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   537
  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
   538
  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
   539
  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
   540
  \begin{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   541
    \item Since @{text "x @ z \<in>  L\<^isub>1*"} holds, a prefix @{text "xa"} of @{text "x"} can be found
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   542
          such that @{text "xa \<in> L\<^isub>1*"} and @{text "(x - xa)@z \<in> L\<^isub>1*"}, as shown in Fig. \ref{first_split}.
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   543
          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
   544
    \item There could be many but fintie many of such @{text "xa"}, from which we can find the longest
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   545
          and name it @{text "xa_max"}, as shown in Fig. \ref{max_split}.
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   546
    \item The next step is to split @{text "z"} into @{text "za"} and @{text "zb"} such that
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   547
           @{text "(x - xa_max) @ za \<in> L\<^isub>1"} and @{text "zb \<in> L\<^isub>1*"}  as shown in Fig. \ref{last_split}.
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   548
          Such a split always exists because:
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   549
          \begin{enumerate}
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   550
            \item Because @{text "(x - x_max) @ z \<in> L\<^isub>1*"}, it can always be splitted into prefix @{text "a"}
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   551
              and suffix @{text "b"}, such that @{text "a \<in> L\<^isub>1"} and @{text "b \<in> L\<^isub>1*"},
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   552
              as shown in Fig. \ref{ab_split}.
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   553
            \item But the prefix @{text "a"} CANNOT be shorter than @{text "x - xa_max"} 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   554
              (as shown in Fig. \ref{ab_split_wrong}), becasue otherwise,
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   555
                   @{text "ma_max@a"} would be in the same kind as @{text "xa_max"} but with 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   556
                   a larger size, conflicting with the fact that @{text "xa_max"} is the longest.
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   557
          \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   558
    \item  \label{tansfer_step} 
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   559
         By the assumption that @{text "x"} and @{text "y"} have the same tag, the structure on @{text "x @ z"}
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   560
          can be transferred to @{text "y @ z"} as shown in Fig. \ref{trans_split}. The detailed steps are:
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   561
          \begin{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   562
            \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
   563
                  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
   564
            \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
   565
                  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
   566
            \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
   567
          \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   568
  \end{enumerate}
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   569
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   570
  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
   571
  while the tagging function @{text "tag_str_STAR"} is defined to make the transfer in step
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   572
  \ref{ansfer_step} feasible.
49
59936c012add Illustration added together with renewed explainations for case STAR.
zhang
parents: 48
diff changeset
   573
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   574
  \input{fig_star}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   575
*} 
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
   576
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
   577
definition 
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   578
  tag_str_STAR :: "lang \<Rightarrow> string \<Rightarrow> lang set"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   579
where
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   580
  "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
   581
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   582
text {* A technical lemma. *}
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   583
lemma finite_set_has_max: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   584
           (\<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
   585
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
   586
  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
   587
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
   588
  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
   589
  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
   590
  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
   591
    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
   592
  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
   593
    case False
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   594
    with insertI.hyps and False  
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   595
    obtain max 
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
   596
      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
   597
      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
   598
    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
   599
    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
   600
      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
   601
      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
   602
    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
   603
      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
   604
      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
   605
    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
   606
  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
   607
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
   608
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   609
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   610
text {* The following is a technical lemma.which helps to show the range finiteness of tag function. *}
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
   611
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
   612
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
   613
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
   614
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
   615
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   616
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
   617
lemma tag_str_STAR_injI:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   618
  fixes v w
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   619
  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
   620
  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
   621
proof-
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   622
    -- {* As explained before, a pattern for just one direction needs to be dealt with:*}
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
   623
  { fix x y z
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   624
    assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>" 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   625
      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
   626
    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
   627
    proof(cases "x = []")
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   628
      -- {* 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   629
        The degenerated case when @{text "x"} is a null string is easy to prove:
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   630
        *}
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
   631
      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
   632
      with tag_xy have "y = []" 
56
b3898315e687 removed the inductive definition of Star and replaced it by a definition in terms of pow
urbanc
parents: 55
diff changeset
   633
        by (auto simp add: tag_str_STAR_def strict_prefix_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
   634
      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
   635
    next
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   636
        -- {* The nontrival case:
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   637
        *}
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
   638
      case False
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   639
      -- {* 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   640
        \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   641
        Since @{text "x @ z \<in> L\<^isub>1\<star>"}, @{text "x"} can always be splitted
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   642
        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
   643
        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
   644
        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
   645
        is nonempty, and finite as well:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   646
        \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   647
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   648
      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
   649
      have "finite ?S"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   650
        by (rule_tac B = "{xa. xa < x}" in finite_subset, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   651
          auto simp:finite_strict_prefix_set)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   652
      moreover have "?S \<noteq> {}" using False xz_in_star
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   653
        by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   654
      -- {* \begin{minipage}{0.7\textwidth} 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   655
            Since @{text "?S"} is finite, we can always single out the longest and name it @{text "xa_max"}: 
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   656
            \end{minipage}
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   657
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   658
      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
   659
        using finite_set_has_max by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   660
      then obtain xa_max 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   661
        where h1: "xa_max < x" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   662
        and h2: "xa_max \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   663
        and h3: "(x - xa_max) @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   664
        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
   665
                                     \<longrightarrow> length xa \<le> length xa_max"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   666
        by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   667
      -- {*
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   668
          \begin{minipage}{0.8\textwidth}
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   669
          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
   670
          @{text "y"}-prefixes, named @{text "ya"}, can be found:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   671
          \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   672
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   673
      obtain ya 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   674
        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
   675
        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
   676
      proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   677
        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
   678
          {\<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
   679
          by (auto simp:tag_str_STAR_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   680
        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
   681
        ultimately have "\<approx>L\<^isub>1 `` {x - xa_max} \<in> ?right" by simp
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   682
        thus ?thesis using that 
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   683
          apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   684
      qed 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   685
      -- {*
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   686
          \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   687
          The @{text "?thesis"}, @{prop "y @ z \<in> L\<^isub>1\<star>"}, is a simple consequence
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   688
          of the following proposition:
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   689
          \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   690
          *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   691
      have "(y - ya) @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   692
      proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   693
        -- {* 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
   694
          such that: *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   695
        obtain za zb where eq_zab: "z = za @ zb" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   696
          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
   697
        proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   698
          -- {* 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   699
            \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   700
            Since @{thm "h1"}, @{text "x"} can be splitted into
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   701
            @{text "a"} and @{text "b"} such that:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   702
            \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   703
            *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   704
          from h1 have "(x - xa_max) @ z \<noteq> []" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   705
            by (auto simp:strict_prefix_def elim:prefixE)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   706
          from star_decom [OF h3 this]
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   707
          obtain a b where a_in: "a \<in> L\<^isub>1" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   708
            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
   709
            and ab_max: "(x - xa_max) @ z = a @ b" by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   710
          -- {* Now the candiates for @{text "za"} and @{text "zb"} are found:*}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   711
          let ?za = "a - (x - xa_max)" and ?zb = "b"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   712
          have pfx: "(x - xa_max) \<le> a" (is "?P1") 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   713
            and eq_z: "z = ?za @ ?zb" (is "?P2")
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   714
          proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   715
            -- {* 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   716
              \begin{minipage}{0.8\textwidth}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   717
              Since @{text "(x - xa_max) @ z = a @ b"}, string @{text "(x - xa_max) @ z"}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   718
              can be splitted in two ways:
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   719
              \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   720
              *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   721
            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
   722
              (a < (x - xa_max) \<and> ((x - xa_max) - a) @ z = b)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   723
              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
   724
            moreover {
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   725
              -- {* However, the undsired way can be refuted by absurdity: *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   726
              assume np: "a < (x - xa_max)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   727
                and b_eqs: "((x - xa_max) - a) @ z = b"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   728
              have "False"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   729
              proof -
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   730
                let ?xa_max' = "xa_max @ a"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   731
                have "?xa_max' < x" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   732
                  using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   733
                moreover have "?xa_max' \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   734
                  using a_in h2 by (simp add:star_intro3) 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   735
                moreover have "(x - ?xa_max') @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   736
                  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
   737
                moreover have "\<not> (length ?xa_max' \<le> length xa_max)" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   738
                  using a_neq by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   739
                ultimately show ?thesis using h4 by blast
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   740
              qed }
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   741
            -- {* 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
   742
            ultimately show ?P1 and ?P2 by auto
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   743
          qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   744
          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
   745
          -- {* 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
   746
          with eq_xya have "(y - ya) @ ?za \<in> L\<^isub>1" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   747
            by (auto simp:str_eq_def str_eq_rel_def)
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   748
           with eq_z and b_in 
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   749
          show ?thesis using that by blast
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   750
        qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   751
        -- {* 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   752
           @{text "?thesis"} can easily be shown using properties of @{text "za"} and @{text "zb"}:
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   753
            *}
56
b3898315e687 removed the inductive definition of Star and replaced it by a definition in terms of pow
urbanc
parents: 55
diff changeset
   754
        have "((y - ya) @ za) @ zb \<in> L\<^isub>1\<star>" using  l_za ls_zb by blast
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   755
        with eq_zab show ?thesis by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   756
      qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   757
      with h5 h6 show ?thesis 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   758
        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
   759
    qed
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
  -- {* By instantiating the reasoning pattern just derived for both directions:*} 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   762
  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
   763
  -- {* The thesis is proved as a trival consequence: *} 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   764
    show  ?thesis unfolding str_eq_def str_eq_rel_def by blast
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   765
qed
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   766
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   767
lemma -- {* The oringal version with less explicit details. *}
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   768
  fixes v w
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   769
  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
   770
  shows "(v::string) \<approx>(L\<^isub>1\<star>) w"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   771
proof-
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   772
    -- {* 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   773
    \begin{minipage}{0.8\textwidth}
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   774
    According to the definition of @{text "\<approx>Lang"}, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   775
    proving @{text "v \<approx>(L\<^isub>1\<star>) w"} amounts to
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   776
    showing: for any string @{text "u"},
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   777
    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
   778
    The reasoning pattern for both directions are the same, as derived
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   779
    in the following:
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   780
    \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   781
    *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   782
  { fix x y z
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   783
    assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   784
      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
   785
    have "y @ z \<in> L\<^isub>1\<star>"
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   786
    proof(cases "x = []")
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   787
      -- {* 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   788
        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
   789
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   790
      case True
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   791
      with tag_xy have "y = []" 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   792
        by (auto simp:tag_str_STAR_def strict_prefix_def)
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   793
      thus ?thesis using xz_in_star True by simp
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   794
    next
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   795
        -- {*
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   796
        \begin{minipage}{0.8\textwidth}
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   797
        The case when @{text "x"} is not null, and
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   798
        @{text "x @ z"} is in @{text "L\<^isub>1\<star>"}, 
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   799
        \end{minipage}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   800
        *}
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   801
      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
   802
      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
   803
        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
   804
        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
   805
        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
   806
        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
   807
                                     \<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
   808
      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
   809
        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
   810
        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
   811
          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
   812
                                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
   813
        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
   814
          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
   815
        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
   816
          using finite_set_has_max by blast
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   817
        thus ?thesis using that 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
   818
      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
   819
      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
   820
        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
   821
      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
   822
        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
   823
          {\<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
   824
          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
   825
        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
   826
        ultimately have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?right" by simp
57
76ab7c09d575 Myhill.pdf modified
zhang
parents: 56
diff changeset
   827
        with that show ?thesis apply 
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
   828
          (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
   829
      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
   830
      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
   831
      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
   832
        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
   833
          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
   834
          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
   835
          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
   836
        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
   837
        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
   838
          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
   839
                            (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
   840
            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
   841
          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
   842
            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
   843
            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
   844
            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
   845
              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
   846
              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
   847
                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
   848
              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
   849
                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
   850
              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
   851
                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
   852
              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
   853
                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
   854
              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
   855
            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
   856
          } 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
   857
        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
   858
        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
   859
          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
   860
          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
   861
        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
   862
          by (auto simp:str_eq_def str_eq_rel_def)
56
b3898315e687 removed the inductive definition of Star and replaced it by a definition in terms of pow
urbanc
parents: 55
diff changeset
   863
	with b_in have "((y - ya) @ za) @ b \<in> L\<^isub>1\<star>" by blast
b3898315e687 removed the inductive definition of Star and replaced it by a definition in terms of pow
urbanc
parents: 55
diff changeset
   864
        with z_decom show ?thesis 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
   865
      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
   866
      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
   867
        by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
45
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   868
    qed
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   869
  } 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   870
  -- {* By instantiating the reasoning pattern just derived for both directions:*} 
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   871
  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
7aa6c20e6d31 More improvement
zhang
parents: 43
diff changeset
   872
  -- {* The thesis is proved as a trival consequence: *} 
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   873
    show  ?thesis unfolding str_eq_def str_eq_rel_def 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
   874
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
   875
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   876
lemma quot_star_finiteI [intro]:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   877
  fixes L1::"lang"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   878
  assumes finite1: "finite (UNIV // \<approx>L1)"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   879
  shows "finite (UNIV // \<approx>(L1\<star>))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   880
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
   881
  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
   882
    by (rule tag_str_STAR_injI)
40
50d00d7dc413 tuned a bit more the last STAR-proof
wu
parents: 39
diff changeset
   883
next
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   884
  have *: "finite (Pow (UNIV // \<approx>L1))" 
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   885
    using finite1 by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   886
  show "finite (range (tag_str_STAR L1))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   887
    unfolding tag_str_STAR_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   888
    apply(rule finite_subset[OF _ *])
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   889
    unfolding quotient_def
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   890
    by auto
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   891
qed
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   892
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
   893
subsubsection{* The conclusion *}
47
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   894
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   895
lemma rexp_imp_finite:
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   896
  fixes r::"rexp"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   897
  shows "finite (UNIV // \<approx>(L r))"
bea2466a6084 slightly tuned the main lemma and the finiteness proofs
urbanc
parents: 45
diff changeset
   898
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
   899
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
end