# HG changeset patch # User nipkow # Date 1513152283 -3600 # Node ID bea1258d9a27fc803fcee911dbd0e36082bdc34b # Parent 88d1c9d86f4853bfb2cabe6da3553ea720572da6 added lemmas diff -r 88d1c9d86f48 -r bea1258d9a27 src/HOL/Data_Structures/List_Ins_Del.thy --- a/src/HOL/Data_Structures/List_Ins_Del.thy Tue Dec 12 10:01:14 2017 +0100 +++ b/src/HOL/Data_Structures/List_Ins_Del.thy Wed Dec 13 09:04:43 2017 +0100 @@ -53,12 +53,12 @@ by(induction xs) auto lemma distinct_if_sorted: "sorted xs \ distinct xs" -apply(induction xs rule: sorted_wrt_induct) +apply(induction xs rule: induct_list012) apply auto by (metis in_set_conv_decomp_first less_imp_not_less sorted_mid_iff2) lemma sorted_ins_list: "sorted xs \ sorted(ins_list x xs)" -by(induction xs rule: sorted_wrt_induct) auto +by(induction xs rule: induct_list012) auto lemma ins_list_sorted: "sorted (xs @ [a]) \ ins_list x (xs @ a # ys) = @@ -105,7 +105,7 @@ done lemma sorted_del_list: "sorted xs \ sorted(del_list x xs)" -apply(induction xs rule: sorted_wrt_induct) +apply(induction xs rule: induct_list012) apply auto by (meson order.strict_trans sorted_Cons_iff) diff -r 88d1c9d86f48 -r bea1258d9a27 src/HOL/List.thy --- a/src/HOL/List.thy Tue Dec 12 10:01:14 2017 +0100 +++ b/src/HOL/List.thy Wed Dec 13 09:04:43 2017 +0100 @@ -793,28 +793,13 @@ "(\xs. \ys. length ys < length xs \ P ys \ P xs) \ P xs" by (fact measure_induct) +lemma induct_list012: + "\P []; \x. P [x]; \x y zs. P (y # zs) \ P (x # y # zs)\ \ P xs" +by induction_schema (pat_completeness, lexicographic_order) + lemma list_nonempty_induct [consumes 1, case_names single cons]: - assumes "xs \ []" - assumes single: "\x. P [x]" - assumes cons: "\x xs. xs \ [] \ P xs \ P (x # xs)" - shows "P xs" -using \xs \ []\ proof (induct xs) - case Nil then show ?case by simp -next - case (Cons x xs) - show ?case - proof (cases xs) - case Nil - with single show ?thesis by simp - next - case Cons - show ?thesis - proof (rule cons) - from Cons show "xs \ []" by simp - with Cons.hyps show "P xs" . - qed - qed -qed + "\ xs \ []; \x. P [x]; \x xs. xs \ [] \ P xs \ P (x # xs)\ \ P xs" +by(induction xs rule: induct_list012) auto lemma inj_split_Cons: "inj_on (\(xs, n). n#xs) X" by (auto intro!: inj_onI) @@ -1986,7 +1971,7 @@ by(induct xs)(auto simp:neq_Nil_conv) lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" -by (induct xs, simp, case_tac xs, simp_all) +by (induction xs rule: induct_list012) simp_all lemma last_list_update: "xs \ [] \ last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)" @@ -2430,11 +2415,7 @@ lemma takeWhile_not_last: "distinct xs \ takeWhile (\y. y \ last xs) xs = butlast xs" -apply(induct xs) - apply simp -apply(case_tac xs) -apply(auto) -done +by(induction xs rule: induct_list012) auto lemma takeWhile_cong [fundef_cong]: "[| l = k; !!x. x : set l ==> P x = Q x |] @@ -3229,6 +3210,9 @@ lemma upto_Nil[simp]: "[i..j] = [] \ j < i" by (simp add: upto.simps) +lemma upto_Nil2[simp]: "[] = [i..j] \ j < i" +by (simp add: upto.simps) + lemma upto_rec1: "i \ j \ [i..j] = i#[i+1..j]" by(simp add: upto.simps) @@ -4910,10 +4894,6 @@ subsubsection \@{const sorted_wrt}\ -lemma sorted_wrt_induct: - "\P []; \x. P [x]; \x y zs. P (y # zs) \ P (x # y # zs)\ \ P xs" -by induction_schema (pat_completeness, lexicographic_order) - lemma sorted_wrt_Cons: assumes "transp P" shows "sorted_wrt P (x # xs) = ((\y \ set xs. P x y) \ sorted_wrt P xs)" @@ -4921,7 +4901,7 @@ lemma sorted_wrt_ConsI: "\ \y. y \ set xs \ P x y; sorted_wrt P xs \ \ sorted_wrt P (x # xs)" -by (induction xs rule: sorted_wrt_induct) simp_all +by (induction xs rule: induct_list012) simp_all lemma sorted_wrt_append: assumes "transp P" @@ -4931,15 +4911,15 @@ lemma sorted_wrt_backwards: "sorted_wrt P (xs @ [y, z]) = (P y z \ sorted_wrt P (xs @ [y]))" -by (induction xs rule: sorted_wrt_induct) auto +by (induction xs rule: induct_list012) auto lemma sorted_wrt_rev: "sorted_wrt P (rev xs) = sorted_wrt (\x y. P y x) xs" -by (induction xs rule: sorted_wrt_induct) (simp_all add: sorted_wrt_backwards) +by (induction xs rule: induct_list012) (simp_all add: sorted_wrt_backwards) lemma sorted_wrt_mono: "(\x y. P x y \ Q x y) \ sorted_wrt P xs \ sorted_wrt Q xs" -by(induction xs rule: sorted_wrt_induct)(auto) +by(induction xs rule: induct_list012)(auto) text \Strictly Ascending Sequences of Natural Numbers\ @@ -4976,7 +4956,7 @@ proof (induct xs rule: sorted.induct) case (Cons xs x) thus ?case by (cases xs) simp_all qed simp -qed (induct xs rule: sorted_wrt_induct, simp_all) +qed (induct xs rule: induct_list012, simp_all) lemma sorted_tl: "sorted xs \ sorted (tl xs)"