diff -r f5db9e08effc -r b6815473ee2e Prefix_subtract.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Prefix_subtract.thy Tue Jan 25 12:14:31 2011 +0000 @@ -0,0 +1,58 @@ +theory Prefix_subtract + imports Main List_Prefix +begin + +section {* A small theory of prefix subtraction *} + +text {* + The notion of @{text "prefix_subtract"} is need to make proofs more readable. + *} + +fun prefix_subtract :: "'a list \ 'a list \ 'a list" (infix "-" 51) +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))" + +lemma [simp]: "(x @ y) - x = y" +apply (induct x) +by (case_tac y, simp+) + +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) + +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 diff_prefix: + "\c \ a - b; b \ a\ \ b @ c \ a" +by (auto elim:prefixE) + +lemma diff_diff_appd: + "\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 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) + +end