diff -r a8a442ba0dbf -r e93760534354 Prefix_subtract.thy --- a/Prefix_subtract.thy Thu May 12 05:55:05 2011 +0000 +++ b/Prefix_subtract.thy Wed May 18 19:54:43 2011 +0000 @@ -1,59 +1,60 @@ theory Prefix_subtract - imports Main - "~~/src/HOL/Library/List_Prefix" + imports Main "~~/src/HOL/Library/List_Prefix" begin + section {* A small theory of prefix subtraction *} text {* - The notion of @{text "prefix_subtract"} is need to make proofs more readable. - *} + The notion of @{text "prefix_subtract"} makes + the second direction of the Myhill-Nerode theorem + more readable. +*} + +instantiation list :: (type) minus +begin -fun prefix_subtract :: "'a list \ 'a list \ 'a list" (infix "-" 51) +fun minus_list :: "'a list \ 'a list \ 'a list" where - "prefix_subtract [] xs = []" -| "prefix_subtract (x#xs) [] = x#xs" -| "prefix_subtract (x#xs) (y#ys) = (if x = y then prefix_subtract xs ys else (x#xs))" + "minus_list [] xs = []" +| "minus_list (x#xs) [] = x#xs" +| "minus_list (x#xs) (y#ys) = (if x = y then minus_list xs ys else (x#xs))" + +instance by default + +end + +lemma [simp]: "x - [] = x" +by (induct x) (auto) lemma [simp]: "(x @ y) - x = y" -apply (induct x) -by (case_tac y, simp+) +by (induct x) (auto) lemma [simp]: "x - x = []" -by (induct x, auto) - -lemma [simp]: "x = xa @ y \ x - xa = y " -by (induct x, auto) - -lemma [simp]: "x - [] = x" -by (induct x, auto) +by (induct x) (auto) -lemma [simp]: "(x - y = []) \ (x \ y)" -proof- - have "\xa. x = xa @ (x - y) \ xa \ y" - apply (rule prefix_subtract.induct[of _ x y], simp+) - by (clarsimp, rule_tac x = "y # xa" in exI, simp+) - thus "(x - y = []) \ (x \ y)" by simp -qed +lemma [simp]: "x = z @ y \ x - z = y " +by (induct x) (auto) lemma diff_prefix: "\c \ a - b; b \ a\ \ b @ c \ a" -by (auto elim:prefixE) +by (auto elim: prefixE) -lemma diff_diff_appd: +lemma diff_diff_append: "\c < a - b; b < a\ \ (a - b) - c = a - (b @ c)" apply (clarsimp simp:strict_prefix_def) by (drule diff_prefix, auto elim:prefixE) -lemma app_eq_cases[rule_format]: - "\ x . x @ y = m @ n \ (x \ m \ m \ x)" -apply (induct y, simp) -apply (clarify, drule_tac x = "x @ [a]" in spec) -by (clarsimp, auto simp:prefix_def) +lemma append_eq_cases: + assumes a: "x @ y = m @ n" + shows "x \ m \ m \ x" +unfolding prefix_def using a +by (auto simp add: append_eq_append_conv2) -lemma app_eq_dest: - "x @ y = m @ n \ - (x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" -by (frule_tac app_eq_cases, auto elim:prefixE) +lemma append_eq_dest: + assumes a: "x @ y = m @ n" + shows "(x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" +using append_eq_cases[OF a] a +by (auto elim: prefixE) end