diff -r 97090fc7aa9f -r 560712a29a36 Attic/Prefix_subtract.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/Prefix_subtract.thy Tue Aug 02 15:27:37 2011 +0000 @@ -0,0 +1,60 @@ +theory Prefix_subtract + imports Main "~~/src/HOL/Library/List_Prefix" +begin + + +section {* A small theory of prefix subtraction *} + +text {* + The notion of @{text "prefix_subtract"} makes + the second direction of the Myhill-Nerode theorem + more readable. +*} + +instantiation list :: (type) minus +begin + +fun minus_list :: "'a list \ 'a list \ 'a list" +where + "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" +by (induct x) (auto) + +lemma [simp]: "x - x = []" +by (induct x) (auto) + +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) + +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 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 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