diff --git a/thys/Card_Avoid.thy b/thys/Card_Avoid.thy deleted file mode 100644 index 30e8c3b4..00000000 --- a/thys/Card_Avoid.thy +++ /dev/null @@ -1,47 +0,0 @@ -theory Card_Avoid -imports - "Prelim.Prelim" -begin - -definition "suppGr f \ {(x, f x) | x. f x \ x}" - - -lemma supp_incl_suppGr: -"suppGr f \ suppGr g \ supp f \ supp g" -unfolding supp_def suppGr_def by auto - - -lemma extend_id_on': -assumes g: "bij g" "g o g = id" "id_on A g" -and A': "A \ A'" -shows "\f. bij f \ suppGr f \ suppGr g \ id_on A' f" -proof- - define f where "f \ - \a. if a \ (A'-A) \ g ` (A' - A) then a else g a" - show ?thesis apply(rule exI[of _ f], intro conjI) - subgoal using g(1) unfolding bij_def inj_def f_def apply safe - subgoal by (metis (mono_tags, opaque_lifting) Un_iff comp_def g(2) id_apply image_iff) - subgoal by auto - subgoal by (smt (verit) \\y x. \\x y. g x = g y \ x = y; surj g; - (if x \ A' - A \ g ` (A' - A) then x else g x) = (if y \ A' - A \ g ` (A' - A) then y else g y)\ - \ x = y\ g(2) image_iff pointfree_idE) . - subgoal unfolding suppGr_def f_def by auto - subgoal using g(3) unfolding id_on_def f_def by auto . -qed - - -lemma extend_id_on: -"bij g \ |supp (g::'a \ 'a)| \ \|A| \ - id_on (A - B1) g \ B2 \ B1 \ \ \ g ` B1 \ A = {}\ - \ \ added: \ g o g = id - \ - \f. bij f \ |supp f| suppGr f \ suppGr g \ id_on (A - B2) f" -apply(drule extend_id_on'[of g "A - B1" "A-B2"]) - subgoal . - subgoal by (meson card_of_diff ordLeq_ordLess_trans) - subgoal by auto - subgoal apply safe subgoal for f apply(rule exI[of _ f]) - subgoal using supp_incl_suppGr[of f g] by (meson card_of_mono1 ordLeq_ordLess_trans) . . . - - -end \ No newline at end of file diff --git a/thys/Prelim/Prelim.thy b/thys/Prelim/Prelim.thy index ac6aa5be..88ae7d0e 100644 --- a/thys/Prelim/Prelim.thy +++ b/thys/Prelim/Prelim.thy @@ -273,16 +273,16 @@ proof- have 0: "|AA| BB = {}" "AA \ CC = {}" "BB \ CC = {}" subgoal using AA bound card_of_diff ordLeq_ordLess_trans by blast - subgoal by (smt (verit, del_insts) AA AB BB DiffE DiffI Diff_Diff_Int bij_betw_iff_bijections + subgoal by (smt (verit, del_insts) AA AB BB DiffE DiffI Diff_Diff_Int bij_betw_iff_bijections eq_on_def id_apply int) subgoal unfolding defs by auto subgoal unfolding defs by auto subgoal unfolding defs by auto . show ?thesis using ex_bij_betw_supp[OF i 0] apply safe subgoal for u apply(rule exI[of _ u], safe) - subgoal by (smt (verit, ccfv_SIG) AA AB CC DiffI Diff_Diff_Int bij_betw_cong - eq_on_def id_apply imsupp_empty_IntD2 inf_commute int) - subgoal by (smt (verit, best) AA CC DiffI Diff_Diff_Int Int_commute eq_on_def + subgoal by (smt (verit, ccfv_SIG) AA AB CC DiffI Diff_Diff_Int bij_betw_cong + eq_on_def id_apply imsupp_empty_IntD2 inf_commute int) + subgoal by (smt (verit, best) AA CC DiffI Diff_Diff_Int Int_commute eq_on_def id_apply imsupp_empty_IntD2 int) . . qed @@ -780,7 +780,7 @@ lemma cmin1: shows "cmin r s \o r" unfolding cmin_def proof (cases "r o r" by (simp add: assms(1) csum_czero1 ordIso_imp_ordLeq) + then show "r +c czero \o r" by (simp add: assms(1) csum_czero1 ordIso_imp_ordLeq) next case False have "Well_order r" "Well_order s" using assms by auto @@ -794,10 +794,10 @@ lemma cmin2: shows "cmin r s \o s" unfolding cmin_def proof (cases "r o s" using assms(1) csum_czero1 ordIso_ordLeq_trans ordLess_imp_ordLeq by blast + then show "r +c czero \o s" using assms(1) csum_czero1 ordIso_ordLeq_trans ordLess_imp_ordLeq by blast next case False - then show "czero +c s \o s" using ordIso_ordLeq_trans using assms(2) csum_czero2 ordIso_imp_ordLeq by blast + then show "czero +c s \o s" using ordIso_ordLeq_trans using assms(2) csum_czero2 ordIso_imp_ordLeq by blast qed lemma cmin_greater: @@ -841,7 +841,7 @@ qed (* *) -definition natOf :: "nat list \ nat" where +definition natOf :: "nat list \ nat" where "natOf \ SOME f . inj f" lemma inject_natOf: "inj natOf" @@ -850,43 +850,96 @@ by (metis ex_inj natOf_def someI_ex) lemma inj_natOf[simp]: "natOf p = natOf p \ p = p" by (meson injD inject_natOf) -(* *) - lemma ex_push_inwards: "(\x. P x \ (\y. Q x y)) \ (\y. \x. Q x y \ P x)" by blast lemma ex_conjunct2: "\x. P x \ Q x \ \x. Q x" by auto -lemma eextend_fresh: +lemma eextend_fresh: fixes A A' B::"'a set" assumes "|B| A" "B \ A' = {}" -shows "\\. bij \ \ |supp \| \ ` B \ A = {} \ id_on A' \" +shows "\\. bij \ \ |supp \| \ ` B \ A = {} \ id_on A' \ \ \ \ \ = id" proof- have "|- (A \ B)| =o |UNIV::'a set|" using card_of_Un_diff_infinite[OF assms(3), unfolded Compl_eq_Diff_UNIV[symmetric]] assms(1) assms(2) assms(3) card_of_Un_ordLess_infinite by blast - hence "|B| B)|" + hence "|B| B)|" using assms(1) ordIso_symmetric ordLess_ordIso_trans by blast - then obtain f where f: "inj_on f B" "f ` B \ - (A \ B)" + then obtain f where f: "inj_on f B" "f ` B \ - (A \ B)" by (meson card_of_ordLeq ordLeq_iff_ordLess_or_ordIso) - define g where "g \ \a. if a \ B then f a else a" - have g: "inj_on g (B \ A')" "g ` (B \ A') \ - (A \ B) \ A'" using f - unfolding g_def inj_on_def using assms(3,4) by auto - define C where C: "C = g ` (B \ A')" - have b: "bij_betw g (B \ A') C" unfolding C bij_betw_def using g by simp - - have 0: "Cinfinite |UNIV::'a set|" "|B \ A'| A') \ C) g id" - subgoal by (simp add: assms(3) cinfinite_iff_infinite) - subgoal by (meson assms(1-4) card_of_Un_ordLess_infinite card_of_ordLess order.trans) - subgoal using assms(3) f unfolding eq_on_def C g_def by auto . - - show ?thesis using ex_bij_betw_supp'[OF 0(1,2) b 0(3)] apply safe - subgoal for \ apply(rule exI[of _ \]) - unfolding id_on_def apply auto - apply (metis ComplD UnCI eq_on_def f(2) g_def image_subset_iff) - by (metis Int_iff Un_Int_eq(1) assms(5) empty_iff eq_on_def g_def sup_commute) . + + then have f1: "f ` B \ B = {}" "f ` B \ A = {}" by blast+ + + define g where "g \ \a. if a \ B then f a else (if a \ f ` B then inv_into B f a else a)" + have 1: "g \ g = id" + unfolding g_def comp_def id_def + apply (rule ext) + subgoal for x + apply (cases "x \ B") + apply (subst if_P, assumption)+ + apply (subst if_not_P) + using f1 apply blast + apply (subst if_P) + apply (erule imageI) + apply (rule inv_into_f_f) + apply (rule f) + apply assumption + apply (subst if_not_P, assumption)+ + apply (cases "x \ f ` B") + apply (subst if_P, assumption)+ + apply (subst if_P) + apply (erule inv_into_into) + apply (erule f_inv_into_f) + apply (subst if_not_P, assumption)+ + apply (rule refl) + done + done + then have 2: "bij g" using o_bij by blast + + have 3: "id_on A' g" unfolding id_on_def g_def using f1 assms(4,5) by auto + have 4: "g ` B \ A = {}" unfolding g_def using f1 by simp + + have "supp g \ B \ g ` B" unfolding g_def supp_def by auto + then have 5: "|supp g| {(x, f x) | x. f x \ x}" + +lemma supp_incl_suppGr: + "suppGr f \ suppGr g \ supp f \ supp g" + unfolding supp_def suppGr_def by auto + +lemma extend_id_on': +assumes g: "bij g" "g o g = id" "id_on A g" +and A': "A \ A'" +shows "\f. bij f \ suppGr f \ suppGr g \ id_on A' f" +proof- + define f where "f \ \a. if a \ (A'-A) \ g ` (A' - A) then a else g a" + show ?thesis apply(rule exI[of _ f], intro conjI) + subgoal using g(1) unfolding bij_def inj_def f_def apply safe + subgoal by (metis (mono_tags, opaque_lifting) Un_iff comp_def g(2) id_apply image_iff) + subgoal by auto + subgoal by (smt (verit) \\y x. \\x y. g x = g y \ x = y; surj g; + (if x \ A' - A \ g ` (A' - A) then x else g x) = (if y \ A' - A \ g ` (A' - A) then y else g y)\ + \ x = y\ g(2) image_iff pointfree_idE) . + subgoal unfolding suppGr_def f_def by auto + subgoal using g(3) unfolding id_on_def f_def by auto . +qed + +lemma extend_id_on: + "bij g \ |supp (g::'a \ 'a)| + id_on (A - B1) g \ B2 \ B1 \ g \ g = id \ + \f. bij f \ |supp f| suppGr f \ suppGr g \ id_on (A - B2) f" +apply(drule extend_id_on'[of g "A - B1" "A-B2"]) + subgoal . + subgoal by (meson card_of_diff ordLeq_ordLess_trans) + subgoal by auto + subgoal apply safe subgoal for f apply(rule exI[of _ f]) + subgoal using supp_incl_suppGr[of f g] by (meson card_of_mono1 ordLeq_ordLess_trans) . . . + end \ No newline at end of file