Prefix_subtract.thy
author urbanc
Wed, 18 May 2011 19:54:43 +0000
changeset 162 e93760534354
parent 149 e122cb146ecc
permissions -rw-r--r--
added directory for journal version; took uptodate version of the theory files
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:
diff changeset
     1
theory Prefix_subtract
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
     2
  imports Main "~~/src/HOL/Library/List_Prefix"
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:
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:
diff changeset
     4
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
     5
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:
diff changeset
     6
section {* A small theory of prefix subtraction *}
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
     7
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
     8
text {*
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
     9
  The notion of @{text "prefix_subtract"} makes 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    10
  the second direction of the Myhill-Nerode theorem 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    11
  more readable.
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    12
*}
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    13
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    14
instantiation list :: (type) minus
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    15
begin
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:
diff changeset
    16
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    17
fun minus_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" 
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:
diff changeset
    18
where
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    19
  "minus_list []     xs     = []" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    20
| "minus_list (x#xs) []     = x#xs" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    21
| "minus_list (x#xs) (y#ys) = (if x = y then minus_list xs ys else (x#xs))"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    22
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    23
instance by default
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    24
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    25
end
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    26
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    27
lemma [simp]: "x - [] = x"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    28
by (induct x) (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:
diff changeset
    29
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
    30
lemma [simp]: "(x @ y) - x = y"
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    31
by (induct x) (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:
diff changeset
    32
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
    33
lemma [simp]: "x - x = []"
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    34
by (induct x) (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:
diff changeset
    35
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    36
lemma [simp]: "x = z @ y \<Longrightarrow> x - z = y "
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    37
by (induct x) (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:
diff changeset
    38
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
    39
lemma 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:
diff changeset
    40
  "\<lbrakk>c \<le> a - b; b \<le> a\<rbrakk> \<Longrightarrow> b @ c \<le> a"
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    41
by (auto elim: prefixE)
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:
diff changeset
    42
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    43
lemma diff_diff_append: 
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:
diff changeset
    44
  "\<lbrakk>c < a - b; b < a\<rbrakk> \<Longrightarrow> (a - b) - c = a - (b @ 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:
diff changeset
    45
apply (clarsimp 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:
diff changeset
    46
by (drule diff_prefix, 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:
diff changeset
    47
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    48
lemma append_eq_cases:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    49
  assumes a: "x @ y = m @ n"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    50
  shows "x \<le> m \<or> m \<le> x"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    51
unfolding prefix_def using a
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    52
by (auto simp add: append_eq_append_conv2)
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:
diff changeset
    53
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    54
lemma append_eq_dest:
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    55
  assumes a: "x @ y = m @ n" 
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    56
  shows "(x \<le> m \<and> (m - x) @ n = y) \<or> (m \<le> x \<and> (x - m) @ y = n)"
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    57
using append_eq_cases[OF a] a
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 149
diff changeset
    58
by (auto elim: prefixE)
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:
diff changeset
    59
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents:
diff changeset
    60
end