# HG changeset patch # User kleing # Date 1047639803 -3600 # Node ID ec901a432ea10c532859b87a9d57296cb13118ec # Parent 7cbc89aa79db633120a0560ce5399d50ef94265f more about list_all2 diff -r 7cbc89aa79db -r ec901a432ea1 src/HOL/List.thy --- a/src/HOL/List.thy Fri Mar 14 12:02:14 2003 +0100 +++ b/src/HOL/List.thy Fri Mar 14 12:03:23 2003 +0100 @@ -993,6 +993,10 @@ "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys" by (simp add: list_all2_def zip_rev cong: conj_cong) +lemma list_all2_rev1: +"list_all2 P (rev xs) ys = list_all2 P xs (rev ys)" +by (subst list_all2_rev [symmetric]) simp + lemma list_all2_append1: "list_all2 P (xs @ ys) zs = (EX us vs. zs = us @ vs \ length us = length xs \ length vs = length ys \ @@ -1019,11 +1023,40 @@ apply (simp add: ball_Un) done +lemma list_all2_append: + "\b. length a = length b \ + list_all2 P (a@c) (b@d) = (list_all2 P a b \ list_all2 P c d)" + apply (induct a) + apply simp + apply (case_tac b) + apply auto + done + +lemma list_all2_appendI [intro?, trans]: + "\ list_all2 P a b; list_all2 P c d \ \ list_all2 P (a@c) (b@d)" + by (simp add: list_all2_append list_all2_lengthD) + lemma list_all2_conv_all_nth: "list_all2 P xs ys = (length xs = length ys \ (\i < length xs. P (xs!i) (ys!i)))" by (force simp add: list_all2_def set_zip) +lemma list_all2_all_nthI [intro?]: + "length a = length b \ (\n. n < length a \ P (a!n) (b!n)) \ list_all2 P a b" + by (simp add: list_all2_conv_all_nth) + +lemma list_all2_nthD [dest?]: + "\ list_all2 P xs ys; p < size xs \ \ P (xs!p) (ys!p)" + by (simp add: list_all2_conv_all_nth) + +lemma list_all2_map1: + "list_all2 P (map f as) bs = list_all2 (\x y. P (f x) y) as bs" + by (simp add: list_all2_conv_all_nth) + +lemma list_all2_map2: + "list_all2 P as (map f bs) = list_all2 (\x y. P x (f y)) as bs" + by (auto simp add: list_all2_conv_all_nth) + lemma list_all2_trans[rule_format]: "\a b c. P1 a b --> P2 b c --> P3 a c ==> \bs cs. list_all2 P1 as bs --> list_all2 P2 bs cs --> list_all2 P3 as cs" @@ -1037,6 +1070,36 @@ apply auto done +lemma list_all2_refl: + "(\x. P x x) \ list_all2 P xs xs" + by (simp add: list_all2_conv_all_nth) + +lemma list_all2_update_cong: + "\ i \ list_all2 P (xs[i:=x]) (ys[i:=y])" + by (simp add: list_all2_conv_all_nth nth_list_update) + +lemma list_all2_update_cong2: + "\list_all2 P xs ys; P x y; i < length ys\ \ list_all2 P (xs[i:=x]) (ys[i:=y])" + by (simp add: list_all2_lengthD list_all2_update_cong) + +lemma list_all2_dropI [intro?]: + "\n bs. list_all2 P as bs \ list_all2 P (drop n as) (drop n bs)" + apply (induct as) + apply simp + apply (clarsimp simp add: list_all2_Cons1) + apply (case_tac n) + apply simp + apply simp + done + +lemma list_all2_mono [intro?]: + "\y. list_all2 P x y \ (\x y. P x y \ Q x y) \ list_all2 Q x y" + apply (induct x) + apply simp + apply (case_tac y) + apply auto + done + subsection {* @{text foldl} *} @@ -1135,6 +1198,16 @@ apply (simp_all add: take_all) done +(* needs nth_equalityI *) +lemma list_all2_antisym: + "\ (\x y. \P x y; Q y x\ \ x = y); list_all2 P xs ys; list_all2 Q ys xs \ + \ xs = ys" + apply (simp add: list_all2_conv_all_nth) + apply (rule nth_equalityI) + apply blast + apply simp + done + lemma take_equalityI: "(\i. take i xs = take i ys) ==> xs = ys" -- {* The famous take-lemma. *} apply (drule_tac x = "max (length xs) (length ys)" in spec)