(* Title: HOL/Library/List_Prefix.thy
ID: $Id: List_Prefix.thy,v 1.16 2005/08/31 13:46:38 wenzelm Exp $
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
*)
header {* List prefixes and postfixes *}
theory List_Prefix
imports Main
begin
subsection {* Prefix order on lists *}
instance list :: (type) ord ..
defs (overloaded)
prefix_def: "xs ≤ ys == ∃zs. ys = xs @ zs"
strict_prefix_def: "xs < ys == xs ≤ ys ∧ xs ≠ (ys::'a list)"
instance list :: (type) order
by intro_classes (auto simp add: prefix_def strict_prefix_def)
lemma prefixI [intro?]: "ys = xs @ zs ==> xs ≤ ys"
by (unfold prefix_def) blast
lemma prefixE [elim?]: "xs ≤ ys ==> (!!zs. ys = xs @ zs ==> C) ==> C"
by (unfold prefix_def) blast
lemma strict_prefixI' [intro?]: "ys = xs @ z # zs ==> xs < ys"
by (unfold strict_prefix_def prefix_def) blast
lemma strict_prefixE' [elim?]:
assumes lt: "xs < ys"
and r: "!!z zs. ys = xs @ z # zs ==> C"
shows C
proof -
from lt obtain us where "ys = xs @ us" and "xs ≠ ys"
by (unfold strict_prefix_def prefix_def) blast
with r show ?thesis by (auto simp add: neq_Nil_conv)
qed
lemma strict_prefixI [intro?]: "xs ≤ ys ==> xs ≠ ys ==> xs < (ys::'a list)"
by (unfold strict_prefix_def) blast
lemma strict_prefixE [elim?]:
"xs < ys ==> (xs ≤ ys ==> xs ≠ (ys::'a list) ==> C) ==> C"
by (unfold strict_prefix_def) blast
subsection {* Basic properties of prefixes *}
theorem Nil_prefix [iff]: "[] ≤ xs"
by (simp add: prefix_def)
theorem prefix_Nil [simp]: "(xs ≤ []) = (xs = [])"
by (induct xs) (simp_all add: prefix_def)
lemma prefix_snoc [simp]: "(xs ≤ ys @ [y]) = (xs = ys @ [y] ∨ xs ≤ ys)"
proof
assume "xs ≤ ys @ [y]"
then obtain zs where zs: "ys @ [y] = xs @ zs" ..
show "xs = ys @ [y] ∨ xs ≤ ys"
proof (cases zs rule: rev_cases)
assume "zs = []"
with zs have "xs = ys @ [y]" by simp
thus ?thesis ..
next
fix z zs' assume "zs = zs' @ [z]"
with zs have "ys = xs @ zs'" by simp
hence "xs ≤ ys" ..
thus ?thesis ..
qed
next
assume "xs = ys @ [y] ∨ xs ≤ ys"
thus "xs ≤ ys @ [y]"
proof
assume "xs = ys @ [y]"
thus ?thesis by simp
next
assume "xs ≤ ys"
then obtain zs where "ys = xs @ zs" ..
hence "ys @ [y] = xs @ (zs @ [y])" by simp
thus ?thesis ..
qed
qed
lemma Cons_prefix_Cons [simp]: "(x # xs ≤ y # ys) = (x = y ∧ xs ≤ ys)"
by (auto simp add: prefix_def)
lemma same_prefix_prefix [simp]: "(xs @ ys ≤ xs @ zs) = (ys ≤ zs)"
by (induct xs) simp_all
lemma same_prefix_nil [iff]: "(xs @ ys ≤ xs) = (ys = [])"
proof -
have "(xs @ ys ≤ xs @ []) = (ys ≤ [])" by (rule same_prefix_prefix)
thus ?thesis by simp
qed
lemma prefix_prefix [simp]: "xs ≤ ys ==> xs ≤ ys @ zs"
proof -
assume "xs ≤ ys"
then obtain us where "ys = xs @ us" ..
hence "ys @ zs = xs @ (us @ zs)" by simp
thus ?thesis ..
qed
lemma append_prefixD: "xs @ ys ≤ zs ==> xs ≤ zs"
by (auto simp add: prefix_def)
theorem prefix_Cons: "(xs ≤ y # ys) = (xs = [] ∨ (∃zs. xs = y # zs ∧ zs ≤ ys))"
by (cases xs) (auto simp add: prefix_def)
theorem prefix_append:
"(xs ≤ ys @ zs) = (xs ≤ ys ∨ (∃us. xs = ys @ us ∧ us ≤ zs))"
apply (induct zs rule: rev_induct)
apply force
apply (simp del: append_assoc add: append_assoc [symmetric])
apply simp
apply blast
done
lemma append_one_prefix:
"xs ≤ ys ==> length xs < length ys ==> xs @ [ys ! length xs] ≤ ys"
apply (unfold prefix_def)
apply (auto simp add: nth_append)
apply (case_tac zs)
apply auto
done
theorem prefix_length_le: "xs ≤ ys ==> length xs ≤ length ys"
by (auto simp add: prefix_def)
lemma prefix_same_cases:
"(xs1::'a list) ≤ ys ==> xs2 ≤ ys ==> xs1 ≤ xs2 ∨ xs2 ≤ xs1"
apply (simp add: prefix_def)
apply (erule exE)+
apply (simp add: append_eq_append_conv_if split: if_splits)
apply (rule disjI2)
apply (rule_tac x = "drop (size xs2) xs1" in exI)
apply clarify
apply (drule sym)
apply (insert append_take_drop_id [of "length xs2" xs1])
apply simp
apply (rule disjI1)
apply (rule_tac x = "drop (size xs1) xs2" in exI)
apply clarify
apply (insert append_take_drop_id [of "length xs1" xs2])
apply simp
done
lemma set_mono_prefix:
"xs ≤ ys ==> set xs ⊆ set ys"
by (auto simp add: prefix_def)
subsection {* Parallel lists *}
constdefs
parallel :: "'a list => 'a list => bool" (infixl "\<parallel>" 50)
"xs \<parallel> ys == ¬ xs ≤ ys ∧ ¬ ys ≤ xs"
lemma parallelI [intro]: "¬ xs ≤ ys ==> ¬ ys ≤ xs ==> xs \<parallel> ys"
by (unfold parallel_def) blast
lemma parallelE [elim]:
"xs \<parallel> ys ==> (¬ xs ≤ ys ==> ¬ ys ≤ xs ==> C) ==> C"
by (unfold parallel_def) blast
theorem prefix_cases:
"(xs ≤ ys ==> C) ==>
(ys < xs ==> C) ==>
(xs \<parallel> ys ==> C) ==> C"
by (unfold parallel_def strict_prefix_def) blast
theorem parallel_decomp:
"xs \<parallel> ys ==> ∃as b bs c cs. b ≠ c ∧ xs = as @ b # bs ∧ ys = as @ c # cs"
proof (induct xs rule: rev_induct)
case Nil
hence False by auto
thus ?case ..
next
case (snoc x xs)
show ?case
proof (rule prefix_cases)
assume le: "xs ≤ ys"
then obtain ys' where ys: "ys = xs @ ys'" ..
show ?thesis
proof (cases ys')
assume "ys' = []" with ys have "xs = ys" by simp
with snoc have "[x] \<parallel> []" by auto
hence False by blast
thus ?thesis ..
next
fix c cs assume ys': "ys' = c # cs"
with snoc ys have "xs @ [x] \<parallel> xs @ c # cs" by (simp only:)
hence "x ≠ c" by auto
moreover have "xs @ [x] = xs @ x # []" by simp
moreover from ys ys' have "ys = xs @ c # cs" by (simp only:)
ultimately show ?thesis by blast
qed
next
assume "ys < xs" hence "ys ≤ xs @ [x]" by (simp add: strict_prefix_def)
with snoc have False by blast
thus ?thesis ..
next
assume "xs \<parallel> ys"
with snoc obtain as b bs c cs where neq: "(b::'a) ≠ c"
and xs: "xs = as @ b # bs" and ys: "ys = as @ c # cs"
by blast
from xs have "xs @ [x] = as @ b # (bs @ [x])" by simp
with neq ys show ?thesis by blast
qed
qed
subsection {* Postfix order on lists *}
constdefs
postfix :: "'a list => 'a list => bool" ("(_/ >>= _)" [51, 50] 50)
"xs >>= ys == ∃zs. xs = zs @ ys"
lemma postfix_refl [simp, intro!]: "xs >>= xs"
by (auto simp add: postfix_def)
lemma postfix_trans: "[|xs >>= ys; ys >>= zs|] ==> xs >>= zs"
by (auto simp add: postfix_def)
lemma postfix_antisym: "[|xs >>= ys; ys >>= xs|] ==> xs = ys"
by (auto simp add: postfix_def)
lemma Nil_postfix [iff]: "xs >>= []"
by (simp add: postfix_def)
lemma postfix_Nil [simp]: "([] >>= xs) = (xs = [])"
by (auto simp add:postfix_def)
lemma postfix_ConsI: "xs >>= ys ==> x#xs >>= ys"
by (auto simp add: postfix_def)
lemma postfix_ConsD: "xs >>= y#ys ==> xs >>= ys"
by (auto simp add: postfix_def)
lemma postfix_appendI: "xs >>= ys ==> zs @ xs >>= ys"
by (auto simp add: postfix_def)
lemma postfix_appendD: "xs >>= zs @ ys ==> xs >>= ys"
by(auto simp add: postfix_def)
lemma postfix_is_subset_lemma: "xs = zs @ ys ==> set ys ⊆ set xs"
by (induct zs, auto)
lemma postfix_is_subset: "xs >>= ys ==> set ys ⊆ set xs"
by (unfold postfix_def, erule exE, erule postfix_is_subset_lemma)
lemma postfix_ConsD2_lemma [rule_format]: "x#xs = zs @ y#ys --> xs >>= ys"
by (induct zs, auto intro!: postfix_appendI postfix_ConsI)
lemma postfix_ConsD2: "x#xs >>= y#ys ==> xs >>= ys"
by (auto simp add: postfix_def dest!: postfix_ConsD2_lemma)
lemma postfix2prefix: "(xs >>= ys) = (rev ys <= rev xs)"
apply (unfold prefix_def postfix_def, safe)
apply (rule_tac x = "rev zs" in exI, simp)
apply (rule_tac x = "rev zs" in exI)
apply (rule rev_is_rev_conv [THEN iffD1], simp)
done
end
lemma prefixI:
ys = xs @ zs ==> xs ≤ ys
lemma prefixE:
xs ≤ ys ==> (!!zs. ys = xs @ zs ==> C) ==> C
lemma strict_prefixI':
ys = xs @ z # zs ==> xs < ys
lemma strict_prefixE':
xs < ys ==> (!!z zs. ys = xs @ z # zs ==> C) ==> C
lemma strict_prefixI:
xs ≤ ys ==> xs ≠ ys ==> xs < ys
lemma strict_prefixE:
xs < ys ==> (xs ≤ ys ==> xs ≠ ys ==> C) ==> C
theorem Nil_prefix:
[] ≤ xs
theorem prefix_Nil:
(xs ≤ []) = (xs = [])
lemma prefix_snoc:
(xs ≤ ys @ [y]) = (xs = ys @ [y] ∨ xs ≤ ys)
lemma Cons_prefix_Cons:
(x # xs ≤ y # ys) = (x = y ∧ xs ≤ ys)
lemma same_prefix_prefix:
(xs @ ys ≤ xs @ zs) = (ys ≤ zs)
lemma same_prefix_nil:
(xs @ ys ≤ xs) = (ys = [])
lemma prefix_prefix:
xs ≤ ys ==> xs ≤ ys @ zs
lemma append_prefixD:
xs @ ys ≤ zs ==> xs ≤ zs
theorem prefix_Cons:
(xs ≤ y # ys) = (xs = [] ∨ (∃zs. xs = y # zs ∧ zs ≤ ys))
theorem prefix_append:
(xs ≤ ys @ zs) = (xs ≤ ys ∨ (∃us. xs = ys @ us ∧ us ≤ zs))
lemma append_one_prefix:
xs ≤ ys ==> length xs < length ys ==> xs @ [ys ! length xs] ≤ ys
theorem prefix_length_le:
xs ≤ ys ==> length xs ≤ length ys
lemma prefix_same_cases:
xs1 ≤ ys ==> xs2 ≤ ys ==> xs1 ≤ xs2 ∨ xs2 ≤ xs1
lemma set_mono_prefix:
xs ≤ ys ==> set xs ⊆ set ys
lemma parallelI:
¬ xs ≤ ys ==> ¬ ys ≤ xs ==> xs \<parallel> ys
lemma parallelE:
xs \<parallel> ys ==> (¬ xs ≤ ys ==> ¬ ys ≤ xs ==> C) ==> C
theorem prefix_cases:
(xs ≤ ys ==> C) ==> (ys < xs ==> C) ==> (xs \<parallel> ys ==> C) ==> C
theorem parallel_decomp:
xs \<parallel> ys ==> ∃as b bs c cs. b ≠ c ∧ xs = as @ b # bs ∧ ys = as @ c # cs
lemma postfix_refl:
xs >>= xs
lemma postfix_trans:
xs >>= ys ==> ys >>= zs ==> xs >>= zs
lemma postfix_antisym:
xs >>= ys ==> ys >>= xs ==> xs = ys
lemma Nil_postfix:
xs >>= []
lemma postfix_Nil:
([] >>= xs) = (xs = [])
lemma postfix_ConsI:
xs >>= ys ==> x # xs >>= ys
lemma postfix_ConsD:
xs >>= y # ys ==> xs >>= ys
lemma postfix_appendI:
xs >>= ys ==> zs @ xs >>= ys
lemma postfix_appendD:
xs >>= zs @ ys ==> xs >>= ys
lemma postfix_is_subset_lemma:
xs = zs @ ys ==> set ys ⊆ set xs
lemma postfix_is_subset:
xs >>= ys ==> set ys ⊆ set xs
lemma postfix_ConsD2_lemma:
x # xs = zs @ y # ys ==> xs >>= ys
lemma postfix_ConsD2:
x # xs >>= y # ys ==> xs >>= ys
lemma postfix2prefix:
(xs >>= ys) = (rev ys ≤ rev xs)