# HG changeset patch # User paulson # Date 1447868220 0 # Node ID 286c1741285c190116311a20e820622f8bc141ba # Parent a81dc5c4d6a917067cb020b7b3bd4b42d972246d# Parent c4a6edbfec49395b44f5be069fbadf9fc41fcd7b Merge diff -r c4a6edbfec49 -r 286c1741285c src/HOL/Fun.thy --- a/src/HOL/Fun.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/Fun.thy Wed Nov 18 17:37:00 2015 +0000 @@ -120,6 +120,9 @@ lemma fcomp_id [simp]: "f \> id = f" by (simp add: fcomp_def) +lemma fcomp_comp: "fcomp f g = comp g f" + by (simp add: ext) + code_printing constant fcomp \ (Eval) infixl 1 "#>" diff -r c4a6edbfec49 -r 286c1741285c src/HOL/Library/Permutation.thy --- a/src/HOL/Library/Permutation.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/Library/Permutation.thy Wed Nov 18 17:37:00 2015 +0000 @@ -15,25 +15,25 @@ | Cons [intro!]: "xs <~~> ys \ z # xs <~~> z # ys" | trans [intro]: "xs <~~> ys \ ys <~~> zs \ xs <~~> zs" -lemma perm_refl [iff]: "l <~~> l" +proposition perm_refl [iff]: "l <~~> l" by (induct l) auto subsection \Some examples of rule induction on permutations\ -lemma xperm_empty_imp: "[] <~~> ys \ ys = []" +proposition xperm_empty_imp: "[] <~~> ys \ ys = []" by (induct xs == "[] :: 'a list" ys pred: perm) simp_all text \\medskip This more general theorem is easier to understand!\ -lemma perm_length: "xs <~~> ys \ length xs = length ys" +proposition perm_length: "xs <~~> ys \ length xs = length ys" by (induct pred: perm) simp_all -lemma perm_empty_imp: "[] <~~> xs \ xs = []" +proposition perm_empty_imp: "[] <~~> xs \ xs = []" by (drule perm_length) auto -lemma perm_sym: "xs <~~> ys \ ys <~~> xs" +proposition perm_sym: "xs <~~> ys \ ys <~~> xs" by (induct pred: perm) auto @@ -41,78 +41,72 @@ text \We can insert the head anywhere in the list.\ -lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" +proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" by (induct xs) auto -lemma perm_append_swap: "xs @ ys <~~> ys @ xs" - apply (induct xs) - apply simp_all - apply (blast intro: perm_append_Cons) - done +proposition perm_append_swap: "xs @ ys <~~> ys @ xs" + by (induct xs) (auto intro: perm_append_Cons) -lemma perm_append_single: "a # xs <~~> xs @ [a]" +proposition perm_append_single: "a # xs <~~> xs @ [a]" by (rule perm.trans [OF _ perm_append_swap]) simp -lemma perm_rev: "rev xs <~~> xs" - apply (induct xs) - apply simp_all - apply (blast intro!: perm_append_single intro: perm_sym) - done +proposition perm_rev: "rev xs <~~> xs" + by (induct xs) (auto intro!: perm_append_single intro: perm_sym) -lemma perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys" +proposition perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys" by (induct l) auto -lemma perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l" +proposition perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l" by (blast intro!: perm_append_swap perm_append1) subsection \Further results\ -lemma perm_empty [iff]: "[] <~~> xs \ xs = []" +proposition perm_empty [iff]: "[] <~~> xs \ xs = []" by (blast intro: perm_empty_imp) -lemma perm_empty2 [iff]: "xs <~~> [] \ xs = []" +proposition perm_empty2 [iff]: "xs <~~> [] \ xs = []" apply auto apply (erule perm_sym [THEN perm_empty_imp]) done -lemma perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]" +proposition perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]" by (induct pred: perm) auto -lemma perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]" +proposition perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]" by (blast intro: perm_sing_imp) -lemma perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]" +proposition perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]" by (blast dest: perm_sym) subsection \Removing elements\ -lemma perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys" +proposition perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys" by (induct ys) auto text \\medskip Congruence rule\ -lemma perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys" +proposition perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys" by (induct pred: perm) auto -lemma remove_hd [simp]: "remove1 z (z # xs) = xs" +proposition remove_hd [simp]: "remove1 z (z # xs) = xs" by auto -lemma cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys" +proposition cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys" by (drule_tac z = z in perm_remove_perm) auto -lemma cons_perm_eq [iff]: "z#xs <~~> z#ys \ xs <~~> ys" +proposition cons_perm_eq [iff]: "z#xs <~~> z#ys \ xs <~~> ys" by (blast intro: cons_perm_imp_perm) -lemma append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys" +proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys" by (induct zs arbitrary: xs ys rule: rev_induct) auto -lemma perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys" +proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys" by (blast intro: append_perm_imp_perm perm_append1) -lemma perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys" +proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys" apply (safe intro!: perm_append2) apply (rule append_perm_imp_perm) apply (rule perm_append_swap [THEN perm.trans]) @@ -120,7 +114,7 @@ apply (blast intro: perm_append_swap) done -lemma mset_eq_perm: "mset xs = mset ys \ xs <~~> ys" +theorem mset_eq_perm: "mset xs = mset ys \ xs <~~> ys" apply (rule iffI) apply (erule_tac [2] perm.induct) apply (simp_all add: union_ac) @@ -140,24 +134,24 @@ apply simp done -lemma mset_le_perm_append: "mset xs \# mset ys \ (\zs. xs @ zs <~~> ys)" +proposition mset_le_perm_append: "mset xs \# mset ys \ (\zs. xs @ zs <~~> ys)" apply (auto simp: mset_eq_perm[THEN sym] mset_le_exists_conv) apply (insert surj_mset) apply (drule surjD) apply (blast intro: sym)+ done -lemma perm_set_eq: "xs <~~> ys \ set xs = set ys" +proposition perm_set_eq: "xs <~~> ys \ set xs = set ys" by (metis mset_eq_perm mset_eq_setD) -lemma perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys" +proposition perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys" apply (induct pred: perm) apply simp_all apply fastforce apply (metis perm_set_eq) done -lemma eq_set_perm_remdups: "set xs = set ys \ remdups xs <~~> remdups ys" +theorem eq_set_perm_remdups: "set xs = set ys \ remdups xs <~~> remdups ys" apply (induct xs arbitrary: ys rule: length_induct) apply (case_tac "remdups xs") apply simp_all @@ -182,10 +176,10 @@ apply (rule length_remdups_leq) done -lemma perm_remdups_iff_eq_set: "remdups x <~~> remdups y \ set x = set y" +proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \ set x = set y" by (metis List.set_remdups perm_set_eq eq_set_perm_remdups) -lemma permutation_Ex_bij: +theorem permutation_Ex_bij: assumes "xs <~~> ys" shows "\f. bij_betw f {.. (\i A}" +proof (rule finite_subset[where B="{xs. set xs \ set A \ length xs \ length A}"]) + show "finite {xs. set xs \ set A \ length xs \ length A}" + apply (cases A, simp) + apply (rule card_ge_0_finite) + apply (auto simp: card_lists_length_le) + done +next + show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}" + by (clarsimp simp add: perm_length perm_set_eq) +qed + +proposition perm_swap: + assumes "i < length xs" "j < length xs" + shows "xs[i := xs ! j, j := xs ! i] <~~> xs" + using assms by (simp add: mset_eq_perm[symmetric] mset_swap) + end diff -r c4a6edbfec49 -r 286c1741285c src/HOL/List.thy --- a/src/HOL/List.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/List.thy Wed Nov 18 17:37:00 2015 +0000 @@ -2081,6 +2081,12 @@ apply (case_tac i, auto) done +lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)" + by (cases "length xs < n") (auto simp: rev_take) + +lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)" + by (cases "length xs < n") (auto simp: rev_drop) + lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i" apply (induct xs arbitrary: i n, auto) apply (case_tac n, blast) diff -r c4a6edbfec49 -r 286c1741285c src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Wed Nov 18 17:37:00 2015 +0000 @@ -2695,6 +2695,9 @@ shows "homotopic_with P X Y f g \ homotopic_with P X Y g f" using homotopic_with_symD by blast +lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" + by force + lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" by force diff -r c4a6edbfec49 -r 286c1741285c src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Nov 18 17:37:00 2015 +0000 @@ -6882,7 +6882,6 @@ fixes s :: "real set" assumes "continuous_on {t \ s. t \ a} f" and "continuous_on {t \ s. a \ t} g" - and "continuous_on s h" and "a \ s \ f a = g a" shows "continuous_on s (\t. if t \ a then f(t) else g(t))" using assms diff -r c4a6edbfec49 -r 286c1741285c src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Wed Nov 18 14:28:45 2015 +0100 +++ b/src/HOL/Orderings.thy Wed Nov 18 17:37:00 2015 +0000 @@ -1379,8 +1379,14 @@ fix x assume "P x" "\y. P y \ x \ y" thus "Q x" by (rule assms(2)) qed +lemma LeastI2_wellorder_ex: + assumes "\x. P x" + and "\a. \ P a; \b. P b \ a \ b \ \ Q a" + shows "Q (Least P)" +using assms by clarify (blast intro!: LeastI2_wellorder) + lemma not_less_Least: "k < (LEAST x. P x) \ \ P k" -apply (simp (no_asm_use) add: not_le [symmetric]) +apply (simp add: not_le [symmetric]) apply (erule contrapos_nn) apply (erule Least_le) done