# HG changeset patch # User berghofe # Date 1048582731 -3600 # Node ID 0451e0fb3f220911b9030d51973a5b6d089c6941 # Parent 2266550ab31627bafbe1b3284fbaba4fdd0adc83 Re-structured some proofs in order to get rid of rule_format attribute. diff -r 2266550ab316 -r 0451e0fb3f22 src/HOL/List.thy --- a/src/HOL/List.thy Tue Mar 25 09:56:01 2003 +0100 +++ b/src/HOL/List.thy Tue Mar 25 09:58:51 2003 +0100 @@ -225,11 +225,9 @@ lemma lists_mono [mono]: "A \ B ==> lists A \ lists B" by (unfold lists.defs) (blast intro!: lfp_mono) -lemma lists_IntI [rule_format]: -"l: lists A ==> l: lists B --> l: lists (A Int B)" -apply (erule lists.induct) -apply blast+ -done +lemma lists_IntI: + assumes l: "l: lists A" shows "l: lists B ==> l: lists (A Int B)" using l + by induct blast+ lemma lists_Int_eq [simp]: "lists (A \ B) = lists A \ lists B" apply (rule mono_Int [THEN equalityI]) @@ -292,15 +290,13 @@ lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])" by (induct xs) auto -lemma append_eq_append_conv [rule_format, simp]: - "\ys. length xs = length ys \ length us = length vs - --> (xs@us = ys@vs) = (xs=ys \ us=vs)" -apply (induct_tac xs) - apply(rule allI) +lemma append_eq_append_conv [simp]: + "!!ys. length xs = length ys \ length us = length vs + ==> (xs@us = ys@vs) = (xs=ys \ us=vs)" +apply (induct xs) apply (case_tac ys) apply simp apply force -apply (rule allI) apply (case_tac ys) apply force apply simp @@ -1041,6 +1037,20 @@ (length xs = length ys \ (\i < length xs. P (xs!i) (ys!i)))" by (force simp add: list_all2_def set_zip) +lemma list_all2_trans: + assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c" + shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs" + (is "!!bs cs. PROP ?Q as bs cs") +proof (induct as) + fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs" + show "!!cs. PROP ?Q (x # xs) bs cs" + proof (induct bs) + fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs" + show "PROP ?Q (x # xs) (y # ys) cs" + by (induct cs) (auto intro: tr I1 I2) + qed simp +qed simp + 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) @@ -1057,19 +1067,6 @@ "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" -apply(induct_tac as) - apply simp -apply(rule allI) -apply(induct_tac bs) - apply simp -apply(rule allI) -apply(induct_tac cs) - apply auto -done - lemma list_all2_refl: "(\x. P x x) \ list_all2 P xs xs" by (simp add: list_all2_conv_all_nth) @@ -1173,11 +1170,10 @@ apply (auto simp add: less_diff_conv nth_upt) done -lemma nth_take_lemma [rule_format]: -"ALL xs ys. k <= length xs --> k <= length ys ---> (ALL i. i < k --> xs!i = ys!i) ---> take k xs = take k ys" -apply (induct k) +lemma nth_take_lemma: + "!!xs ys. k <= length xs ==> k <= length ys ==> + (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys" +apply (atomize, induct k) apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib) apply clarify txt {* Both lists must be non-empty *}