diff --git a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy --- a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy +++ b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy @@ -1,1114 +1,1114 @@ section \\isaheader{Implementing Unique Priority Queues by Annotated Lists}\ theory PrioUniqueByAnnotatedList imports "../spec/AnnotatedListSpec" "../spec/PrioUniqueSpec" begin text \ In this theory we use annotated lists to implement unique priority queues with totally ordered elements. This theory is written as a generic adapter from the AnnotatedList interface to the unique priority queue interface. The annotated list stores a sequence of elements annotated with priorities\footnote{Technically, the annotated list elements are of unit-type, and the annotations hold both, the priority queue elements and the priorities. This is required as we defined annotated lists to only sum up the elements annotations.} The monoids operations forms the maximum over the elements and the minimum over the priorities. The sequence of pairs is ordered by ascending elements' order. The insertion point for a new element, or the priority of an existing element can be found by splitting the sequence at the point where the maximum of the elements read so far gets bigger than the element to be inserted. The minimum priority can be read out as the sum over the whole sequence. Finding the element with minimum priority is done by splitting the sequence at the point where the minimum priority of the elements read so far becomes equal to the minimum priority of the whole sequence. \ subsection "Definitions" subsubsection "Monoid" datatype ('e, 'a) LP = Infty | LP 'e 'a fun p_unwrap :: "('e,'a) LP \ ('e \ 'a)" where "p_unwrap (LP e a) = (e , a)" fun p_min :: "('e::linorder, 'a::linorder) LP \ ('e, 'a) LP \ ('e, 'a) LP" where "p_min Infty Infty = Infty"| "p_min Infty (LP e a) = LP e a"| "p_min (LP e a) Infty = LP e a"| "p_min (LP e1 a) (LP e2 b) = (LP (max e1 e2) (min a b))" fun e_less_eq :: "'e \ ('e::linorder, 'a::linorder) LP \ bool" where "e_less_eq e Infty = False"| "e_less_eq e (LP e' _) = (e \ e')" text_raw\\paragraph{Instantiation of classes}\ \\\ lemma p_min_re_neut[simp]: "p_min a Infty = a" by (induct a) auto lemma p_min_le_neut[simp]: "p_min Infty a = a" by (induct a) auto lemma p_min_asso: "p_min (p_min a b) c = p_min a (p_min b c)" apply(induct a b rule: p_min.induct ) apply (auto) apply (induct c) apply (auto) apply (metis max.assoc) apply (metis min.assoc) done lemma lp_mono: "class.monoid_add p_min Infty" by unfold_locales (auto simp add: p_min_asso) instantiation LP :: (linorder,linorder) monoid_add begin definition zero_def: "0 == Infty" definition plus_def: "a+b == p_min a b" instance by intro_classes (auto simp add: p_min_asso zero_def plus_def) end fun p_less_eq :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less_eq (LP e a) (LP f b) = (a \ b)"| "p_less_eq _ Infty = True"| "p_less_eq Infty (LP e a) = False" fun p_less :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less (LP e a) (LP f b) = (a < b)"| "p_less (LP e a) Infty = True"| "p_less Infty _ = False" lemma p_less_le_not_le : "p_less x y \ p_less_eq x y \ \ (p_less_eq y x)" by (induct x y rule: p_less.induct) auto lemma p_order_refl : "p_less_eq x x" by (induct x) auto lemma p_le_inf : "p_less_eq Infty x \ x = Infty" by (induct x) auto lemma p_order_trans : "\p_less_eq x y; p_less_eq y z\ \ p_less_eq x z" apply (induct y z rule: p_less.induct) apply auto apply (induct x) apply auto apply (cases x) apply auto apply(induct x) apply (auto simp add: p_le_inf) apply (metis p_le_inf p_less_eq.simps(2)) done lemma p_linear2 : "p_less_eq x y \ p_less_eq y x" apply (induct x y rule: p_less_eq.induct) apply auto done instantiation LP :: (type, linorder) preorder begin definition plesseq_def: "less_eq = p_less_eq" definition pless_def: "less = p_less" instance apply (intro_classes) apply (simp only: p_less_le_not_le pless_def plesseq_def) apply (simp only: p_order_refl plesseq_def pless_def) apply (simp only: plesseq_def) apply (metis p_order_trans) done end subsubsection "Operations" definition aluprio_\ :: "('s \ (unit \ ('e::linorder,'a::linorder) LP) list) \ 's \ ('e::linorder \ 'a::linorder)" where "aluprio_\ \ ft == (map_of (map p_unwrap (map snd (\ ft))))" definition aluprio_invar :: "('s \ (unit \ ('c::linorder, 'd::linorder) LP) list) \ ('s \ bool) \ 's \ bool" where "aluprio_invar \ invar ft == invar ft \ (\ x\set (\ ft). snd x\Infty) \ sorted (map fst (map p_unwrap (map snd (\ ft)))) \ distinct (map fst (map p_unwrap (map snd (\ ft)))) " definition aluprio_empty where "aluprio_empty empt = empt" definition aluprio_isEmpty where "aluprio_isEmpty isEmpty = isEmpty" definition aluprio_insert :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ ('s \ 's \ 's) \ ('s \ unit \ ('e,'a) LP \ 's) \ 's \ 'e \ 'a \ 's" where " aluprio_insert splits annot isEmpty app consr s e a = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e < fst (p_unwrap lp) then app (consr (consr l () (LP e a)) () lp) r else app (consr l () (LP e a)) r )) else consr s () (LP e a)) " definition aluprio_pop :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ 's \ 's) \ 's \ 'e \'a \'s" where "aluprio_pop splits annot app s = (let (l, (_,lp) , r) = splits (\ x. x \ (annot s)) Infty s in (case lp of (LP e a) \ (e, a, app l r) ))" definition aluprio_prio :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ 's \ 'e \ 'a option" where " aluprio_prio splits annot isEmpty s e = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e = fst (p_unwrap lp) then Some (snd (p_unwrap lp)) else None)) else None) " lemmas aluprio_defs = aluprio_invar_def aluprio_\_def aluprio_empty_def aluprio_isEmpty_def aluprio_insert_def aluprio_pop_def aluprio_prio_def subsection "Correctness" subsubsection "Auxiliary Lemmas" lemma p_linear: "(x::('e, 'a::linorder) LP) \ y \ y \ x" by (unfold plesseq_def) (simp only: p_linear2) lemma e_less_eq_mon1: "e_less_eq e x \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI1) done lemma e_less_eq_mon2: "e_less_eq e y \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI2) done lemmas e_less_eq_mon = e_less_eq_mon1 e_less_eq_mon2 lemma p_less_eq_mon: "(x::('e::linorder,'a::linorder) LP) \ z \ (x + y) \ z" apply(cases y) apply(auto simp add: plus_def) apply (cases x) apply (cases z) apply (auto simp add: plesseq_def) apply (cases z) apply (auto simp add: min.coboundedI1) done lemma p_less_eq_lem1: "\\ (x::('e::linorder,'a::linorder) LP) \ z; (x + y) \ z\ \ y \ z " apply (cases x,auto simp add: plus_def) apply (cases y, auto) apply (cases z, auto simp add: plesseq_def) apply (metis min_le_iff_disj) done lemma infadd: "x \ Infty \x + y \ Infty" apply (unfold plus_def) apply (induct x y rule: p_min.induct) apply auto done lemma e_less_eq_sum_list: "\\ e_less_eq e (sum_list xs)\ \ \x \ set xs. \ e_less_eq e x" proof (induct xs) case Nil thus ?case by simp next case (Cons a xs) hence "\ e_less_eq e (sum_list xs)" by (auto simp add: e_less_eq_mon) hence v1: "\x\set xs. \ e_less_eq e x" using Cons.hyps by simp from Cons.prems have "\ e_less_eq e a" by (auto simp add: e_less_eq_mon) with v1 show "\x\set (a#xs). \ e_less_eq e x" by simp qed lemma e_less_eq_p_unwrap: "\x \ Infty;\ e_less_eq e x\ \ fst (p_unwrap x) < e" by (cases x) auto lemma e_less_eq_refl : "b \ Infty \ e_less_eq (fst (p_unwrap b)) b" by (cases b) auto lemma e_less_eq_sum_list2: assumes "\x\set (\s). snd x \ Infty" "((), b) \ set (\s)" shows "e_less_eq (fst (p_unwrap b)) (sum_list (map snd (\s)))" apply(insert assms) apply (induct "\s") apply (auto simp add: zero_def e_less_eq_mon e_less_eq_refl) done lemma e_less_eq_lem1: "\\ e_less_eq e a;e_less_eq e (a + b)\ \ e_less_eq e b" apply (auto simp add: plus_def) apply (cases a) apply auto apply (cases b) apply auto apply (metis le_max_iff_disj) done lemma p_unwrap_less_sum: "snd (p_unwrap ((LP e aa) + b)) \ aa" apply (cases b) apply (auto simp add: plus_def) done lemma sum_list_less_elems: "\x\set xs. snd x \ Infty \ \y\set (map snd (map p_unwrap (map snd xs))). snd (p_unwrap (sum_list (map snd xs))) \ y" proof (induct xs) case Nil thus ?case by simp next case (Cons a as) thus ?case apply auto apply (cases "(snd a)" rule: p_unwrap.cases) apply auto apply (cases "sum_list (map snd as)") apply auto apply (metis linorder_linear p_min_re_neut p_unwrap.simps plus_def[abs_def] snd_eqD) apply (auto simp add: p_unwrap_less_sum) apply (unfold plus_def) apply (cases "(snd a, sum_list (map snd as))" rule: p_min.cases) apply auto apply (cases "map snd as") apply (auto simp add: infadd) apply (metis min.coboundedI2 snd_conv) done qed lemma distinct_sortet_list_app: "\sorted xs; distinct xs; xs = as @ b # cs\ \ \ x\ set cs. b < x" by (metis distinct.simps(2) distinct_append - linorder_antisym_conv2 sorted.simps(2) sorted_append) + antisym_conv2 sorted.simps(2) sorted_append) lemma distinct_sorted_list_lem1: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" " \ x \ set xs. x < e" " \ y \ set ys. e < y" shows "sorted (xs @ e # ys)" "distinct (xs @ e # ys)" proof - from assms (5,6) have "\x\set xs. \y\set ys. x \ y" by force thus "sorted (xs @ e # ys)" using assms by (auto simp add: sorted_append) have "set xs \ set ys = {}" using assms (5,6) by force thus "distinct (xs @ e # ys)" using assms by (auto) qed lemma distinct_sorted_list_lem2: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" "e < e'" " \ x \ set xs. x < e" " \ y \ set ys. e' < y" shows "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" proof - have "sorted (e' # ys)" "distinct (e' # ys)" "\ y \ set (e' # ys). e < y" using assms(2,4,5,7) by (auto) thus "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" using assms(1,3,6) distinct_sorted_list_lem1[of xs "e' # ys" e] by auto qed lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" apply(insert assms) apply(induct xs) apply clarsimp apply (metis dom_map_of_conv_image_fst fun_upd_None_restrict restrict_complement_singleton_eq restrict_map_self) apply (auto simp add: map_of_eq_None_iff) [] done lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed subsubsection "Finite" lemma aluprio_finite_correct: "uprio_finite (aluprio_\ \) (aluprio_invar \ invar)" by(unfold_locales) (simp add: aluprio_defs finite_dom_map_of) subsubsection "Empty" lemma aluprio_empty_correct: assumes "al_empty \ invar empt" shows "uprio_empty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_empty empt)" proof - interpret al_empty \ invar empt by fact show ?thesis apply (unfold_locales) apply (auto simp add: empty_correct aluprio_defs) done qed subsubsection "Is Empty" lemma aluprio_isEmpty_correct: assumes "al_isEmpty \ invar isEmpty" shows "uprio_isEmpty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_isEmpty isEmpty)" proof - interpret al_isEmpty \ invar isEmpty by fact show ?thesis apply (unfold_locales) apply (auto simp add: aluprio_defs isEmpty_correct) done qed subsubsection "Insert" lemma annot_inf: assumes A: "invar s" "\x\set (\ s). snd x \ Infty" "al_annot \ invar annot" shows "annot s = Infty \ \ s = [] " proof - from A have invs: "invar s" by (simp add: aluprio_defs) interpret al_annot \ invar annot by fact show "annot s = Infty \ \ s = []" proof (cases "\ s = []") case True hence "map snd (\ s) = []" by simp hence "sum_list (map snd (\ s)) = Infty" by (auto simp add: zero_def) with invs have "annot s = Infty" by (auto simp add: annot_correct) with True show ?thesis by simp next case False hence " \x xs. (\ s) = x # xs" by (cases "\ s") auto from this obtain x xs where [simp]: "(\ s) = x # xs" by blast from this assms(2) have "snd x \ Infty" by (auto simp add: aluprio_defs) hence "sum_list (map snd (\ s)) \ Infty" by (auto simp add: infadd) thus ?thesis using annot_correct invs False by simp qed qed lemma e_less_eq_annot: assumes "al_annot \ invar annot" "invar s" "\x\set (\ s). snd x \ Infty" "\ e_less_eq e (annot s)" shows "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" proof - interpret al_annot \ invar annot by fact from assms(2) have "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct) with assms(4) have "\x \ set (map snd (\ s)). \ e_less_eq e x" by (metis e_less_eq_sum_list) with assms(3) show ?thesis by (auto simp add: e_less_eq_p_unwrap) qed lemma aluprio_insert_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" "al_app \ invar app" "al_consr \ invar consr" shows "uprio_insert (aluprio_\ \) (aluprio_invar \ invar) (aluprio_insert splits annot isEmpty app consr)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact interpret al_app \ invar app by fact interpret al_consr \ invar consr by fact show ?thesis proof (unfold_locales, unfold aluprio_defs, goal_cases) case g1asms: (1 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct ) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) with prems(3) show ?case by(auto simp add: sorted_append) next case prems: 2 hence "annot s = sum_list (map snd (\ s))" by (simp add: annot_correct) with prems show ?case by (auto simp add: e_less_eq_sum_list2) next case prems: 3 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp next case prems: 4 hence "\ s = []" by (auto simp add: isEmpty_correct) with prems show ?case by simp qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto have v10: "sorted (map fst (map p_unwrap (map snd (\ l))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" "sorted (map fst (map p_unwrap (map snd (\ r))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" using g1asms v8 by (auto simp add: sorted_append) from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr (consr l () (LP e a)) () lp) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr (consr l () (LP e a)) () lp) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp hence v16: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" using v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem2 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp next case False hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr l () (LP e a)) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr l () (LP e a)) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp have v16: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto hence v17: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" using v16 v15 v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem1 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp qed qed next case g1asms: (2 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) hence "e \ set (map fst ((map (p_unwrap \ snd)) (\ s)))" by auto thus ?case by (auto simp add: map_of_distinct_upd) next case prems: 2 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp qed next case True note T1 = this obtain l lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 g1asms by auto from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap (LP e a))" by simp have v15: "e \ set (map fst (map p_unwrap (map snd(((),lp)#\ r))))" using v11(2) True by auto note map_of_distinct_upd2[OF v11(1) v15] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) = map_of (map p_unwrap (map snd (\ s)))(e \ a)" using v12 v13 v4(1) by simp next case False hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto note v15 = map_of_distinct_upd3[OF v11(1) v11(2)] have v16:"(map p_unwrap (map snd (\ s))) = (map p_unwrap (map snd (\ l))) @ (e,snd(p_unwrap lp)) # (map p_unwrap (map snd (\ r)))" using v4(1) v14 by simp note v15[of a "snd(p_unwrap lp)"] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) = map_of (map p_unwrap (map snd (\ s)))(e \ a)" using v12 v13 v16 by simp qed qed qed qed subsubsection "Prio" lemma aluprio_prio_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" shows "uprio_prio (aluprio_\ \) (aluprio_invar \ invar) (aluprio_prio splits annot isEmpty)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact show ?thesis proof (unfold_locales) fix s e assume inv1: "aluprio_invar \ invar s" hence sinv: "invar s" "(\ x\set (\ s). snd x\Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" by (auto simp add: aluprio_defs) show "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" proof(cases "e_less_eq e (annot s) \ \ isEmpty s") case False note F1 = this thus ?thesis proof(cases "isEmpty s") case True hence "\ s = []" using sinv isEmpty_correct by simp hence "aluprio_\ \ s = Map.empty" by (simp add:aluprio_defs) hence "aluprio_\ \ s e = None" by simp thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) next case False hence v3:"\ e_less_eq e (annot s)" using F1 by simp note v4=e_less_eq_annot[OF assms(2)] note v4[OF sinv(1) sinv(2) v3] hence v5:"e\set (map (fst \ (p_unwrap \ snd)) (\ s))" by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v5) thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 sinv annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using sinv v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) sinv v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 sinv by auto from l_lp_r T1 sinv show ?thesis proof (cases "e = fst (p_unwrap lp)") case False have v12: "e \ set (map fst (map p_unwrap (map snd(\ s))))" using v11 False v4(1) by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v12) thus ?thesis using T1 False l_lp_r by (auto simp add: aluprio_defs) next case True have v12: "map (p_unwrap \ snd) (\ s) = map p_unwrap (map snd (\ l)) @ (e,snd (p_unwrap lp)) # map p_unwrap (map snd (\ r))" using v4(1) True by simp note map_of_distinct_lookup[OF v11] hence "map_of (map (p_unwrap \ snd) (\ s)) e = Some (snd (p_unwrap lp))" using v12 by simp thus ?thesis using T1 True l_lp_r by (auto simp add: aluprio_defs) qed qed qed qed subsubsection "Pop" lemma aluprio_pop_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_app \ invar app" shows "uprio_pop (aluprio_\ \) (aluprio_invar \ invar) (aluprio_pop splits annot app)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_app \ invar app by fact show ?thesis proof (unfold_locales) fix s e a s' assume A: "aluprio_invar \ invar s" "aluprio_\ \ s \ Map.empty" "aluprio_pop splits annot app s = (e, a, s')" hence v1: "\ s \ []" by (auto simp add: aluprio_defs) obtain l lp r where l_lp_r: "splits (\ x. x\annot s) Infty s = (l,((),lp),r)" by (cases "splits (\ x. x\annot s) Infty s", auto) have invs: "invar s" "(\x\set (\ s). snd x \ Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" using A by (auto simp add:aluprio_defs) note a1 = annot_inf[of invar s \ annot] note a1[OF invs(1) invs(2) assms(2)] hence v2: "annot s \ Infty" using v1 by simp hence v3: "\ Infty \ annot s" by(cases "annot s") (auto simp add: plesseq_def) have v4: "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct invs(1)) hence v5: "(Infty + sum_list (map snd (\ s))) \ annot s" by (auto simp add: plus_def) note p_mon = p_less_eq_mon[of _ "annot s"] note v6 = splits_correct[OF invs(1)] note v7 = v6[of "\ x. x \ annot s"] note v7[OF _ v3 v5 l_lp_r] p_mon hence v8: " \ s = \ l @ ((), lp) # \ r" "\ Infty + sum_list (map snd (\ l)) \ annot s" "Infty + sum_list (map snd (\ l)) + lp \ annot s" "invar l" "invar r" by auto hence v9: "lp \ Infty" using invs(2) by auto hence v10: "s' = app l r" "(e,a) = p_unwrap lp" using l_lp_r A(3) apply (auto simp add: aluprio_defs) apply (cases lp) apply auto apply (cases lp) apply auto done have "lp \ annot s" using v8(2,3) p_less_eq_lem1 by auto hence v11: "a \ snd (p_unwrap (annot s))" using v10(2) v2 v9 apply (cases "annot s") apply auto apply (cases lp) apply (auto simp add: plesseq_def) done note sum_list_less_elems[OF invs(2)] hence v12: "\y\set (map snd (map p_unwrap (map snd (\ s)))). a \ y" using v4 v11 by auto have "ran (aluprio_\ \ s) = set (map snd (map p_unwrap (map snd (\ s))))" using ran_distinct[OF invs(4)] apply (unfold aluprio_defs) apply (simp only: set_map) done hence ziel1: "\y\ran (aluprio_\ \ s). a \ y" using v12 by simp have v13: "map p_unwrap (map snd (\ s)) = map p_unwrap (map snd (\ l)) @ (e,a) # map p_unwrap (map snd (\ r))" using v8(1) v10 by auto hence v14: "map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ e # map fst (map p_unwrap (map snd (\ r)))" by auto hence v15: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using invs(4) by auto note map_of_distinct_lookup[OF v15] note this[of a] hence ziel2: "aluprio_\ \ s e = Some a" using v13 by (unfold aluprio_defs, auto) have v16: "\ s' = \ l @ \ r" "invar s'" using v8(4,5) app_correct v10 by auto note map_of_distinct_upd4[OF v15] note this[of a] hence ziel3: "aluprio_\ \ s' = (aluprio_\ \ s)(e := None)" unfolding aluprio_defs using v16(1) v13 by auto have ziel4: "aluprio_invar \ invar s'" using v16 v8(1) invs(2,3,4) unfolding aluprio_defs by (auto simp add: sorted_append) show "aluprio_invar \ invar s' \ aluprio_\ \ s' = (aluprio_\ \ s)(e := None) \ aluprio_\ \ s e = Some a \ (\y\ran (aluprio_\ \ s). a \ y)" using ziel1 ziel2 ziel3 ziel4 by simp qed qed lemmas aluprio_correct = aluprio_finite_correct aluprio_empty_correct aluprio_isEmpty_correct aluprio_insert_correct aluprio_pop_correct aluprio_prio_correct locale aluprio_defs = StdALDefs ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin definition [icf_rec_def]: "aluprio_ops \ \ upr_\ = aluprio_\ \, upr_invar = aluprio_invar \ invar, upr_empty = aluprio_empty empty, upr_isEmpty = aluprio_isEmpty isEmpty, upr_insert = aluprio_insert splits annot isEmpty app consr, upr_pop = aluprio_pop splits annot app, upr_prio = aluprio_prio splits annot isEmpty \" end locale aluprio = aluprio_defs ops + StdAL ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin lemma aluprio_ops_impl: "StdUprio aluprio_ops" apply (rule StdUprio.intro) apply (simp_all add: icf_rec_unf) apply (rule aluprio_correct) apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] done end end diff --git a/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy b/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy --- a/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy +++ b/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy @@ -1,1162 +1,1162 @@ (* Author: Florian Messner Author: Julian Parsert Author: Jonas Schöpf Author: Christian Sternagel License: LGPL *) section \Homogeneous Linear Diophantine Equations\ theory Linear_Diophantine_Equations imports List_Vector begin (*TODO: move*) lemma lcm_div_le: fixes a :: nat shows "lcm a b div b \ a" by (metis div_by_0 div_le_dividend div_le_mono div_mult_self_is_m lcm_nat_def neq0_conv) (*TODO: move*) lemma lcm_div_le': fixes a :: nat shows "lcm a b div a \ b" by (metis lcm.commute lcm_div_le) (*TODO: move*) lemma lcm_div_gt_0: fixes a :: nat assumes "a > 0" and "b > 0" shows "lcm a b div a > 0" proof - have "lcm a b = (a * b) div (gcd a b)" using lcm_nat_def by blast moreover have "\ > 0" using assms by (metis assms calculation lcm_pos_nat) ultimately show ?thesis using assms by (metis div_by_0 div_mult2_eq div_positive gcd_le2_nat nat_mult_div_cancel_disj neq0_conv semiring_normalization_rules(7)) qed (*TODO: move*) lemma sum_list_list_update_Suc: assumes "i < length u" shows "sum_list (u[i := Suc (u ! i)]) = Suc (sum_list u)" using assms proof (induct u arbitrary: i) case (Cons x xs) then show ?case by (simp_all split: nat.splits) qed (simp) (*TODO: move*) lemma lessThan_conv: assumes "card A = n" and "\x\A. x < n" shows "A = {.. Given a non-empty list \xs\ of \n\ natural numbers, either there is a value in \xs\ that is \0\ modulo \n\, or there are two values whose moduli coincide. \ lemma list_mod_cases: assumes "length xs = n" and "n > 0" shows "(\x\set xs. x mod n = 0) \ (\ij j \ (xs ! i) mod n = (xs ! j) mod n)" proof - let ?f = "\x. x mod n" and ?X = "set xs" have *: "\x \ ?f ` ?X. x < n" using \n > 0\ by auto consider (eq) "card (?f ` ?X) = card ?X" | (less) "card (?f ` ?X) < card ?X" using antisym_conv2 and card_image_le by blast then show ?thesis proof (cases) case eq show ?thesis proof (cases "distinct xs") assume "distinct xs" with eq have "card (?f ` ?X) = n" using \distinct xs\ by (simp add: assms card_distinct distinct_card) from lessThan_conv [OF this *] and \n > 0\ have "\x\set xs. x mod n = 0" by (metis imageE lessThan_iff) then show ?thesis .. next assume "\ distinct xs" then show ?thesis by (auto) (metis distinct_conv_nth) qed next case less from pigeonhole [OF this] show ?thesis by (auto simp: inj_on_def iff: in_set_conv_nth) qed qed text \ Homogeneous linear Diophantine equations: \a\<^sub>1x\<^sub>1 + \ + a\<^sub>mx\<^sub>m = b\<^sub>1y\<^sub>1 + \ + b\<^sub>ny\<^sub>n\ \ locale hlde_ops = fixes a b :: "nat list" begin abbreviation "m \ length a" abbreviation "n \ length b" \ \The set of all solutions.\ definition Solutions :: "(nat list \ nat list) set" where "Solutions = {(x, y). a \ x = b \ y \ length x = m \ length y = n}" lemma in_Solutions_iff: "(x, y) \ Solutions \ length x = m \ length y = n \ a \ x = b \ y" by (auto simp: Solutions_def) \ \The set of pointwise minimal solutions.\ definition Minimal_Solutions :: "(nat list \ nat list) set" where "Minimal_Solutions = {(x, y) \ Solutions. nonzero x \ \ (\(u, v) \ Solutions. nonzero u \ u @ v <\<^sub>v x @ y)}" definition dij :: "nat \ nat \ nat" where "dij i j = lcm (a ! i) (b ! j) div (a ! i)" definition eij :: "nat \ nat \ nat" where "eij i j = lcm (a ! i) (b ! j) div (b ! j)" definition sij :: "nat \ nat \ (nat list \ nat list)" where "sij i j = ((zeroes m)[i := dij i j], (zeroes n)[j := eij i j])" subsection \Further Constraints on Minimal Solutions\ definition Ej :: "nat \ nat list \ nat set" where "Ej j x = { eij i j - 1 | i. i < length x \ x ! i \ dij i j }" definition Di :: "nat \ nat list \ nat set" where "Di i y = { dij i j - 1 | j. j < length y \ y ! j \ eij i j }" definition Di' :: "nat \ nat list \ nat set" where "Di' i y = { dij i (j + length b - length y) - 1 | j. j < length y \ y ! j \ eij i (j + length b - length y) }" lemma Ej_take_subset: "Ej j (take k x) \ Ej j x" by (auto simp: Ej_def) lemma Di_take_subset: "Di i (take l y) \ Di i y" by (auto simp: Di_def) lemma Di'_drop_subset: "Di' i (drop l y) \ Di' i y" by (auto simp: Di'_def) (metis add.assoc add.commute less_diff_conv) lemma finite_Ej: "finite (Ej j x)" by (rule finite_subset [of _ "(\i. eij i j - 1) ` {0 ..< length x}"]) (auto simp: Ej_def) lemma finite_Di: "finite (Di i y)" by (rule finite_subset [of _ "(\j. dij i j - 1) ` {0 ..< length y}"]) (auto simp: Di_def) lemma finite_Di': "finite (Di' i y)" by (rule finite_subset [of _ "(\j. dij i (j + length b - length y) - 1) ` {0 ..< length y}"]) (auto simp: Di'_def) definition max_y :: "nat list \ nat \ nat" where "max_y x j = (if j < n \ Ej j x \ {} then Min (Ej j x) else Max (set a))" definition max_x :: "nat list \ nat \ nat" where "max_x y i = (if i < m \ Di i y \ {} then Min (Di i y) else Max (set b))" definition max_x' :: "nat list \ nat \ nat" where "max_x' y i = (if i < m \ Di' i y \ {} then Min (Di' i y) else Max (set b))" lemma Min_Ej_le: assumes "j < n" and "e \ Ej j x" and "length x \ m" shows "Min (Ej j x) \ Max (set a)" (is "?m \ _") proof - have "?m \ Ej j x" using assms and finite_Ej and Min_in by blast then obtain i where i: "?m = eij i j - 1" "i < length x" "x ! i \ dij i j" by (auto simp: Ej_def) have "lcm (a ! i) (b ! j) div b ! j \ a ! i" by (rule lcm_div_le) then show ?thesis using i and assms by (auto simp: eij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma Min_Di_le: assumes "i < m" and "e \ Di i y" and "length y \ n" shows "Min (Di i y) \ Max (set b)" (is "?m \ _") proof - have "?m \ Di i y" using assms and finite_Di and Min_in by blast then obtain j where j: "?m = dij i j - 1" "j < length y" "y ! j \ eij i j" by (auto simp: Di_def) have "lcm (a ! i) (b ! j) div a ! i \ b ! j" by (rule lcm_div_le') then show ?thesis using j and assms by (auto simp: dij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma Min_Di'_le: assumes "i < m" and "e \ Di' i y" and "length y \ n" shows "Min (Di' i y) \ Max (set b)" (is "?m \ _") proof - have "?m \ Di' i y" using assms and finite_Di' and Min_in by blast then obtain j where j: "?m = dij i (j + length b - length y) - 1" "j < length y" "y ! j \ eij i (j + length b - length y)" by (auto simp: Di'_def) then have "j + length b - length y < length b" using assms by auto moreover have "lcm (a ! i) (b ! (j + length b - length y)) div a ! i \ b ! (j + length b - length y)" by (rule lcm_div_le') ultimately show ?thesis using j and assms by (auto simp: dij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma max_y_le_take: assumes "length x \ m" shows "max_y x j \ max_y (take k x) j" using assms and Min_Ej_le and Ej_take_subset and Min.subset_imp [OF _ _ finite_Ej] by (auto simp: max_y_def) blast lemma max_x_le_take: assumes "length y \ n" shows "max_x y i \ max_x (take l y) i" using assms and Min_Di_le and Di_take_subset and Min.subset_imp [OF _ _ finite_Di] by (auto simp: max_x_def) blast lemma max_x'_le_drop: assumes "length y \ n" shows "max_x' y i \ max_x' (drop l y) i" using assms and Min_Di'_le and Di'_drop_subset and Min.subset_imp [OF _ _ finite_Di'] by (auto simp: max_x'_def) blast end abbreviation "Solutions \ hlde_ops.Solutions" abbreviation "Minimal_Solutions \ hlde_ops.Minimal_Solutions" abbreviation "dij \ hlde_ops.dij" abbreviation "eij \ hlde_ops.eij" abbreviation "sij \ hlde_ops.sij" declare hlde_ops.dij_def [code] declare hlde_ops.eij_def [code] declare hlde_ops.sij_def [code] lemma Solutions_sym: "(x, y) \ Solutions a b \ (y, x) \ Solutions b a" by (auto simp: hlde_ops.in_Solutions_iff) lemma Minimal_Solutions_imp_Solutions: "(x, y) \ Minimal_Solutions a b \ (x, y) \ Solutions a b" by (auto simp: hlde_ops.Minimal_Solutions_def) lemma Minimal_SolutionsI: assumes "(x, y) \ Solutions a b" and "nonzero x" and "\ (\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y)" shows "(x, y) \ Minimal_Solutions a b" using assms by (auto simp: hlde_ops.Minimal_Solutions_def) lemma minimize_nonzero_solution: assumes "(x, y) \ Solutions a b" and "nonzero x" obtains u and v where "u @ v \\<^sub>v x @ y" and "(u, v) \ Minimal_Solutions a b" using assms proof (induct "x @ y" arbitrary: x y thesis rule: wf_induct [OF wf_less]) case 1 then show ?case proof (cases "(x, y) \ Minimal_Solutions a b") case False then obtain u and v where "nonzero u" and "(u, v) \ Solutions a b" and uv: "u @ v <\<^sub>v x @ y" using 1(3,4) by (auto simp: hlde_ops.Minimal_Solutions_def) with 1(1) [rule_format, of "u @ v" u v] obtain u' and v' where uv': "u' @ v' \\<^sub>v u @ v" and "(u', v') \ Minimal_Solutions a b" by blast moreover have "u' @ v' \\<^sub>v x @ y" using uv and uv' by auto ultimately show ?thesis by (intro 1(2)) qed blast qed lemma Minimal_SolutionsI': assumes "(x, y) \ Solutions a b" and "nonzero x" and "\ (\(u, v) \ Minimal_Solutions a b. u @ v <\<^sub>v x @ y)" shows "(x, y) \ Minimal_Solutions a b" proof (rule Minimal_SolutionsI [OF assms(1,2)]) show "\ (\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y)" proof assume "\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y" then obtain u and v where "(u, v) \ Solutions a b" and "nonzero u" and uv: "u @ v <\<^sub>v x @ y" by blast then obtain u' and v' where "(u', v') \ Minimal_Solutions a b" and uv': "u' @ v' \\<^sub>v u @ v" by (blast elim: minimize_nonzero_solution) moreover have "u' @ v' <\<^sub>v x @ y" using uv and uv' by auto ultimately show False using assms by blast qed qed lemma Minimal_Solutions_length: "(x, y) \ Minimal_Solutions a b \ length x = length a \ length y = length b" by (auto simp: hlde_ops.Minimal_Solutions_def hlde_ops.in_Solutions_iff) lemma Minimal_Solutions_gt0: "(x, y) \ Minimal_Solutions a b \ zeroes (length x) <\<^sub>v x" using zero_less by (auto simp: hlde_ops.Minimal_Solutions_def) lemma Minimal_Solutions_sym: assumes "0 \ set a" and "0 \ set b" shows "(xs, ys) \ Minimal_Solutions a b \ (ys, xs) \ Minimal_Solutions b a" using assms by (auto simp: hlde_ops.Minimal_Solutions_def hlde_ops.Solutions_def dest: dotprod_eq_nonzero_iff dest!: less_append_swap [of _ _ ys xs]) locale hlde = hlde_ops + assumes no0: "0 \ set a" "0 \ set b" begin lemma nonzero_Solutions_iff: assumes "(x, y) \ Solutions" shows "nonzero x \ nonzero y" using assms and no0 by (auto simp: in_Solutions_iff dest: dotprod_eq_nonzero_iff) lemma Minimal_Solutions_min: assumes "(x, y) \ Minimal_Solutions" and "u @ v <\<^sub>v x @ y" and "a \ u = b \ v" and [simp]: "length u = m" and non0: "nonzero (u @ v)" shows False proof - have [simp]: "length v = n" using assms by (force dest: less_appendD Minimal_Solutions_length) have "(u, v) \ Solutions" using \a \ u = b \ v\ by (simp add: in_Solutions_iff) moreover from nonzero_Solutions_iff [OF this] have "nonzero u" using non0 by auto ultimately show False using assms by (auto simp: hlde_ops.Minimal_Solutions_def) qed lemma Solutions_snd_not_0: assumes "(x, y) \ Solutions" and "nonzero x" shows "nonzero y" using assms by (metis nonzero_Solutions_iff) end subsection \Pointwise Restricting Solutions\ text \ Constructing the list of \u\ vectors from Huet's proof @{cite "Huet1978"}, satisfying \<^item> \\i y ! i\ and \<^item> \0 < sum_list u \ a\<^sub>k\. \ text \ Given \y\, increment a "previous" \u\ vector at first position starting from \i\ where \u\ is strictly smaller than \y\. If this is not possible, return \u\ unchanged. \ function inc :: "nat list \ nat \ nat list \ nat list" where "inc y i u = (if i < length y then if u ! i < y ! i then u[i := u ! i + 1] else inc y (Suc i) u else u)" by (pat_completeness) auto termination inc by (relation "measure (\(y, i, u). max (length y) (length u) - i)") auto (*inc.simps may cause simplification to loop*) declare inc.simps [simp del] text \ Starting from the 0-vector produce \u\s by iteratively incrementing with respect to \y\. \ definition huets_us :: "nat list \ nat \ nat list" ("\<^bold>u" 1000) where "\<^bold>u y i = ((inc y 0) ^^ Suc i) (zeroes (length y))" lemma huets_us_simps [simp]: "\<^bold>u y 0 = inc y 0 (zeroes (length y))" "\<^bold>u y (Suc i) = inc y 0 (\<^bold>u y i)" by (auto simp: huets_us_def) lemma length_inc [simp]: "length (inc y i u) = length u" by (induct y i u rule: inc.induct) (simp add: inc.simps) lemma length_us [simp]: "length (\<^bold>u y i) = length y" by (induct i) (simp_all) text \ \inc\ produces vectors that are pointwise smaller than \y\ \ lemma inc_le: assumes "length u = length y" and "i < length y" and "u \\<^sub>v y" shows "inc y i u \\<^sub>v y" using assms by (induct y i u rule: inc.induct) (auto simp: inc.simps nth_list_update less_eq_def) lemma us_le: assumes "length y > 0" shows "\<^bold>u y i \\<^sub>v y" using assms by (induct i) (auto simp: inc_le le_length) lemma sum_list_inc_le: "u \\<^sub>v y \ sum_list (inc y i u) \ sum_list y" by (induct y i u rule: inc.induct) (auto simp: inc.simps intro: le_sum_list_mono) lemma sum_list_inc_gt0: assumes "sum_list u > 0" and "length y = length u" shows "sum_list (inc y i u) > 0" using assms proof (induct y i u rule: inc.induct) case (1 y i u) then show ?case by (auto simp add: inc.simps) (meson Suc_neq_Zero gr_zeroI set_update_memI sum_list_eq_0_iff) qed lemma sum_list_inc_gt0': assumes "length u = length y" and "i < length y" and "y ! i > 0" and "j \ i" shows "sum_list (inc y j u) > 0" using assms proof (induct y j u rule: inc.induct) case (1 y i u) then show ?case by (auto simp: inc.simps [of y i] sum_list_update) (metis elem_le_sum_list le_antisym le_zero_eq neq0_conv not_less_eq_eq sum_list_inc_gt0) qed lemma sum_list_us_gt0: assumes "sum_list y \ 0" shows "0 < sum_list (\<^bold>u y i)" using assms by (induct i) (auto simp: in_set_conv_nth sum_list_inc_gt0' sum_list_inc_gt0) lemma sum_list_inc_le': assumes "length u = length y" shows "sum_list (inc y i u) \ sum_list u + 1" using assms by (induct y i u rule: inc.induct) (auto simp: inc.simps sum_list_update) lemma sum_list_us_le: "sum_list (\<^bold>u y i) \ i + 1" proof (induct i) case 0 then show ?case by (auto simp: sum_list_update) (metis Suc_eq_plus1 in_set_replicate length_replicate sum_list_eq_0_iff sum_list_inc_le') next case (Suc i) then show ?case by auto (metis Suc_le_mono add.commute le_trans length_us plus_1_eq_Suc sum_list_inc_le') qed lemma sum_list_us_bounded: assumes "i < k" shows "sum_list (\<^bold>u y i) \ k" using assms and sum_list_us_le [of y i] by force lemma sum_list_inc_eq_sum_list_Suc: assumes "length u = length y" and "i < length y" and "\j\i. j < length y \ u ! j < y ! j" shows "sum_list (inc y i u) = Suc (sum_list u)" using assms by (induct y i u rule: inc.induct) (metis inc.simps Suc_eq_plus1 Suc_leI antisym_conv2 leD sum_list_list_update_Suc) lemma sum_list_us_eq: assumes "i < sum_list y" shows "sum_list (\<^bold>u y i) = i + 1" using assms proof (induct i) case (Suc i) then show ?case by (auto) (metis (no_types, lifting) Suc_eq_plus1 gr_implies_not0 length_pos_if_in_set - length_us less_Suc_eq_le less_imp_le_nat linorder_antisym_conv2 not_less_eq_eq + length_us less_Suc_eq_le less_imp_le_nat antisym_conv2 not_less_eq_eq sum_list_eq_0_iff sum_list_inc_eq_sum_list_Suc sum_list_less_diff_Ex us_le) qed (metis Suc_eq_plus1 Suc_leI antisym_conv gr_implies_not0 sum_list_us_gt0 sum_list_us_le) lemma inc_ge: "length u = length y \ u \\<^sub>v inc y i u" by (induct y i u rule: inc.induct) (auto simp: inc.simps nth_list_update less_eq_def) lemma us_le_mono: assumes "i < j" shows "\<^bold>u y i \\<^sub>v \<^bold>u y j" using assms proof (induct "j - i" arbitrary: j i) case (Suc n) then show ?case by (simp add: Suc.prems inc_ge order.strict_implies_order order_vec.lift_Suc_mono_le) qed simp lemma us_mono: assumes "i < j" and "j < sum_list y" shows "\<^bold>u y i <\<^sub>v \<^bold>u y j" proof - let ?u = "\<^bold>u y i" and ?v = "\<^bold>u y j" have "?u \\<^sub>v ?v" using us_le_mono [OF \i < j\] by simp moreover have "sum_list ?u < sum_list ?v" using assms by (auto simp: sum_list_us_eq) ultimately show ?thesis by (intro le_sum_list_less) (auto simp: less_eq_def) qed context hlde begin lemma max_coeff_bound_right: assumes "(xs, ys) \ Minimal_Solutions" shows "\x \ set xs. x \ maxne0 ys b" (is "\x\set xs. x \ ?m") proof (rule ccontr) assume "\ ?thesis" then obtain k where k_def: "k < length xs \ \ (xs ! k \ ?m)" by (metis in_set_conv_nth) have sol: "(xs, ys) \ Solutions" using assms Minimal_Solutions_def by auto then have len: "m = length xs" by (simp add: in_Solutions_iff) have max_suml: "?m * sum_list ys \ b \ ys" using maxne0_times_sum_list_gt_dotprod sol by (auto simp: in_Solutions_iff) then have is_sol: "b \ ys = a \ xs" using sol by (auto simp: in_Solutions_iff) then have a_ge_ak: "a \ xs \ a ! k * xs ! k" using dotprod_pointwise_le k_def len by auto then have ak_gt_max: "a ! k * xs ! k > a ! k * ?m" using no0 in_set_conv_nth k_def len by fastforce then have sl_ys_g_ak: "sum_list ys > a ! k" by (metis a_ge_ak is_sol less_le_trans max_suml mult.commute mult_le_mono1 not_le) define Seq where Seq_def: "Seq = map (\<^bold>u ys) [0 ..< a ! k]" have ak_n0: "a ! k \ 0" using \a ! k * ?m < a ! k * xs ! k\ by auto have "zeroes (length ys) <\<^sub>v ys" by (intro zero_less) (metis gr_implies_not0 nonzero_iff sl_ys_g_ak sum_list_eq_0_iff) then have "length Seq > 0" using ak_n0 Seq_def by auto have u_in_nton: "\u \ set Seq. length u = length ys" by (simp add: Seq_def) have prop_3: "\u \ set Seq. u \\<^sub>v ys" proof - have "length ys > 0" using sl_ys_g_ak by auto then show ?thesis using us_le [of ys ] less_eq_def Seq_def by (simp) qed have prop_4_1: "\u \ set Seq. sum_list u > 0" by (metis Seq_def sl_ys_g_ak gr_implies_not_zero imageE set_map sum_list_us_gt0) have prop_4_2: "\u \ set Seq. sum_list u \ a ! k" by (simp add: Seq_def sum_list_us_bounded) have prop_5: "\u. length u = length ys \ u \\<^sub>v ys \ sum_list u > 0 \ sum_list u \ a ! k" using \0 < length Seq\ nth_mem prop_3 prop_4_1 prop_4_2 u_in_nton by blast define Us where "Us = {u. length u = length ys \ u \\<^sub>v ys \ sum_list u > 0 \ sum_list u \ a ! k}" have "\u \ Us. b \ u mod a ! k = 0" proof (rule ccontr) assume neg_th: "\ ?thesis" define Seq_p where "Seq_p = map (dotprod b) Seq" have "length Seq = a ! k" by (simp add: Seq_def) then consider (eq_0) "(\x\set Seq_p. x mod (a ! k) = 0)" | (not_0) "(\ij j \ (Seq_p ! i) mod (a!k) = (Seq_p ! j) mod (a!k))" using list_mod_cases[of Seq_p] Seq_p_def ak_n0 by auto then show False proof (cases) case eq_0 have "\u \ set Seq. b \ u mod a ! k = 0" using Seq_p_def eq_0 by auto then show False by (metis (mono_tags, lifting) Us_def mem_Collect_eq neg_th prop_3 prop_4_1 prop_4_2 u_in_nton) next case not_0 obtain i and j where i_j: "i j" " Seq_p ! i mod a ! k = Seq_p ! j mod a ! k" using not_0 by blast define v where v_def: "v = Seq!i" define w where w_def: "w = Seq!j" have mod_eq: "b \ v mod a!k = b \ w mod a!k" using Seq_p_def i_j w_def v_def i_j by auto have "v <\<^sub>v w \ w <\<^sub>v v" using \i \ j\ and i_j proof (cases "i < j") case True then show ?thesis using Seq_p_def sl_ys_g_ak i_j(2) local.Seq_def us_mono v_def w_def by auto next case False then show ?thesis using Seq_p_def sl_ys_g_ak \i \ j\ i_j(1) local.Seq_def us_mono v_def w_def by auto qed then show False proof assume ass: "v <\<^sub>v w" define u where u_def: "u = w -\<^sub>v v" have "w \\<^sub>v ys" using Seq_p_def w_def i_j(2) prop_3 by force then have prop_3: "less_eq u ys" using vdiff_le ass less_eq_def order_vec.less_imp_le u_def by auto have prop_4_1: "sum_list u > 0" using le_sum_list_mono [of v w] ass u_def sum_list_vdiff_distr [of v w] by (simp add: less_vec_sum_list_less) have prop_4_2: "sum_list u \ a ! k" proof - have "u \\<^sub>v w" using u_def using ass less_eq_def order_vec.less_imp_le vdiff_le by auto then show ?thesis by (metis Seq_p_def i_j(2) length_map le_sum_list_mono less_le_trans not_le nth_mem prop_4_2 w_def) qed have "b \ u mod a ! k = 0" by (metis (mono_tags, lifting) in_Solutions_iff \w \\<^sub>v ys\ u_def ass no0(2) less_eq_def mem_Collect_eq mod_eq mods_with_vec_2 prod.simps(2) sol) then show False using neg_th by (metis (mono_tags, lifting) Us_def less_eq_def mem_Collect_eq prop_3 prop_4_1 prop_4_2) next assume ass: "w <\<^sub>v v" define u where u_def: "u = v -\<^sub>v w" have "v \\<^sub>v ys" using Seq_p_def v_def i_j(1) prop_3 by force then have prop_3: "u \\<^sub>v ys" using vdiff_le ass less_eq_def order_vec.less_imp_le u_def by auto have prop_4_1: "sum_list u > 0" using le_sum_list_mono [of w v] sum_list_vdiff_distr [of w v] \u \ v -\<^sub>v w\ ass less_vec_sum_list_less by auto have prop_4_2: "sum_list u \ a!k" proof - have "u \\<^sub>v v" using u_def using ass less_eq_def order_vec.less_imp_le vdiff_le by auto then show ?thesis by (metis Seq_p_def i_j(1) le_neq_implies_less length_map less_imp_le_nat less_le_trans nth_mem prop_4_2 le_sum_list_mono v_def) qed have "b \ u mod a ! k = 0" by (metis (mono_tags, lifting) in_Solutions_iff \v \\<^sub>v ys\ u_def ass no0(2) less_eq_def mem_Collect_eq mod_eq mods_with_vec_2 prod.simps(2) sol) then show False by (metis (mono_tags, lifting) neg_th Us_def less_eq_def mem_Collect_eq prop_3 prop_4_1 prop_4_2) qed qed qed then obtain u where u3_4: "u \\<^sub>v ys" "sum_list u > 0" "sum_list u \ a ! k" " b \ u mod (a ! k) = 0" "length u = length ys" unfolding Us_def by auto have u_b_len: "length u = n" using less_eq_def u3_4 in_Solutions_iff sol by simp have "b \ u \ maxne0 u b * sum_list u" by (simp add: maxne0_times_sum_list_gt_dotprod u_b_len) also have "... \ ?m * a ! k" by (intro mult_le_mono) (simp_all add: u3_4 maxne0_mono) also have "... < a ! k * xs ! k" using ak_gt_max by auto then obtain zk where zk: "b \ u = zk * a ! k" using u3_4(4) by auto have "length xs > k" by (simp add: k_def) have "zk \ 0" proof - have "\e \ set u. e \ 0" using u3_4 by (metis neq0_conv sum_list_eq_0_iff) then have "b \ u > 0" using assms no0 u3_4 unfolding dotprod_gt0_iff[OF u_b_len [symmetric]] by (fastforce simp add: in_set_conv_nth u_b_len) then have "a ! k > 0" using \a ! k \ 0\ by blast then show ?thesis using \0 < b \ u\ zk by auto qed define z where z_def: "z = (zeroes (length xs))[k := zk]" then have zk_zk: "z ! k = zk" by (auto simp add: \k < length xs\) have "length z = length xs" using assms z_def \k < length xs\ by auto then have bu_eq_akzk: "b \ u = a ! k * z ! k" by (simp add: \b \ u = zk * a ! k\ zk_zk) then have "z!k < xs!k" using ak_gt_max calculation by auto then have z_less_xs: "z <\<^sub>v xs" by (auto simp add: z_def) (metis \k < length xs\ le0 le_list_update less_def less_imp_le order_vec.dual_order.antisym nat_neq_iff z_def zk_zk) then have "z @ u <\<^sub>v xs @ ys" by (intro less_append) (auto simp add: u3_4(1) z_less_xs) moreover have "(z, u) \ Solutions" by (auto simp add: bu_eq_akzk in_Solutions_iff z_def u_b_len \k < length xs\ len) moreover have "nonzero z" using \length z = length xs\ and \zk \ 0\ and k_def and zk_zk by (auto simp: nonzero_iff) ultimately show False using assms by (auto simp: Minimal_Solutions_def) qed text \Proof of Lemma 1 of Huet's paper.\ lemma max_coeff_bound: assumes "(xs, ys) \ Minimal_Solutions" shows "(\x \ set xs. x \ maxne0 ys b) \ (\y \ set ys. y \ maxne0 xs a)" proof - interpret ba: hlde b a by (standard) (auto simp: no0) show ?thesis using assms and Minimal_Solutions_sym [OF no0, of xs ys] by (auto simp: max_coeff_bound_right ba.max_coeff_bound_right) qed lemma max_coeff_bound': assumes "(x, y) \ Minimal_Solutions" shows "\i Max (set b)" and "\j Max (set a)" using max_coeff_bound [OF assms] and maxne0_le_Max by auto (metis le_eq_less_or_eq less_le_trans nth_mem)+ lemma Minimal_Solutions_alt_def: "Minimal_Solutions = {(x, y)\Solutions. (x, y) \ (zeroes m, zeroes n) \ x \\<^sub>v replicate m (Max (set b)) \ y \\<^sub>v replicate n (Max (set a)) \ \ (\(u, v)\Solutions. nonzero u \ u @ v <\<^sub>v x @ y)}" by (auto simp: not_nonzero_iff Minimal_Solutions_imp_Solutions less_eq_def Minimal_Solutions_length max_coeff_bound' intro!: Minimal_SolutionsI' dest: Minimal_Solutions_gt0) (auto simp: Minimal_Solutions_def nonzero_Solutions_iff not_nonzero_iff) subsection \Special Solutions\ definition Special_Solutions :: "(nat list \ nat list) set" where "Special_Solutions = {sij i j | i j. i < m \ j < n}" lemma dij_neq_0: assumes "i < m" and "j < n" shows "dij i j \ 0" proof - have "a ! i > 0" and "b ! j > 0" using assms and no0 by (simp_all add: in_set_conv_nth) then have "dij i j > 0" using lcm_div_gt_0 [of "a ! i" "b ! j"] by (simp add: dij_def) then show ?thesis by simp qed lemma eij_neq_0: assumes "i < m" and "j < n" shows "eij i j \ 0" proof - have "a ! i > 0" and "b ! j > 0" using assms and no0 by (simp_all add: in_set_conv_nth) then have "eij i j > 0" using lcm_div_gt_0[of "b ! j" "a ! i"] by (simp add: eij_def lcm.commute) then show ?thesis by simp qed lemma Special_Solutions_in_Solutions: "x \ Special_Solutions \ x \ Solutions" by (auto simp: in_Solutions_iff Special_Solutions_def sij_def dij_def eij_def) lemma Special_Solutions_in_Minimal_Solutions: assumes "(x, y) \ Special_Solutions" shows "(x, y) \ Minimal_Solutions" proof (intro Minimal_SolutionsI') show "(x, y) \ Solutions" by (fact Special_Solutions_in_Solutions [OF assms]) then have [simp]: "length x = m" "length y = n" by (auto simp: in_Solutions_iff) show "nonzero x" using assms and dij_neq_0 by (auto simp: Special_Solutions_def sij_def nonzero_iff) (metis length_replicate set_update_memI) show "\ (\(u, v)\Minimal_Solutions. u @ v <\<^sub>v x @ y)" proof assume "\(u, v)\Minimal_Solutions. u @ v <\<^sub>v x @ y" then obtain u and v where uv: "(u, v) \ Minimal_Solutions" and "u @ v <\<^sub>v x @ y" and [simp]: "length u = m" "length v = n" and "nonzero u" by (auto simp: Minimal_Solutions_def in_Solutions_iff) then consider "u <\<^sub>v x" and "v \\<^sub>v y" | "v <\<^sub>v y" and "u \\<^sub>v x" by (auto elim: less_append_cases) then show False proof (cases) case 1 then obtain i and j where ij: "i < m" "j < n" and less_dij: "u ! i < dij i j" and "u \\<^sub>v (zeroes m)[i := dij i j]" and "v \\<^sub>v (zeroes n)[j := eij i j]" using assms by (auto simp: Special_Solutions_def sij_def unit_less) then have u: "u = (zeroes m)[i := u ! i]" and v: "v = (zeroes n)[j := v ! j]" by (auto simp: less_eq_def list_eq_iff_nth_eq) (metis le_zero_eq length_list_update length_replicate rep_upd_unit)+ then have "u ! i > 0" using \nonzero u\ and ij by (metis gr_implies_not0 neq0_conv unit_less zero_less) define c where "c = a ! i * u ! i" then have ac: "a ! i dvd c" by simp have "a \ u = b \ v" using uv by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "c = b ! j * v ! j" using ij unfolding c_def by (subst (asm) u, subst (asm)v, subst u, subst v) auto then have bc: "b ! j dvd c" by simp have "a ! i * u ! i < a ! i * dij i j" using less_dij and no0 and ij by (auto simp: in_set_conv_nth) then have "c < lcm (a ! i) (b ! j)" by (auto simp: dij_def c_def) moreover have "lcm (a ! i) (b ! j) dvd c" by (simp add: ac bc) moreover have "c > 0" using \u ! i > 0\ and no0 and ij by (auto simp: c_def in_set_conv_nth) ultimately show False using ac and bc by (auto dest: nat_dvd_not_less) next case 2 then obtain i and j where ij: "i < m" "j < n" and less_dij: "v ! j < eij i j" and "u \\<^sub>v (zeroes m)[i := dij i j]" and "v \\<^sub>v (zeroes n)[j := eij i j]" using assms by (auto simp: Special_Solutions_def sij_def unit_less) then have u: "u = (zeroes m)[i := u ! i]" and v: "v = (zeroes n)[j := v ! j]" by (auto simp: less_eq_def list_eq_iff_nth_eq) (metis le_zero_eq length_list_update length_replicate rep_upd_unit)+ moreover have "nonzero v" using \nonzero u\ and \(u, v) \ Minimal_Solutions\ and Minimal_Solutions_imp_Solutions Solutions_snd_not_0 by blast ultimately have "v ! j > 0" using ij by (metis gr_implies_not0 neq0_conv unit_less zero_less) define c where "c = b ! j * v ! j" then have bc: "b ! j dvd c" by simp have "a \ u = b \ v" using uv by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "c = a ! i * u ! i" using ij unfolding c_def by (subst (asm) u, subst (asm)v, subst u, subst v) auto then have ac: "a ! i dvd c" by simp have "b ! j * v ! j < b ! j * eij i j" using less_dij and no0 and ij by (auto simp: in_set_conv_nth) then have "c < lcm (a ! i) (b ! j)" by (auto simp: eij_def c_def) moreover have "lcm (a ! i) (b ! j) dvd c" by (simp add: ac bc) moreover have "c > 0" using \v ! j > 0\ and no0 and ij by (auto simp: c_def in_set_conv_nth) ultimately show False using ac and bc by (auto dest: nat_dvd_not_less) qed qed qed (*Lemma 2 of Huet*) lemma non_special_solution_non_minimal: assumes "(x, y) \ Solutions - Special_Solutions" and ij: "i < m" "j < n" and "x ! i \ dij i j" and "y ! j \ eij i j" shows "(x, y) \ Minimal_Solutions" proof assume min: "(x, y) \ Minimal_Solutions" moreover have "sij i j \ Solutions" using ij by (intro Special_Solutions_in_Solutions) (auto simp: Special_Solutions_def) moreover have "(case sij i j of (u, v) \ u @ v) <\<^sub>v x @ y" using assms and min apply (cases "sij i j") apply (auto simp: sij_def Special_Solutions_def) by (metis List_Vector.le0 Minimal_Solutions_length le_append le_list_update less_append order_vec.dual_order.strict_iff_order same_append_eq) moreover have "(case sij i j of (u, v) \ nonzero u)" apply (auto simp: sij_def) by (metis dij_neq_0 ij length_replicate nonzero_iff set_update_memI) ultimately show False by (auto simp: Minimal_Solutions_def) qed subsection \Huet's conditions\ (*A*) definition "cond_A xs ys \ (\x\set xs. x \ maxne0 ys b)" (*B*) definition "cond_B x \ (\k\m. take k a \ take k x \ b \ map (max_y (take k x)) [0 ..< n])" (*C*) definition "boundr x y \ (\j max_y x j)" (*D*) definition "cond_D x y \ (\l\n. take l b \ take l y \ a \ x)" subsection \New conditions: facilitating generation of candidates from right to left\ (*condition on right sub-dotproduct*) definition "subdprodr y \ (\l\n. take l b \ take l y \ a \ map (max_x (take l y)) [0 ..< m])" (*condition on left sub-dotproduct*) definition "subdprodl x y \ (\k\m. take k a \ take k x \ b \ y)" (*bound on elements of left vector*) definition "boundl x y \ (\i max_x y i)" lemma boundr: assumes min: "(x, y) \ Minimal_Solutions" and "(x, y) \ Special_Solutions" shows "boundr x y" proof (unfold boundr_def, intro allI impI) fix j assume ass: "j < n" have ln: "m = length x \ n = length y" using assms Minimal_Solutions_def in_Solutions_iff min by auto have is_sol: "(x, y) \ Solutions" using assms Minimal_Solutions_def min by auto have j_less_l: "j < n" using assms ass le_less_trans by linarith consider (notemp) "Ej j x \ {}" | (empty) " Ej j x = {}" by blast then show "y ! j \ max_y x j" proof (cases) case notemp have max_y_def: "max_y x j = Min (Ej j x)" using j_less_l max_y_def notemp by auto have fin_e: "finite (Ej j x)" using finite_Ej [of j x] by auto have e_def': "\e \ Ej j x. (\i dij i j \ eij i j - 1 = e)" using Ej_def [of j x] by auto then have "\i dij i j \ eij i j - 1 = Min (Ej j x)" using notemp Min_in e_def' fin_e by blast then obtain i where i: "i < length x" "x ! i \ dij i j" "eij i j - 1 = Min (Ej j x)" by blast show ?thesis proof (rule ccontr) assume "\ ?thesis" with non_special_solution_non_minimal [of x y i j] and i and ln and assms and is_sol and j_less_l have "case sij i j of (u, v) \ u @ v \\<^sub>v x @ y" by (force simp: max_y_def) then have cs:"case sij i j of (u, v) \ u @ v <\<^sub>v x @ y" using assms by(auto simp: Special_Solutions_def) (metis append_eq_append_conv i(1) j_less_l length_list_update length_replicate sij_def order_vec.le_neq_trans ln prod.sel(1)) then obtain u v where u_v: "sij i j = (u, v)" "u @ v <\<^sub>v x @ y" by blast have dij_gt0: "dij i j > 0" using assms(1) assms(2) dij_neq_0 i(1) j_less_l ln by auto then have not_0_u: "nonzero u" proof (unfold nonzero_iff) have "i < length (zeroes m)" by (simp add: i(1) ln) then show "\i\set u. i \ 0" by (metis (no_types) Pair_inject dij_gt0 set_update_memI sij_def u_v(1) neq0_conv) qed then have "sij i j \ Solutions" by (metis (mono_tags, lifting) Special_Solutions_def i(1) Special_Solutions_in_Solutions j_less_l ln mem_Collect_eq u_v(1)) then show False using assms cs u_v not_0_u Minimal_Solutions_def min by auto qed next case empty have "\y\set y. y \ Max (set a)" using assms and max_coeff_bound and maxne0_le_Max using le_trans by blast then show ?thesis using empty j_less_l ln max_y_def by auto qed qed lemma boundl: assumes min: "(x, y) \ Minimal_Solutions" and "(x, y) \ Special_Solutions" shows "boundl x y" proof (unfold boundl_def, intro allI impI) fix i assume ass: "i < m" have ln: "n = length y \ m = length x" using assms Minimal_Solutions_def in_Solutions_iff min by auto have is_sol: "(x, y) \ Solutions" using assms Minimal_Solutions_def min by auto have i_less_l: "i < m" using assms ass le_less_trans by linarith consider (notemp) "Di i y \ {}" | (empty) " Di i y = {}" by blast then show "x ! i \ max_x y i" proof (cases) case notemp have max_x_def: "max_x y i = Min (Di i y)" using i_less_l max_x_def notemp by auto have fin_e: "finite (Di i y)" using finite_Di [of i y] by auto have e_def': "\e \ Di i y. (\j eij i j \ dij i j - 1 = e)" using Di_def [of i y] by auto then have "\j eij i j \ dij i j - 1 = Min (Di i y)" using notemp Min_in e_def' fin_e by blast then obtain j where j: "j < length y" "y ! j \ eij i j" "dij i j - 1 = Min (Di i y)" by blast show ?thesis proof (rule ccontr) assume "\ ?thesis" with non_special_solution_non_minimal [of x y i j] and j and ln and assms and is_sol and i_less_l have "case sij i j of (u, v) \ u @ v \\<^sub>v x @ y" by (force simp: max_x_def) then have cs: "case sij i j of (u, v) \ u @ v <\<^sub>v x @ y" using assms by(auto simp: Special_Solutions_def) (metis append_eq_append_conv j(1) i_less_l length_list_update length_replicate sij_def order_vec.le_neq_trans ln prod.sel(1)) then obtain u v where u_v: "sij i j = (u, v)" "u @ v <\<^sub>v x @ y" by blast have dij_gt0: "dij i j > 0" using assms(1) assms(2) dij_neq_0 j(1) i_less_l ln by auto then have not_0_u: "nonzero u" proof (unfold nonzero_iff) have "i < length (zeroes m)" using ass by simp then show "\i\set u. i \ 0" by (metis (no_types) Pair_inject dij_gt0 set_update_memI sij_def u_v(1) neq0_conv) qed then have "sij i j \ Solutions" by (metis (mono_tags, lifting) Special_Solutions_def j(1) Special_Solutions_in_Solutions i_less_l ln mem_Collect_eq u_v(1)) then show False using assms cs u_v not_0_u Minimal_Solutions_def min by auto qed next case empty have "\x\set x. x \ Max (set b)" using assms and max_coeff_bound and maxne0_le_Max using le_trans by blast then show ?thesis using empty i_less_l ln max_x_def by auto qed qed lemma Solution_imp_cond_D: assumes "(x, y) \ Solutions" shows "cond_D x y" using assms and dotprod_le_take by (auto simp: cond_D_def in_Solutions_iff) lemma Solution_imp_subdprodl: assumes "(x, y) \ Solutions" shows "subdprodl x y" using assms and dotprod_le_take by (auto simp: subdprodl_def in_Solutions_iff) metis theorem conds: assumes min: "(x, y) \ Minimal_Solutions" shows cond_A: "cond_A x y" and cond_B: "(x, y) \ Special_Solutions \ cond_B x" and "(x, y) \ Special_Solutions \ boundr x y" and cond_D: "cond_D x y" and subdprodr: "(x, y) \ Special_Solutions \ subdprodr y" and subdprodl: "subdprodl x y" proof - have sol: "a \ x = b \ y" and ln: "m = length x \ n = length y" using min by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "\i maxne0 y b" by (metis min max_coeff_bound_right nth_mem) then show "cond_A x y" using min and le_less_trans by (auto simp: cond_A_def max_coeff_bound) show "(x, y) \ Special_Solutions \ cond_B x" proof (unfold cond_B_def, intro allI impI) fix k assume non_spec: "(x, y) \ Special_Solutions" and k: "k \ m" from k have "take k a \ take k x \ a \ x" using dotprod_le_take ln by blast also have "... = b \ y" by fact also have map_b_dot_p: "... \ b \ map (max_y x) [0.. _ b \ ?nt") using non_spec and less_eq_def and ln and boundr and min by (fastforce intro!: dotprod_le_right simp: boundr_def) also have "... \ b \ map (max_y (take k x)) [0.. _ \ ?t") proof - have "\j ?t!j" using min and ln and max_y_le_take and k by auto then have "?nt \\<^sub>v ?t" using less_eq_def by auto then show ?thesis by (simp add: dotprod_le_right) qed finally show "take k a \ take k x \ b \ map (max_y (take k x)) [0.. Special_Solutions \ subdprodr y" proof (unfold subdprodr_def, intro allI impI) fix l assume non_spec: "(x, y) \ Special_Solutions" and l: "l \ n" from l have "take l b \ take l y \ b \ y" using dotprod_le_take ln by blast also have "... = a \ x" by (simp add: sol) also have map_b_dot_p: "... \ a \ map (max_x y) [0.. _ a \ ?nt") using non_spec and less_eq_def and ln and boundl and min by (fastforce intro!: dotprod_le_right simp: boundl_def) also have "... \ a \ map (max_x (take l y)) [0.. _ \ ?t") proof - have "\i ?t ! i" using min and ln and max_x_le_take and l by auto then have "?nt \\<^sub>v ?t" using less_eq_def by auto then show ?thesis by (simp add: dotprod_le_right) qed finally show "take l b \ take l y \ a \ map (max_x (take l y)) [0.. Special_Solutions \ boundr x y" using boundr [of x y] and min by blast show "cond_D x y" using ln and dotprod_le_take and sol by (auto simp: cond_D_def) show "subdprodl x y" using ln and dotprod_le_take and sol by (force simp: subdprodl_def) qed lemma le_imp_Ej_subset: assumes "u \\<^sub>v x" shows "Ej j u \ Ej j x" using assms and le_trans by (force simp: Ej_def less_eq_def dij_def eij_def) lemma le_imp_max_y_ge: assumes "u \\<^sub>v x" and "length x \ m" shows "max_y u j \ max_y x j" using assms and le_imp_Ej_subset and Min_Ej_le [of j, OF _ _ assms(2)] by (metis Min.subset_imp Min_in emptyE finite_Ej max_y_def order_refl subsetCE) lemma le_imp_Di_subset: assumes "v \\<^sub>v y" shows "Di i v \ Di i y" using assms and le_trans by (force simp: Di_def less_eq_def dij_def eij_def) lemma le_imp_max_x_ge: assumes "v \\<^sub>v y" and "length y \ n" shows "max_x v i \ max_x y i" using assms and le_imp_Di_subset and Min_Di_le [of i, OF _ _ assms(2)] by (metis Min.subset_imp Min_in emptyE finite_Di max_x_def order_refl subsetCE) end end diff --git a/thys/Group-Ring-Module/Algebra3.thy b/thys/Group-Ring-Module/Algebra3.thy --- a/thys/Group-Ring-Module/Algebra3.thy +++ b/thys/Group-Ring-Module/Algebra3.thy @@ -1,5558 +1,5558 @@ (** author Hidetsune Kobayashi Department of Mathematics Nihon University hikoba@math.cst.nihon-u.ac.jp May 3, 2004. April 6, 2007 (revised) chapter 3. Group Theory. Focused on Jordan Hoelder theorem (continued) section 5. products section 6. preliminary lemmas for Zassenhaus section 7. homomorphism section 8. gkernel, kernel of a group homomorphism section 9. image, image of a group homomorphism section 10. incuded homomorphisms section 11. Homomorphism therems section 12. isomorphisms section 13. Zassenhaus section 14. an automorphism group section 15. chain of groups I section 16. existence of reduced chain section 17. existence of reduced chain and composition series" section 18. chain of groups II section 19. Jordan Hoelder theorem **) theory Algebra3 imports Algebra2 begin section "Setproducts" definition commutators:: "_ \ 'a set" where "commutators G = {z. \ a \ carrier G. \b \ carrier G. ((a \\<^bsub>G\<^esub> b) \\<^bsub>G\<^esub> (\\<^bsub>G\<^esub> a)) \\<^bsub>G\<^esub> (\\<^bsub>G\<^esub> b) = z}" lemma (in Group) contain_commutator:"\G \ H; (commutators G) \ H\ \ G \ H" apply (rule cond_nsg[of "H"], assumption) apply (rule ballI)+ apply (frule_tac h = h in sg_subset_elem[of "H"], assumption, frule_tac a = h in i_closed, frule_tac a = a in i_closed, frule_tac a = a and b = h in mult_closed, assumption+, frule_tac a = "a \ h" and b = "\ a" in mult_closed, assumption+) apply (frule_tac a = "a \ h \ (\ a)" and b = "\ h" and c = h in tassoc, assumption+) apply (simp add:l_i r_unit) apply (rule_tac a = "a \ h \ \ a \ \ h \ h" and A = H and b = "a \ h \ \ a" in eq_elem_in, rule_tac x = "a \ h \ \ a \ \ h" and y = h in sg_mult_closed[of "H"], assumption, rule_tac c = "a \ h \ \ a \ \ h" in subsetD[of "commutators G" "H"], assumption, thin_tac "commutators G \ H", simp add:commutators_def, blast) apply assumption+ done definition s_top :: "[_ , 'a set, 'a set] \ 'a set " where "s_top G H K = {z. \x \ H. \y \ K. (x \\<^bsub>G\<^esub> y = z)}" abbreviation S_TOP :: "[('a, 'm) Group_scheme, 'a set, 'a set] \ 'a set" ("(3_ \\ _)" [66,67]66) where "H \\<^bsub>G\<^esub> K == s_top G H K" lemma (in Group) s_top_induced:"\G \ L; H \ L; K \ L\ \ H \\<^bsub>Gp G L\<^esub> K = H \\<^bsub>G\<^esub> K" by (simp add:s_top_def Gp_def) lemma (in Group) s_top_l_unit:"G \ K \ {\} \\<^bsub>G\<^esub> K = K" apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, erule bexE, frule_tac h = y in sg_subset_elem[of "K"], assumption+, simp add:l_unit) apply (rule subsetI, simp add:s_top_def) apply (frule_tac h = x in sg_subset_elem, assumption, frule_tac a = x in l_unit, blast) done lemma (in Group) s_top_r_unit:"G \ K \ K \\<^bsub>G\<^esub> {\} = K" apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, erule bexE, frule_tac h = xa in sg_subset_elem[of "K"], assumption+, simp add:r_unit) apply (rule subsetI, simp add:s_top_def, frule_tac h = x in sg_subset_elem[of "K"], assumption+, frule_tac a = x in r_unit, blast) done lemma (in Group) s_top_sub:"\G \ H; G \ K\ \ H \\<^bsub>G\<^esub> K \ carrier G" apply (rule subsetI) apply (simp add:s_top_def) apply (erule bexE)+ apply (frule_tac h = xa in sg_subset_elem [of "H"], assumption+, frule_tac h = y in sg_subset_elem[of "K"], assumption+, frule_tac a = xa and b = y in mult_closed, assumption+, simp) done lemma (in Group) sg_inc_set_mult:"\G \ L; H \ L; K \ L\ \ H \\<^bsub>G\<^esub> K \ L" apply (rule subsetI) apply (simp add:s_top_def, (erule bexE)+) apply (frule_tac c = xa in subsetD [of "H" "L"], assumption+, frule_tac c = y in subsetD [of "K" "L"], assumption+, frule_tac x = xa and y = y in sg_mult_closed[of "L"], assumption+) apply simp done lemma (in Group) s_top_sub1:"\H \ (carrier G); K \ (carrier G)\ \ H \\<^bsub>G\<^esub> K \ carrier G" apply (rule subsetI) apply (simp add:s_top_def) apply (erule bexE)+ apply (frule_tac c = xa in subsetD[of "H" "carrier G"], assumption+, frule_tac c = y in subsetD[of "K" "carrier G"], assumption+, frule_tac a = xa and b = y in mult_closed, assumption+, simp) done lemma (in Group) s_top_elem:"\G \ H; G \ K; a \ H; b \ K\ \ a \ b \ H \\<^bsub>G\<^esub> K" by (simp add:s_top_def, blast) lemma (in Group) s_top_elem1:"\H \ carrier G; K \ carrier G; a \ H; b \ K\ \ a \ b \ H \\<^bsub>G\<^esub> K " by (simp add:s_top_def, blast) lemma (in Group) mem_s_top:"\H \ carrier G; K \ carrier G; u \ H \\<^bsub>G\<^esub> K\ \ \a \ H. \b \ K. (a \ b = u)" by (simp add:s_top_def) lemma (in Group) s_top_mono:"\H \ carrier G; K \ carrier G; H1 \ H; K1 \ K\ \ H1 \\<^bsub>G\<^esub> K1 \ H \\<^bsub>G\<^esub> K" by (rule subsetI, simp add:s_top_def, blast) lemma (in Group) s_top_unit_closed:"\G \ H; G \ K\ \ \ \ H \\<^bsub>G\<^esub> K" apply (frule sg_unit_closed [of "H"], frule sg_unit_closed [of "K"]) apply (cut_tac unit_closed, frule l_unit[of "\"]) apply (simp add:s_top_def, blast) done lemma (in Group) s_top_commute:"\G \ H; G \ K; K \\<^bsub>G\<^esub> H = H \\<^bsub>G\<^esub> K; u \ H \\<^bsub>G\<^esub> K; v \ H \\<^bsub>G\<^esub> K\ \ u \ v \ H \\<^bsub>G\<^esub> K" apply (frule sg_subset[of "H"], frule sg_subset[of "K"], frule mem_s_top[of "H" "K" "u"], assumption+, (erule bexE)+, frule mem_s_top[of "H" "K" "v"], assumption+, (erule bexE)+) apply (rotate_tac 4, frule sym, thin_tac "a \ b = u", frule sym, thin_tac "aa \ ba = v", simp, thin_tac "u = a \ b", thin_tac "v = aa \ ba") apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+, frule_tac h = aa in sg_subset_elem[of "H"], assumption+, frule_tac h = b in sg_subset_elem[of "K"], assumption+, frule_tac h = ba in sg_subset_elem[of "K"], assumption+) apply (simp add:tOp_assocTr41[THEN sym], simp add:tOp_assocTr42) apply (frule_tac a = b and b = aa in s_top_elem1[of "K" "H"], assumption+, simp, thin_tac "K \\<^bsub>G\<^esub> H = H \\<^bsub>G\<^esub> K") apply (frule_tac u = "b \ aa" in mem_s_top[of "H" "K"], assumption+, (erule bexE)+, frule sym, thin_tac "ab \ bb = b \ aa", simp, thin_tac "b \ aa = ab \ bb") apply (frule_tac h = ab in sg_subset_elem[of "H"], assumption+, frule_tac h = bb in sg_subset_elem[of "K"], assumption+) apply (simp add:tOp_assocTr42[THEN sym], simp add:tOp_assocTr41) apply (frule_tac x = a and y = ab in sg_mult_closed[of "H"], assumption+, frule_tac x = bb and y = ba in sg_mult_closed[of "K"], assumption+, simp add:s_top_elem1) done lemma (in Group) s_top_commute1:"\G \ H; G \ K; K \\<^bsub>G\<^esub> H = H \\<^bsub>G\<^esub> K; u \ H \\<^bsub>G\<^esub> K\ \ (\ u) \ H \\<^bsub>G\<^esub> K" apply (frule sg_subset[of "H"], frule sg_subset[of "K"], frule mem_s_top[of "H" "K" "u"], assumption+, (erule bexE)+) apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+, frule_tac h = b in sg_subset_elem[of "K"], assumption+, frule_tac a = a and b = b in i_ab, assumption+, rotate_tac 4, frule sym, thin_tac "a \ b = u", simp, thin_tac "\ (a \ b) = \ b \ \ a") apply (frule_tac x = a in sg_i_closed[of "H"], assumption+, frule_tac x = b in sg_i_closed[of "K"], assumption+, frule_tac a = "\ b" and b = "\ a" in s_top_elem1[of "K" "H"], assumption+, simp) done lemma (in Group) s_top_commute_sg:"\G \ H; G \ K; K \\<^bsub>G\<^esub> H = H \\<^bsub>G\<^esub> K\ \ G \ (H \\<^bsub>G\<^esub> K)" apply (subst sg_def) apply (frule s_top_unit_closed[of "H" "K"], assumption, simp add:nonempty, simp add:s_top_sub) apply ((rule ballI)+, frule_tac u = b in s_top_commute1[of "H" "K"], assumption+, rule_tac u = a and v = "\ b" in s_top_commute[of "H" "K"], assumption+) done lemma (in Group) s_top_assoc:"\G \ H; G \ K; G \ L\ \ (H \\<^bsub>G\<^esub> K) \\<^bsub>G\<^esub> L = H \\<^bsub>G\<^esub> (K \\<^bsub>G\<^esub> L)" apply (rule equalityI) apply (rule subsetI, simp add:s_top_def) apply (erule exE) apply (erule conjE) apply (erule bexE)+ apply (rotate_tac -1, frule sym, thin_tac "xb \ ya = xa", simp, thin_tac "xa = xb \ ya", frule sym, thin_tac "xb \ ya \ y = x", simp) apply (frule_tac h = xb in sg_subset_elem[of "H"], assumption+, frule_tac h = y in sg_subset_elem[of "L"], assumption+, frule_tac h = ya in sg_subset_elem[of "K"], assumption+, simp add:tassoc, blast) apply (rule subsetI, simp add:s_top_def, erule bexE, erule exE, erule conjE, (erule bexE)+, rotate_tac -1, frule sym, thin_tac "xb \ ya = y", simp, thin_tac "y = xb \ ya") apply (frule_tac h = xa in sg_subset_elem[of "H"], assumption+, frule_tac h = ya in sg_subset_elem[of "L"], assumption+, frule_tac h = xb in sg_subset_elem[of "K"], assumption+, simp add:tassoc[THEN sym], frule sym, thin_tac "xa \ xb \ ya = x", simp, blast) done lemma (in Group) s_topTr6:"\G \ H1; G \ H2; G \ K; H1 \ K\ \ (H1 \\<^bsub>G\<^esub> H2) \ K = H1 \\<^bsub>G\<^esub> (H2 \ K)" apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, erule conjE, (erule bexE)+, frule sym, thin_tac "xa \ y = x", simp, frule_tac c = xa in subsetD[of "H1" "K"], assumption+, frule_tac x = "xa \ y" in inEx[of _ "K"], erule bexE, frule_tac x = xa in sg_i_closed[of "K"], assumption+, frule_tac x = "\ xa" and y = ya in sg_mult_closed[of "K"], assumption+, simp) apply (frule_tac h = xa in sg_subset_elem[of "K"], assumption+, frule_tac h = "\ xa" in sg_subset_elem[of "K"], assumption+, frule_tac h = y in sg_subset_elem[of "H2"], assumption+, simp add:tassoc[THEN sym] l_i l_unit) apply blast apply (rule subsetI, simp add:s_top_def, (erule bexE)+, simp, frule sym, thin_tac "xa \ y = x", simp, frule_tac c = xa in subsetD[of "H1" "K"], assumption+, frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+, simp) apply blast done lemma (in Group) s_topTr6_1:"\G \ H1; G \ H2; G \ K; H2 \ K\ \ (H1 \\<^bsub>G\<^esub> H2) \ K = (H1 \ K) \\<^bsub>G\<^esub> H2" apply (rule equalityI) apply (rule subsetI) apply (simp add:s_top_def, erule conjE, (erule bexE)+) apply (frule_tac c = y in subsetD [of "H2" "K"], assumption+) apply (frule_tac x = y in sg_i_closed [of "K"], assumption) apply (frule_tac h = xa in sg_subset_elem[of "H1"], assumption+, frule_tac h = x in sg_subset_elem[of "K"], assumption+, frule_tac h = y in sg_subset_elem[of "K"], assumption+, frule_tac h = "\ y" in sg_subset_elem[of "K"], assumption+, frule sym, thin_tac "xa \ y = x", frule_tac x = x and y = "\ y" in sg_mult_closed[of "K"], assumption+, simp add:tassoc r_i r_unit, blast) apply (rule subsetI, simp add:s_top_def, (erule bexE)+, simp, erule conjE, frule_tac c = y in subsetD[of "H2" "K"], assumption+, frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+, simp, blast) done lemma (in Group) l_sub_smult:"\G \ H; G \ K\ \ H \ H \\<^bsub>G\<^esub> K" apply (rule subsetI, simp add:s_top_def) apply (frule sg_unit_closed[of "K"], frule_tac h = x in sg_subset_elem[of "H"], assumption+, frule_tac a = x in r_unit) apply blast done lemma (in Group) r_sub_smult:"\G \ H; G \ K\ \ K \ H \\<^bsub>G\<^esub> K" apply (rule subsetI, simp add:s_top_def) apply (frule sg_unit_closed[of "H"], frule_tac h = x in sg_subset_elem[of "K"], assumption+, frule_tac a = x in l_unit) apply blast done lemma (in Group) s_topTr8:"G \ H \ H = H \\<^bsub>G\<^esub> H" apply (frule l_sub_smult[of "H" "H"], assumption) apply (rule equalityI, assumption) apply (rule subsetI) apply (thin_tac "H \ H \\<^bsub>G\<^esub> H", simp add:s_top_def, (erule bexE)+) apply (frule_tac x = xa and y = y in sg_mult_closed[of "H"], assumption+, simp) done section "Preliminary lemmas for Zassenhaus" lemma (in Group) Gp_sg_subset:"\G \ H; Gp G H \ K\ \ K \ H" by (frule Group_Gp[of "H"], frule Group.sg_subset[of "\H" "K"], assumption, thin_tac "(\H) \ K", thin_tac "Group (\H)", simp add:Gp_def) lemma (in Group) inter_Gp_nsg:"\G \ N; G \ H \ \ (\H) \ (H \ N)" apply (frule Group_Gp[of "H"], rule Group.cond_nsg[of "Gp G H" "H \ N"], assumption+, frule nsg_sg[of "N"], frule inter_sgs[of "H" "N"], assumption+, rule sg_sg [of "H" "H \ N"], assumption+) apply (rule subsetI, simp) apply ((rule ballI)+, simp, simp add:Gp_carrier, simp add:Gp_mult_induced[of "H"], simp add:sg_i_induced[of "H"]) apply (erule conjE, frule_tac x = a in sg_i_closed[of "H"], assumption+, frule_tac x = a and y = h in sg_mult_closed, assumption+, simp add:Gp_mult_induced[of "H"], simp add:sg_mult_closed) apply (frule_tac h = a in sg_subset_elem[of "H"], assumption+, simp add:nsgPr1_1[of "N"]) done lemma (in Group) ZassenhausTr0:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ Gp G (H \ K) \ (H \ K1)" apply (frule inter_sgs[of "H" "K"], assumption, frule inter_sgs[of "H" "K1"], assumption, frule Group_Gp[of "H"], frule Group_Gp[of "K"], frule Group.nsg_sg[of "\H" "H1"], assumption+, frule Group.nsg_sg[of "\K" "K1"], assumption+) apply (rule Group.cond_nsg[of "\(H \ K)" "H \ K1"], simp add:Group_Gp[of "H \ K"]) apply (rule sg_sg[of "H \ K" "H \ K1"], assumption+) apply (frule Gp_sg_subset[of "K" "K1"], assumption+, rule subsetI, simp add:subsetD) apply ((rule ballI)+, simp) apply (frule Gp_sg_subset[of "K" "K1"], assumption+, erule conjE, frule_tac c = h in subsetD[of "K1" "K"], assumption+) apply (rule conjI) apply (simp only:Gp_carrier, subst Gp_mult_induced1[of "H" "K"], assumption+, simp, simp only:sg_i_induced1) apply (frule_tac a = a and b = h in Group.mult_closed[of "\H"], simp add:Gp_carrier, simp add:Gp_carrier, simp only:Gp_carrier) apply (frule_tac a = a in Group.i_closed[of "\H"], simp add:Gp_carrier) apply (simp add:Gp_mult_induced1[of "H" "K"], simp add:Gp_carrier, subst Gp_mult_induced1[of "H" "K"], assumption+, simp add:Gp_mult_induced sg_mult_closed, simp add:sg_i_induced sg_i_closed) apply (simp add:Gp_mult_induced sg_i_induced, simp add:sg_mult_closed) apply (subst Gp_mult_induced2[of "H" "K"], assumption+, simp add:Gp_carrier, simp, subst sg_i_induced2, assumption+, simp add:Gp_carrier) apply (frule_tac a = a and b = h in Group.mult_closed[of "\K"], simp add:Gp_carrier, simp add:Gp_carrier, frule_tac a = a in Group.i_closed[of "\K"], simp add:Gp_carrier) apply (subst Gp_mult_induced2, assumption+, simp add:Gp_carrier, simp add:Gp_mult_induced sg_mult_closed, simp add:Gp_carrier, simp add:sg_i_induced sg_i_closed) apply (rule_tac a = a and h = h in Group.nsgPr1_1[of "\K" "K1"], assumption+, simp add:Gp_carrier, assumption) done lemma (in Group) lcs_sub_s_mult:"\G \ H; G \ N; a \ H\ \ a \ N \ H \\<^bsub>G\<^esub> N" apply (rule subsetI) apply (simp add:lcs_def s_top_def, blast) done lemma (in Group) rcs_sub_smult:"\G \ H; G \ N; a \ H\ \ N \ a \ N \\<^bsub>G\<^esub> H" apply (rule subsetI) apply (simp add:rcs_def s_top_def, blast) done lemma (in Group) smult_commute_sg_nsg:"\G \ H; G \ N\ \ H \\<^bsub>G\<^esub> N = N \\<^bsub>G\<^esub> H" apply (frule nsg_sg[of "N"]) apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, (erule bexE)+, frule_tac h = xa in sg_subset_elem, assumption+, frule_tac h = y in sg_subset_elem, assumption, frule_tac a = xa and b = y in mult_closed, assumption, frule_tac a = xa in i_closed, frule_tac a = "xa \ y" and b = "\ xa" and c = xa in tassoc, assumption+, frule sym, thin_tac "xa \ y = x", simp, thin_tac "x = xa \ y", simp add:l_i r_unit, frule_tac a = xa and h = y in nsgPr1_1[of "N"], assumption+) apply blast apply (rule subsetI) apply (simp add:s_top_def, (erule bexE)+, frule_tac h = xa in sg_subset_elem, assumption+, frule_tac h = y in sg_subset_elem, assumption, frule_tac a = xa and b = y in mult_closed, assumption, frule_tac a = y in i_closed, frule_tac a = y and b = "\ y" and c = "xa \ y" in tassoc, assumption+, frule sym, thin_tac "xa \ y = x", simp, thin_tac "x = xa \ y", simp add:r_i l_unit, frule_tac a = y and h = xa in nsgPr2[of "N"], assumption+, frule sym, thin_tac "xa \ y = y \ (\ y \ (xa \ y))") apply blast done lemma (in Group) smult_sg_nsg:"\G \ H; G \ N\ \ G \ H \\<^bsub>G\<^esub> N" apply (frule smult_commute_sg_nsg[of "H" "N"], assumption+, frule nsg_sg[of "N"], rule s_top_commute_sg[of "H" "N"], assumption+, rule sym, assumption) done lemma (in Group) smult_nsg_sg:"\G \ H; G \ N\ \ G \ N \\<^bsub>G\<^esub> H" apply (frule smult_commute_sg_nsg[THEN sym, of "H" "N"], assumption+) apply (simp add:smult_sg_nsg) done lemma (in Group) Gp_smult_sg_nsg:"\G \ H; G \ N\ \ Group (Gp G (H \\<^bsub>G\<^esub> N))" apply (frule smult_sg_nsg[of "H" "N"], assumption+) apply (simp add:Group_Gp) done lemma (in Group) N_sg_HN:"\G \ H; G \ N\ \ Gp G (H \\<^bsub>G\<^esub> N) \ N" apply (frule smult_sg_nsg[of "H" "N"], assumption+, frule nsg_sg[of "N"], frule r_sub_smult[of "H" "N"], assumption+) apply (rule sg_sg[of "H \\<^bsub>G\<^esub> N" "N"], assumption+) done lemma (in Group) K_absorb_HK:"\G \ H; G \ K; H \ K\ \ H \\<^bsub>G\<^esub> K = K" apply (frule r_sub_smult[of "H" "K"], assumption+) apply (rule equalityI) apply (thin_tac "K \ H \\<^bsub>G\<^esub> K", rule subsetI, simp add:s_top_def, (erule bexE)+, frule_tac c = xa in subsetD[of "H" "K"], assumption+, frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+, simp) apply assumption done lemma (in Group) nsg_Gp_nsg:"\G \ H; G \ N; N \ H\ \ Gp G H \ N" apply (frule Group_Gp[of "H"], frule nsg_sg[of "N"], frule sg_sg[of "H" "N"], assumption+, rule Group.cond_nsg[of "\H" "N"], assumption+) apply ((rule ballI)+, frule_tac c = h in subsetD[of "N" "H"], assumption+, simp add:Gp_carrier, simp add:Gp_mult_induced[of "H"] sg_i_induced[of "H"] sg_mult_closed sg_i_closed) apply (rule_tac a = a and h = h in nsgPr1_1[of "N"], assumption+, rule_tac h = a in sg_subset_elem[of "H"], assumption+) done lemma (in Group) Gp_smult_nsg:"\G \ H; G \ N\ \ Gp G (H \\<^bsub>G\<^esub> N) \ N" apply (frule smult_sg_nsg[of "H" "N"], assumption+, frule nsg_sg[of "N"], frule N_sg_HN[of "H" "N"], assumption+, frule Gp_smult_sg_nsg[of "H" "N"], assumption+, rule Group.cond_nsg[of "\(H \\<^bsub>G\<^esub> N)" "N"], assumption+) apply ((rule ballI)+, frule_tac a = a in Group.i_closed[of "\(H \\<^bsub>G\<^esub> N)"], assumption+, simp add:Gp_carrier) apply (frule r_sub_smult[of "H" "N"], assumption+, frule_tac c = h in subsetD[of "N" "H \\<^bsub>G\<^esub> N"], assumption+, simp add:Gp_mult_induced[of "H \\<^bsub>G\<^esub> N"] sg_i_induced[of "H \\<^bsub>G\<^esub> N"]) apply (frule_tac x = a and y = h in sg_mult_closed[of "H \\<^bsub>G\<^esub> N"], assumption+, simp add:Gp_mult_induced) apply (rule_tac a = a and h = h in nsgPr1_1[of "N"], assumption+, frule sg_subset[of "H \\<^bsub>G\<^esub> N"], frule_tac c = a in subsetD[of "H \\<^bsub>G\<^esub> N" "carrier G"], assumption+) done lemma (in Group) Gp_smult_nsg1:"\G \ H; G \ N\ \ Gp G (N \\<^bsub>G\<^esub> H) \ N" apply (simp add:smult_commute_sg_nsg[THEN sym, of "H" "N"], simp only:Gp_smult_nsg) done lemma (in Group) ZassenhausTr2_3:"\G \ H; G \ H1; Gp G H \ H1\ \ H1 \ H" apply (frule Group_Gp[of "H"], frule Group.nsg_sg[of "\H" "H1"], assumption, frule Group.sg_subset[of "\H" "H1"], assumption, simp add:Gp_carrier) done lemma (in Group) ZassenhausTr2_4:"\G \ H; G \ H1; Gp G H \ H1; h \ H; h1 \ H1\ \ h \ h1 \ (\ h) \ H1" apply (frule Group_Gp[of "H"]) apply (frule_tac a = h and h = h1 in Group.nsgPr1_1[of "\H" "H1"], assumption) apply (simp add:Gp_carrier) apply assumption apply (simp add:Gp_def) done lemma (in Group) ZassenhausTr1:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ H1 \\<^bsub>G\<^esub> (H \ K1) = (H \ K1) \\<^bsub>G\<^esub> H1" apply (frule Group_Gp[of "H"], frule Group.nsg_sg[of "\H" "H1"], assumption, frule Group.sg_subset[of "\H" "H1"], assumption, simp add:Gp_carrier) apply (frule Group_Gp[of "K"], frule Group.nsg_sg[of "\K" "K1"], assumption, frule Group.sg_subset[of "\K" "K1"], assumption, simp add:Gp_carrier) apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, (erule bexE)+, frule_tac h = xa in sg_subset_elem[of "H1"], assumption+, frule_tac h = y in sg_subset_elem[of "H"], simp, frule_tac a = y in i_closed, frule_tac a = xa and b = y in mult_closed, assumption+, frule_tac a1 = y and b1 = "\ y" and c1 = "xa \ y" in tassoc[THEN sym], assumption+) apply (frule sym, thin_tac "xa \ y = x", simp, thin_tac "x = xa \ y", simp add:r_i l_unit, frule_tac x = y in sg_i_closed[of "H"], simp) apply (frule_tac a1 = "\ y" and b1 = xa and c1 = y in tassoc[THEN sym], assumption+, simp, thin_tac "\ y \ (xa \ y) = \ y \ xa \ y") apply (frule_tac h = "\ y" and ?h1.0 = xa in ZassenhausTr2_4[of "H" "H1"], assumption+, simp add:iop_i_i) apply blast apply (rule subsetI) apply (simp add:s_top_def, (erule bexE)+, frule_tac h = xa in sg_subset_elem[of "H"], simp, frule_tac h = y in sg_subset_elem[of "H1"], assumption, frule sym, thin_tac "xa \ y = x", simp, thin_tac "x = xa \ y") apply (frule_tac a = xa in i_closed, frule_tac a = xa and b = y in mult_closed, assumption+, frule_tac a = "xa \ y" and b = "\ xa" and c = xa in tassoc, assumption+) apply (simp add:l_i r_unit, frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "H" "H1"], assumption+, simp, assumption, blast) done lemma (in Group) ZassenhausTr1_1:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ G \ (H1 \\<^bsub>G\<^esub> (H \ K1))" apply (rule s_top_commute_sg, assumption) apply (simp add:inter_sgs[of "H" "K1"]) apply (rule ZassenhausTr1 [THEN sym, of "H" "H1" "K" "K1"], assumption+) done lemma (in Group) ZassenhausTr2:"\G \ H; G \ H1; G \ K; Gp G H \ H1\ \ H1 \\<^bsub>G\<^esub> (H \ K) = (H \ K) \\<^bsub>G\<^esub> H1" apply (frule special_nsg_G1[of "K"]) apply (simp add: ZassenhausTr1 [of "H" "H1" "K" "K"]) done lemma (in Group) ZassenhausTr2_1:"\G \ H; G \ H1; G \ K; Gp G H \ H1\ \ G \ H1 \\<^bsub>G\<^esub> (H \ K)" apply (frule ZassenhausTr2 [of "H" "H1" "K"], assumption+, frule inter_sgs [of "H" "K"], assumption+, rule s_top_commute_sg, assumption+) apply (rule sym, assumption) done lemma (in Group) ZassenhausTr2_2:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ H1 \\<^bsub>G\<^esub> (H \ K1) \ H1 \\<^bsub>G\<^esub> (H \ K)" apply (frule Group_Gp[of "K"], frule Group.nsg_sg[of "\K" "K1"], assumption, frule Group.sg_subset[of "\K" "K1"], assumption, simp add:Gp_carrier, frule Group_Gp[of "H"], frule Group.nsg_sg[of "\H" "H1"], assumption, frule Group.sg_subset[of "\H" "H1"], assumption, simp add:Gp_carrier, frule sg_subset[of "H"], frule sg_subset[of "K"]) apply (rule s_top_mono[of "H1" "H \ K" "H1" "H \ K1"], rule subset_trans[of "H1" "H" "carrier G"], assumption+, blast, simp, blast) done lemma (in Group) ZassenhausTr2_5:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1; a\ H1; b \ H \ K1; c \ H1\ \ a \ b \ c \ H1 \\<^bsub>G\<^esub> (H \ K1)" apply (simp, erule conjE) apply (frule sg_subset_elem[of "H1" "a"], assumption+, frule sg_subset_elem[of "H1" "c"], assumption+, frule sg_subset_elem[of "H" "b"], assumption+, frule i_closed[of "b"], frule mult_closed[of "a" "b"], assumption+, frule mult_closed[of "a \ b" "c"], assumption+, frule tassoc[of "a \ b \ c" "\ b" "b"], assumption+, simp add:l_i r_unit) apply (rule eq_elem_in[of "a \ b \ c \ \ b \ b" "H1 \\<^bsub>G\<^esub> H \ K1" "a \ b \ c"], thin_tac "a \ b \ c \ \ b \ b = a \ b \ c", frule inter_sgs[of "H" "K1"], assumption+, rule s_top_elem[of "H1" "H \ K1" "a \ b \ c \ \ b " "b"], assumption+, subst tOp_assocTr42, assumption+, frule mult_closed[of "b" "c"], assumption+, simp add:tassoc[of "a" "b \ c" "\ b"]) apply (rule sg_mult_closed[of "H1" "a" "b \ c \ \ b"], assumption+, rule ZassenhausTr2_4[of "H" "H1" "b" "c"], assumption+) apply blast apply assumption done lemma (in Group) ZassenhausTr2_6:"\u \ carrier G; v \ carrier G; x \ carrier G; y \ carrier G\ \ (u \ v) \ (x \ y) \ (\ (u \ v)) = u \ v \ x \ (\ v) \ (v \ y \ (\ v)) \ (\ u)" apply (simp add:i_ab) apply (frule i_closed[of "u"], frule i_closed[of "v"]) apply (frule mult_closed[of "u" "v"], assumption+, frule mult_closed[of "u \ v" "x"], assumption+, frule mult_closed[of "v" "y"], assumption+, frule mult_closed[of "v \ y" "\ v"], assumption+, frule mult_closed[of "u \ v \ x" "\ v"], assumption+, simp add:tOp_assocTr42[THEN sym, of "u \ v \ x \ \ v " "v \ y" "\ v" "\ u"]) apply (frule mult_closed[of "x" "y"], assumption+, frule mult_closed[of "u \ v" "x \ y"], assumption+) apply (simp add:tassoc[THEN sym, of "u \ v \ (x \ y)" "\ v" "\ u"]) apply (rule r_mult_eqn[of _ _ "\ u"], rule mult_closed[of "u \ v \ (x \ y)" "\ v"], assumption+, (rule mult_closed)+, assumption+) apply (rule r_mult_eqn[of _ _ "\ v"], assumption+, (rule mult_closed)+, assumption+) apply (simp add:tOp_assocTr41[THEN sym, of "u \ v \ x" "\ v" "v" "y"], simp add:tOp_assocTr42[of "u \ v \ x" "\ v" "v" "y"], simp add:l_i r_unit) apply (simp add:tOp_assocTr41) done lemma (in Group) ZassenhausTr2_7:"\a \ carrier G; x \ carrier G; y \ carrier G\ \ a \ ( x \ y) \ (\ a) = a \ x \ (\ a) \ (a \ y \ (\ a))" apply (frule i_closed[of "a"], frule mult_closed[of "a" "y"], assumption+, frule mult_closed[of "a \ y" "\ a"], assumption+) apply (simp add:tOp_assocTr41[of "a" "x" "\ a" "a \ y \ (\ a)"], simp add:tassoc[THEN sym, of "\ a" "a \ y" "\ a"], simp add:tassoc[THEN sym, of "\ a" "a" "y"] l_i l_unit, simp add:tOp_assocTr41[THEN sym], simp add:tOp_assocTr42[THEN sym]) done lemma (in Group) ZassenhausTr3:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ Gp G (H1 \\<^bsub>G\<^esub> (H \ K)) \ (H1 \\<^bsub>G\<^esub> (H \ K1))" apply (frule ZassenhausTr2_1 [of "H" "H1" "K"], assumption+, frule ZassenhausTr2_1 [of "H" "H1" "K1"], assumption+, frule ZassenhausTr2_2 [of "H" "H1" "K" "K1"], assumption+, frule sg_sg [of "H1 \\<^bsub>G\<^esub> (H \ K)" "H1 \\<^bsub>G\<^esub> (H \ K1)"], assumption+, frule Group_Gp[of "H1 \\<^bsub>G\<^esub> H \ K"]) apply (rule Group.cond_nsg[of "\(H1 \\<^bsub>G\<^esub> H \ K)" "H1 \\<^bsub>G\<^esub> H \ K1"], assumption+, (rule ballI)+, simp add:Gp_carrier) apply (frule_tac c = h in subsetD[of "H1 \\<^bsub>G\<^esub> H \ K1" "H1 \\<^bsub>G\<^esub> H \ K"], assumption+, simp add:Gp_mult_induced[of "H1 \\<^bsub>G\<^esub> H \ K"], simp add:sg_i_induced[of "H1 \\<^bsub>G\<^esub> H \ K"], frule_tac x = a in sg_i_closed[of "H1 \\<^bsub>G\<^esub> H \ K"], assumption+, frule_tac x = a and y = h in sg_mult_closed[of "H1 \\<^bsub>G\<^esub> H \ K"], assumption+, simp add:Gp_mult_induced[of "H1 \\<^bsub>G\<^esub> H \ K"], thin_tac "\ a \ H1 \\<^bsub>G\<^esub> H \ K", thin_tac "a \ h \ H1 \\<^bsub>G\<^esub> H \ K") apply (simp add:s_top_def[of "G" "H1" "H \ K"], (erule bexE)+, simp add:s_top_def[of "G" "H1" "H \ K1"], fold s_top_def, (erule bexE)+, thin_tac "xa \ ya = h", (erule conjE)+, thin_tac "xa \ H1", thin_tac "ya \ H", frule sym, thin_tac "x \ y = a", frule sym, thin_tac "xb \ yb = h", simp, (erule conjE)+, thin_tac "a = x \ y", thin_tac "h = xb \ yb") apply (frule_tac h = x in sg_subset_elem[of "H1"], assumption+, frule_tac h = y in sg_subset_elem[of "H"], assumption+, frule_tac h = xb in sg_subset_elem[of "H1"], assumption+, frule_tac h = yb in sg_subset_elem[of "H"], assumption+, subst ZassenhausTr2_6, assumption+) apply (frule_tac a = y and b = xb in mult_closed, assumption+, frule_tac a = y in i_closed, frule_tac a = "y \ xb" and b = "\ y" in mult_closed, assumption+, frule_tac a = x and b = "y \ xb \ \ y" in mult_closed, assumption+, frule_tac a = y and b = yb in mult_closed, assumption+, frule_tac a = "y \ yb" and b = "\ y" in mult_closed, assumption+, frule_tac a = "x \ y \ xb \ \ y" and b = "y \ yb \ \ y" and c = "\ x" in ZassenhausTr2_5[of "H" "H1" "K" "K1"], assumption+, frule_tac a = x and b = y in mult_closed, assumption+, frule_tac a = "x \ y" and b = xb and c = "\ y" in tassoc, assumption+, simp, thin_tac "x \ y \ xb \ \ y = x \ y \ (xb \ \ y)", frule_tac a = xb and b = "\ y" in mult_closed, assumption+, simp add:tassoc) apply (rule_tac x = x and y = "y \ (xb \ \ y)" in sg_mult_closed, assumption+, simp add:tassoc[THEN sym], rule_tac h = y and ?h1.0 = xb in ZassenhausTr2_4[of "H" "H1"], assumption+) apply (frule_tac x = y and y = yb in sg_mult_closed[of "H"], assumption+, frule_tac x = y in sg_i_closed[of "H"], assumption+, frule_tac x = "y \ yb" and y = "\ y" in sg_mult_closed[of "H"], assumption+, simp, rule_tac h = y and ?h1.0 = yb in ZassenhausTr2_4[of "K" "K1"], assumption+, rule_tac x = x in sg_i_closed[of "H1"], assumption+) apply (simp add:s_top_def[of "G" "H1" "H \ K1"]) done lemma (in Group) ZassenhausTr3_2:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ G \ H1 \\<^bsub>G\<^esub> (H \ K1) \\<^bsub>G\<^esub> (H \ K)" apply (frule s_top_assoc[of "H1" "H \ K1" "H \ K"], (simp add:inter_sgs)+, frule inter_sgs[of "H" "K1"], assumption+, frule inter_sgs[of "H" "K"], assumption+) apply (frule K_absorb_HK[of "H \ K1" "H \ K"], assumption+, frule ZassenhausTr2_3[of "K" "K1"], assumption+, blast, simp, simp add:ZassenhausTr2_1) done lemma (in Group) ZassenhausTr3_3:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ (H1 \ K) \\<^bsub>G\<^esub> (H \ K1) = (K1 \ H) \\<^bsub>G\<^esub> (K \ H1)" apply (rule equalityI) apply (rule subsetI, simp add:s_top_def, (erule bexE)+) apply (frule sym, thin_tac "xa \ y = x", simp, (erule conjE)+) apply (frule_tac h = xa in sg_subset_elem[of "H1"], assumption+, frule_tac h = y in sg_subset_elem[of "K1"], assumption+, frule_tac a = xa in i_closed, frule_tac a = xa and b = y and c = "\ xa" and d = xa in tOp_assocTr41, assumption+, frule_tac a = xa and b = y in mult_closed, assumption, simp add:l_i r_unit) apply (frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "K" "K1"], assumption+) apply (frule ZassenhausTr2_3[of "H" "H1"], assumption+, frule_tac c = xa in subsetD[of "H1" "H"], assumption+) apply (frule_tac x = xa and y = y in sg_mult_closed[of "H"], assumption+) apply (frule_tac x = xa in sg_i_closed[of "H"], assumption+, frule_tac x = "xa \ y" and y = "\ xa" in sg_mult_closed[of "H"], assumption+) apply blast apply (rule subsetI, simp add:s_top_def, (erule bexE)+) apply (frule sym, thin_tac "xa \ y = x", simp, (erule conjE)+) apply (frule_tac h = xa in sg_subset_elem[of "K1"], assumption+, frule_tac h = y in sg_subset_elem[of "H1"], assumption+, frule_tac a = xa in i_closed, frule_tac a = xa and b = y and c = "\ xa" and d = xa in tOp_assocTr41, assumption+, frule_tac a = xa and b = y in mult_closed, assumption, simp add:l_i r_unit) apply (frule_tac h = xa and ?h1.0 = y in ZassenhausTr2_4[of "H" "H1"], assumption+) apply (frule ZassenhausTr2_3[of "K" "K1"], assumption+) apply (frule_tac c = xa in subsetD[of "K1" "K"], assumption+) apply (frule_tac x = xa and y = y in sg_mult_closed[of "K"], assumption+) apply (frule_tac x = xa in sg_i_closed[of "K"], assumption+, frule_tac x = "xa \ y" and y = "\ xa" in sg_mult_closed[of "K"], assumption+) apply blast done lemma (in Group) ZassenhausTr3_4:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1; g \ H \ K; h \ H \ K1\ \ g \ h \ (\ g) \ H \ K1" apply (simp, (erule conjE)+) apply (frule_tac x = g and y = h in sg_mult_closed, assumption+, frule_tac x = g in sg_i_closed[of "H"], assumption+, simp add:sg_mult_closed[of "H" "g \ h" "\ g"]) apply (rule ZassenhausTr2_4[of "K" "K1" "g" "h"], assumption+) done lemma (in Group) ZassenhausTr3_5:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ (Gp G (H \ K)) \ (H1 \ K) \\<^bsub>G\<^esub> (H \ K1)" apply (frule inter_sgs[of "H" "K"], assumption, frule inter_sgs[of "H1" "K"], assumption, frule inter_sgs[of "K" "H"], assumption, frule inter_sgs[of "H" "K1"], assumption+) apply (frule ZassenhausTr3[of "H \ K" "H1 \ K" "K \ H" "H \ K1"], assumption+, frule ZassenhausTr0[of "K" "K1" "H" "H1"], assumption+, simp add:Int_commute, frule ZassenhausTr0[of "H" "H1" "K" "K1"], assumption+, simp add:Int_commute) apply (frule ZassenhausTr2_3 [of "K" "K1"], assumption+, frule ZassenhausTr2_3 [of "H" "H1"], assumption+) apply (simp add:Int_commute[of "K" "H"]) apply (cut_tac Int_mono[of "H" "H" "K1" "K"]) apply (cut_tac Int_mono[of "H1" "H" "K" "K"]) apply (simp only:Int_absorb1[of "H \ K1" "H \ K"], simp only:K_absorb_HK[of "H1 \ K" "H \ K"]) apply simp+ done lemma (in Group) ZassenhausTr4:"\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ (H1 \\<^bsub>G\<^esub> (H \ K1)) \\<^bsub>G\<^esub> (H1 \\<^bsub>G\<^esub> (H \ K)) = H1 \\<^bsub>G\<^esub> (H \ K)" apply (frule ZassenhausTr2 [of "H" "H1" "K"], assumption+, frule ZassenhausTr2 [of "H" "H1" "K1"], assumption+, frule ZassenhausTr1_1 [of "H" "H1" "K" "K1"], assumption+, frule ZassenhausTr2_1 [of "H" "H1" "K"], assumption+, frule ZassenhausTr2_2 [of "H" "H1" "K" "K1"], assumption+) apply (rule K_absorb_HK[of "H1 \\<^bsub>G\<^esub> H \ K1" "H1 \\<^bsub>G\<^esub> H \ K"], assumption+) done lemma (in Group) ZassenhausTr4_0: "\G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ H1 \\<^bsub>G\<^esub> (H \ K) = (H1 \\<^bsub>G\<^esub> (H \ K1)) \\<^bsub>G\<^esub> (H \ K)" apply (frule inter_sgs [of "H" "K1"], assumption+, frule inter_sgs [of "H" "K"], assumption+) apply (subst s_top_assoc [of "H1" "H \ K1" "H \ K"], assumption+, subst K_absorb_HK[of "H \ K1" "H \ K"], assumption+) apply (frule ZassenhausTr2_3[of "K" "K1"], assumption+, rule Int_mono[of "H" "H" "K1" "K"]) apply simp+ done lemma (in Group) ZassenhausTr4_1:"\G \ H; (Gp G H) \ H1; (Gp G H) \ (H \ K)\ \ (Gp G (H1 \\<^bsub>G\<^esub> (H \ K))) \ H1" apply (frule Group_Gp [of "H"], frule Group.nsg_sg[of "Gp G H" "H1"], assumption+, frule Group.Gp_smult_nsg1[of "\H" "H \ K" "H1"], assumption+, frule subg_sg_sg [of "H" "H1"], assumption+, frule Group.sg_subset[of "\H" "H1"], assumption, frule Group.sg_subset[of "\H" "H \ K"], assumption+, frule Group.smult_nsg_sg[of "\H" "H \ K" "H1"], assumption+, frule Group.s_top_sub[of "\H" "H1" "H \ K"], assumption+) apply (simp only: Gp_carrier s_top_induced [of "H" "H1" "H \ K"]) apply (frule subg_sg_sg[of "H" "H1 \\<^bsub>G\<^esub> H \ K"], assumption+, simp add:Gp_inherited[of "H1 \\<^bsub>G\<^esub> H \ K" "H"]) done section "Homomorphism" lemma gHom: "\Group F; Group G; f \ gHom F G ; x \ carrier F; y \ carrier F\ \ f (x \\<^bsub>F\<^esub> y) = (f x) \\<^bsub>G\<^esub> (f y)" apply (simp add: gHom_def) done lemma gHom_mem:"\Group F; Group G; f \ gHom F G ; x \ carrier F\ \ (f x) \ carrier G" apply (simp add:gHom_def, (erule conjE)+) apply (simp add:Pi_def) done lemma gHom_func:"\Group F; Group G; f \ gHom F G\ \ f \ carrier F \ carrier G" by (simp add:gHom_def) (* we have to check the composition of two ghom's *) lemma gHomcomp:"\Group F; Group G; Group H; f \ gHom F G; g \ gHom G H\ \ (g \\<^bsub>F\<^esub> f) \ gHom F H" apply (simp add:gHom_def) apply (erule conjE)+ apply (simp add:cmpghom_def composition) apply (rule ballI)+ apply (simp add:compose_def) apply (frule_tac x = x in funcset_mem [of "f" "carrier F" "carrier G"], assumption+) apply (frule_tac x = y in funcset_mem [of "f" "carrier F" "carrier G"], assumption+) apply (simp add:Group.mult_closed[of "F"]) done lemma gHom_comp_gsurjec:"\Group F; Group G; Group H; gsurj\<^bsub>F,G\<^esub> f; gsurj\<^bsub>G,H\<^esub> g\ \ gsurj\<^bsub>F,H\<^esub> (g \\<^bsub>F\<^esub> f)" apply (simp add:gsurjec_def, simp add:gHomcomp, (erule conjE)+) apply (simp add:cmpghom_def, simp add:gHom_def, (erule conjE)+, rule compose_surj, assumption+) done lemma gHom_comp_ginjec:"\Group F; Group G; Group H; ginj\<^bsub>F,G\<^esub> f; ginj\<^bsub>G,H\<^esub> g\ \ ginj\<^bsub>F,H\<^esub> (g \\<^bsub>F\<^esub> f)" apply (simp add:ginjec_def, simp add:gHomcomp, simp add:gHom_def, (erule conjE)+) apply (simp add:cmpghom_def, simp add:comp_inj) done lemma ghom_unit_unit:"\Group F; Group G; f \ gHom F G \ \ f (\\<^bsub>F\<^esub>) = \\<^bsub>G\<^esub>" apply (frule Group.unit_closed[of "F"]) apply (frule gHom [of "F" "G" "f" "\\<^bsub>F\<^esub>" "\\<^bsub>F\<^esub>"], assumption+) apply (simp add:Group.l_unit[of "F"]) apply (frule gHom_mem[of "F" "G" "f" "\\<^bsub>F\<^esub>"], assumption+) apply (frule sym) apply (rule Group.one_unique[of "G" "f \\<^bsub>F\<^esub>" "f \\<^bsub>F\<^esub>"], assumption+) done lemma ghom_inv_inv:"\Group F; Group G; f \ gHom F G ; x \ carrier F\ \ f (\\<^bsub>F\<^esub> x) = \\<^bsub>G\<^esub> (f x)" apply (frule Group.i_closed[of "F" "x"], assumption+, frule gHom [of "F" "G" "f" "\\<^bsub>F\<^esub> x" "x"], assumption+) apply (simp add:Group.l_i, simp add:ghom_unit_unit) apply (frule sym, frule gHom_mem[of "F" "G" "f" "x"], assumption+ , frule gHom_mem[of "F" "G" "f" "\\<^bsub>F\<^esub> x"], assumption+, rule Group.l_i_unique[THEN sym, of "G" "f x" "f (\\<^bsub>F\<^esub> x)"], assumption+) done lemma ghomTr3:"\Group F; Group G; f \ gHom F G ; x \ carrier F; y \ carrier F; f (x \\<^bsub>F\<^esub> (\\<^bsub>F\<^esub> y)) = \\<^bsub>G\<^esub> \ \ f x = f y" apply (frule Group.i_closed[of "F" "y"], assumption+, simp only:gHom ghom_inv_inv) apply (rule Group.r_div_eq[of "G" "f x" "f y"], assumption, (simp add:gHom_mem)+) done lemma iim_nonempty:"\Group F; Group G; f \ gHom F G; G \ K\ \ (iim F G f K) \ {}" apply (frule Group.sg_unit_closed[of "G" "K"], assumption+, frule Group.unit_closed[of "F"]) apply (frule ghom_unit_unit[of "F" "G" "f"], assumption+, simp add:iim_def, frule sym, thin_tac "f \\<^bsub>F\<^esub> = \\<^bsub>G\<^esub>", simp) apply blast done lemma ghomTr4:"\Group F; Group G; f \ gHom F G; G \ K\ \ F \ (iim F G f K)" apply (rule Group.sg_condition[of "F" "iim F G f K"], assumption+, rule subsetI, simp add:iim_def, simp add:iim_nonempty) apply ((rule allI)+, rule impI, erule conjE) apply (simp add:iim_def) apply (erule conjE)+ apply (frule_tac a = b in Group.i_closed[of "F"], assumption+, frule_tac a = a and b = "\\<^bsub>F\<^esub> b" in Group.mult_closed[of "F"], assumption+, simp) apply (simp add:gHom ghom_inv_inv) apply (frule_tac x = "f b" in Group.sg_i_closed[of "G" "K"], assumption+) apply (simp add:gHom_mem Group.sg_mult_closed) done lemma (in Group) IdTr0: "idmap (carrier G) \ gHom G G" apply (simp add:gHom_def) apply (simp add:idmap_def extensional_def) apply (simp add:Pi_def mult_closed) done abbreviation IDMAP ("(I\<^bsub>_\<^esub>)" [999]1000) where "I\<^bsub>F\<^esub> == idmap (carrier F)" abbreviation INVFUN ("(3Ifn _ _ _)" [88,88,89]88) where "Ifn F G f == invfun (carrier F) (carrier G) f" lemma IdTr1:"\Group F; x \ carrier F\ \ (I\<^bsub>F\<^esub>) x = x" apply (simp add:idmap_def) done lemma IdTr2:"Group F \ gbij\<^bsub>F,F\<^esub> (I\<^bsub>F\<^esub>)" apply (simp add:gbijec_def) apply (rule conjI) (* gsurjec *) apply (simp add:gsurjec_def) apply (rule conjI) apply (simp add:idmap_def gHom_def Group.mult_closed) apply (simp add:surj_to_def image_def idmap_def) (* ginjec *) apply (simp add:ginjec_def) apply (simp add:idmap_def gHom_def Group.mult_closed) done lemma Id_l_unit:"\Group G; gbij\<^bsub>G,G\<^esub> f\ \ I\<^bsub>G\<^esub> \\<^bsub>G\<^esub> f = f" apply (rule funcset_eq[of _ "carrier G"]) apply (simp add:cmpghom_def) apply (simp add:gbijec_def gsurjec_def gHom_def) apply (rule ballI) apply (simp add:cmpghom_def compose_def) apply (frule_tac x = x in gHom_mem[of "G" "G" "f"], assumption+) apply (simp add:gbijec_def gsurjec_def, assumption) apply (simp add:IdTr1) done section "Gkernel" lemma gkernTr1:"\Group F; Group G; f \ gHom F G; x \ gker\<^bsub>F,G\<^esub> f\ \ x \ carrier F" apply (simp add: gkernel_def) done lemma gkernTr1_1:"\Group F; Group G; f \ gHom F G\ \ gker\<^bsub>F,G\<^esub> f \ carrier F" by (rule subsetI, simp add:gkernTr1) lemma gkernTr2:"\Group F; Group G; f \ gHom F G; x \ gker\<^bsub>F,G\<^esub> f; y \ gker\<^bsub>F,G\<^esub> f\ \ (x \\<^bsub>F\<^esub> y) \ gker\<^bsub>F,G\<^esub> f" apply (simp add:gkernel_def) apply (simp add:gHom, (erule conjE)+) apply (simp add:Group.mult_closed, frule Group.unit_closed[of "G"], simp add:Group.l_unit[of "G"]) done lemma gkernTr3:"\Group F; Group G; f \ gHom F G ; x \ gker\<^bsub>F,G\<^esub> f\ \ (\\<^bsub>F\<^esub> x) \ gker\<^bsub>F,G\<^esub> f" apply (simp add:gkernel_def) (* thm ghom_inv_inv [of "F" "G" "f" "x"] *) apply (simp add:ghom_inv_inv [of "F" "G" "f" "x"]) apply (simp add:Group.i_closed) apply (simp add:Group.i_one) done lemma gkernTr6:"\Group F; Group G; f \ gHom F G\ \ (\\<^bsub>F\<^esub>) \ gker\<^bsub>F,G\<^esub> f" apply (simp add:gkernel_def) apply (simp add:Group.unit_closed ghom_unit_unit [of "F" "G" "f" ]) done lemma gkernTr7:"\Group F; Group G; f \ gHom F G \ \ F \ gker\<^bsub>F,G\<^esub> f" apply (rule Group.sg_condition[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+, rule subsetI, simp add:gkernTr1, frule gkernTr6[of "F" "G" "f"], assumption+, simp add:nonempty) apply ((rule allI)+, rule impI, erule conjE, frule_tac x = b in gkernTr3[of "F" "G" "f"], assumption+) apply (simp add:gkernTr2) done lemma gker_normal:"\Group F; Group G; f \ gHom F G\ \ F \ gker\<^bsub>F,G\<^esub> f" apply (rule Group.cond_nsg[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption) apply (simp add:gkernTr7) apply (rule ballI)+ apply (simp add:gkernel_def, erule conjE) apply (frule_tac a = a in Group.i_closed[of "F"], assumption) apply (subst gHom [of "F" "G" "f" _], assumption+) apply (simp add:Group.mult_closed[of "F"]) apply assumption apply (simp add:gHom) apply (simp add:Group.mult_closed[of "F"])+ apply (frule_tac x = a in gHom_mem[of "F" "G" "f"], assumption+, simp add:Group.r_unit[of "G"]) apply (simp add:ghom_inv_inv, simp add:Group.r_i) done lemma Group_coim:"\Group F; Group G; f \ gHom F G\ \ Group ( F / gker\<^bsub>F,G\<^esub> f)" by (frule gker_normal[of "F" "G" "f"], assumption+, simp add:Group.Group_Qg[of "F" "gker\<^bsub>F,G\<^esub> f"]) lemma gkern1:"\Group F; Ugp E; f \ gHom F E\ \ gker\<^bsub>F,E\<^esub> f = carrier F" apply (frule Group_Ugp[of "E"]) apply (frule gkernTr1_1[of "F" "E" "f"], assumption+) apply (rule equalityI, assumption) apply (rule subsetI, thin_tac "gker\<^bsub>F,E\<^esub> f \ carrier F", simp add:gkernel_def, frule_tac x = x in gHom_mem[of "F" "E" "f"], assumption+, simp add:Ugp_def) done lemma gkern2:"\Group F; Group G; f \ gHom F G; ginj\<^bsub>F,G\<^esub> f\ \ gker\<^bsub>F,G\<^esub> f = {\\<^bsub>F\<^esub>}" apply (frule gkernTr6[of "F" "G" "f"], assumption+, frule singleton_sub[of "\\<^bsub>F\<^esub>" "gker\<^bsub>F,G\<^esub> f"], rule equalityI) apply (rule subsetI, thin_tac "{\\<^bsub>F\<^esub>} \ gker\<^bsub>F,G\<^esub> f", simp add:gkernel_def, (erule conjE)+) apply (frule sym, thin_tac "f \\<^bsub>F\<^esub> = \\<^bsub>G\<^esub>", simp, thin_tac "\\<^bsub>G\<^esub> = f \\<^bsub>F\<^esub>", simp add:ginjec_def, simp add:inj_on_def) apply assumption done lemma gkernTr9:"\Group F; Group G; f \ gHom F G; a \ carrier F; b \ carrier F\ \ ((gker\<^bsub>F,G\<^esub> f) \\<^bsub>F\<^esub> a = (gker\<^bsub>F,G\<^esub> f) \\<^bsub>F\<^esub> b) = (f a = f b)" (* thm r_cosEqVar1 [of "F" "ker f (F \ G) " "a" "b"] *) apply (frule gkernTr7[of "F" "G" "f"], assumption+) apply (subst Group.rcs_eq[THEN sym, of "F" "gker\<^bsub>F,G\<^esub> f" "a" "b"], assumption+) apply (thin_tac "F \ gker\<^bsub>F,G\<^esub> f") apply (simp add:gkernel_def) apply (frule Group.i_closed[of "F" "a"], assumption) apply (simp add:Group.mult_closed[of "F"]) apply (simp add:gHom, simp add:ghom_inv_inv) apply (frule gHom_mem[of "F" "G" "f" "a"], assumption+, frule gHom_mem[of "F" "G" "f" "b"], assumption+) apply (rule iffI) apply (rule Group.r_div_eq[THEN sym, of "G" "f b" "f a"], assumption+) apply (simp add:Group.r_i[of "G"]) done lemma gkernTr11:"\Group F; Group G; f \ gHom F G ; a \ carrier F\ \ (iim F G f {f a}) = (gker\<^bsub>F,G\<^esub> f) \\<^bsub>F\<^esub> a" apply (frule gkernTr7[of "F" "G" "f"], assumption+) apply (rule equalityI) (** iim F G f {f a} \ ker\<^sub>F\<^sub>,\<^sub>Gf \<^sub>F a **) apply (rule subsetI, simp add:iim_def, erule conjE) apply (frule_tac a1 = x in gkernTr9[THEN sym, of "F" "G" "f" _ "a"], assumption+, simp, frule_tac a = x in Group.a_in_rcs[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+, simp) apply (rule subsetI) apply (simp add:gkernel_def rcs_def iim_def, erule exE, (erule conjE)+, rotate_tac -1, frule sym, thin_tac "h \\<^bsub>F\<^esub> a = x", simp add:gHom, thin_tac "x = h \\<^bsub>F\<^esub> a", frule gHom_mem[of "F" "G" "f" "a"], assumption+, simp add:Group.mult_closed Group.l_unit) done lemma gbij_comp_bij:"\Group F; Group G; Group H; gbij\<^bsub>F,G\<^esub> f; gbij\<^bsub>G,H\<^esub> g\ \ gbij\<^bsub>F,H\<^esub> (g \\<^bsub>F\<^esub> f)" apply (simp add:gbijec_def) apply (simp add:gHom_comp_gsurjec gHom_comp_ginjec) done lemma gbij_automorph:"\Group G; gbij\<^bsub>G,G\<^esub> f; gbij\<^bsub>G,G\<^esub> g\ \ gbij\<^bsub>G,G\<^esub> (g \\<^bsub>G\<^esub> f)" apply (simp add:gbij_comp_bij) done lemma l_unit_gHom:"\Group F; Group G; f \ gHom F G\ \ (I\<^bsub>G\<^esub>) \\<^bsub>F\<^esub> f = f" apply (simp add:cmpghom_def) apply (rule funcset_eq[of _ "carrier F"], simp add:compose_def, simp add:gHom_def) apply (rule ballI, simp add:idmap_def compose_def, simp add:gHom_mem[of "F" "G" "f"]) done lemma r_unit_gHom:"\Group F; Group G; f \ gHom F G \ \ f \\<^bsub>F\<^esub> (I\<^bsub>F\<^esub>) = f" apply (simp add:cmpghom_def) apply (rule funcset_eq[of _ "carrier F"], simp add:compose_def, simp add:gHom_def) apply (rule ballI, simp add:idmap_def compose_def) done section "Image" lemma inv_gHom:"\Group F; Group G; gbij\<^bsub>F,G\<^esub> f\ \ (Ifn F G f) \ gHom G F" apply (simp add:gHom_def) apply (rule conjI) (** Ifn F G f \ carrier G \ carrier F **) apply (simp add:invfun_def restrict_def extensional_def) apply (rule conjI) apply (rule inv_func) apply (simp add:gbijec_def gsurjec_def gHom_def) apply (simp add:gbijec_def gsurjec_def ginjec_def)+ apply (rule ballI)+ apply (erule conjE)+ apply (frule gHom_func[of "F" "G" "f"], assumption+, frule inv_func[of "f" "carrier F" "carrier G"], assumption+) apply (frule_tac b = x in invfun_mem[of "f" "carrier F" "carrier G"], assumption+, frule_tac b = y in invfun_mem[of "f" "carrier F" "carrier G"], assumption+) apply (frule_tac x = "(Ifn F G f) x" and y = "(Ifn F G f) y" in gHom[of "F" "G" "f"], assumption+) apply (simp only:invfun_r) apply (frule sym, thin_tac "f ((Ifn F G f) x \\<^bsub>F\<^esub> (Ifn F G f) y) = x \\<^bsub>G\<^esub> y") apply (frule_tac a = x and b = y in Group.mult_closed[of "G"], assumption+) apply (frule_tac b = "x \\<^bsub>G\<^esub> y" in invfun_r[of "f" "carrier F" "carrier G"], assumption+) apply (frule_tac r = "f ((Ifn F G f) (x \\<^bsub>G\<^esub> y))" and s = "x \\<^bsub>G\<^esub> y" and t = "f ((Ifn F G f) x \\<^bsub>F\<^esub> (Ifn F G f) y)" in trans, assumption+) apply (thin_tac "f ((Ifn F G f) (x \\<^bsub>G\<^esub> y)) = x \\<^bsub>G\<^esub> y", thin_tac "x \\<^bsub>G\<^esub> y = f ((Ifn F G f) x \\<^bsub>F\<^esub> (Ifn F G f) y)") apply (frule_tac b = "x \\<^bsub>G\<^esub> y" in invfun_mem[of "f" "carrier F" "carrier G"], assumption+, frule_tac a = "(Ifn F G f) x" and b = "(Ifn F G f) y" in Group.mult_closed[of "F"], assumption+) apply (simp add:inj_on_def) done lemma inv_gbijec_gbijec:"\Group F; Group G; gbij\<^bsub>F,G\<^esub> f\ \ gbij\<^bsub>G,F\<^esub> (Ifn F G f)" apply (frule inv_gHom [of "F" "G" "f"], assumption+) apply (simp add:gbijec_def) apply (rule conjI) (* gsurjec *) apply (simp add:gsurjec_def ginjec_def, (erule conjE)+) apply (frule gHom_func[of "F" "G" "f"], simp add:invfun_surj, assumption+, simp add:invfun_surj) (* ginjec *) apply (erule conjE, simp add:gsurjec_def ginjec_def, erule conjE, frule gHom_func[of "F" "G" "f"], assumption+, rule invfun_inj, assumption+) done lemma l_inv_gHom:"\Group F; Group G; gbij\<^bsub>F,G\<^esub> f\ \ (Ifn F G f) \\<^bsub>F\<^esub> f = (I\<^bsub>F\<^esub>)" apply (rule funcset_eq[of _ "carrier F"], simp add:cmpghom_def, simp add:idmap_def, rule ballI) apply (simp add:idmap_def cmpghom_def compose_def, simp add:gbijec_def gsurjec_def ginjec_def, (erule conjE)+, frule gHom_func[of "F" "G" "f"], assumption+) apply (rule invfun_l, assumption+) done lemma img_mult_closed:"\Group F; Group G; f \ gHom F G; u \ f `(carrier F); v \ f `(carrier F)\ \ u \\<^bsub>G\<^esub> v \ f `(carrier F)" apply (simp add:image_def) apply ((erule bexE)+, simp) apply (subst gHom [of "F" "G" "f", THEN sym], assumption+) apply (frule_tac a = x and b = xa in Group.mult_closed [of "F"], assumption+) apply blast done lemma img_unit_closed:"\Group F; Group G; f \ gHom F G\ \ \\<^bsub>G\<^esub> \ f `(carrier F)" apply (cut_tac Group.unit_closed[of "F"], frule ghom_unit_unit[THEN sym, of "F" "G" "f"], assumption+, simp add:image_def, blast, assumption) done lemma imgTr7:"\Group F; Group G; f \ gHom F G; u \ f `(carrier F)\ \ \\<^bsub>G\<^esub> u \ f `(carrier F)" apply (simp add:image_def, erule bexE, simp) apply (subst ghom_inv_inv[THEN sym, of "F" "G" "f"], assumption+) apply (frule_tac a = x in Group.i_closed[of "F"], assumption+) apply blast done lemma imgTr8:"\Group F; Group G; f \ gHom F G; F \ H; u \ f ` H; v \ f` H \ \ u \\<^bsub>G\<^esub> v \ f ` H" apply (simp add:image_def, (erule bexE)+, simp, subst gHom [of "F" "G" "f" _, THEN sym], assumption+) apply (simp only:Group.sg_subset_elem[of "F"], simp only:Group.sg_subset_elem[of "F"]) apply (frule_tac x = x and y = xa in Group.sg_mult_closed[of "F" "H"], assumption+) apply blast done lemma imgTr9:"\Group F; Group G; f \ gHom F G; F \ H; u \ f ` H\ \ \\<^bsub>G\<^esub> u \ f ` H" apply (simp add:image_def, erule bexE, simp) apply (frule_tac h = x in Group.sg_subset_elem[of "F" "H"], assumption+, simp add:ghom_inv_inv[THEN sym]) apply (frule_tac x = x in Group.sg_i_closed [of "F" "H"], assumption+, blast) done lemma imgTr10:"\Group F; Group G; f \ gHom F G; F \ H\ \ \\<^bsub>G\<^esub> \ f ` H" apply (frule Group.sg_unit_closed[of "F" "H"], assumption+, subst ghom_unit_unit[THEN sym, of "F" "G" "f"], assumption+) apply (simp add:image_def, blast) done lemma imgTr11:"\Group F; Group G; f \ gHom F G; F \ H\ \ G \ (f ` H)" apply (frule gHom_func[of "F" "G" "f"], assumption+, frule Group.sg_subset[of "F" "H"], assumption+, frule image_sub[of "f" "carrier F" "carrier G" "H"], assumption+) apply (rule Group.sg_condition[of "G" "f ` H"], assumption+, frule imgTr10[of "F" "G" "f" "H"], assumption+, rule nonempty, assumption) apply ((rule allI)+, rule impI, erule conjE, frule_tac u = b in imgTr9[of "F" "G" "f" "H"], assumption+, frule_tac u = a and v = "\\<^bsub>G\<^esub> b" in imgTr8[of "F" "G" "f"], assumption+) done lemma sg_gimg:"\Group F; Group G; f \ gHom F G \ \ G \ f`(carrier F)" apply (frule Group.special_sg_G [of "F"]) apply (simp add:imgTr11) done lemma Group_Img:"\Group F; Group G; f \ gHom F G \ \ Group (Img\<^bsub>F,G\<^esub> f)" apply (frule sg_gimg [of "F" "G" "f"], assumption+, simp add:Gimage_def Group.Group_Gp[of "G"]) done lemma Img_carrier:"\Group F; Group G; f \ gHom F G \ \ carrier (Img\<^bsub>F,G\<^esub> f) = f ` (carrier F)" by (simp add:Gimage_def Gp_def) lemma hom_to_Img:"\Group F; Group G; f \ gHom F G\ \ f \ gHom F (Img\<^bsub>F,G\<^esub> f)" by (simp add:gHom_def Gimage_def Gp_def) lemma gker_hom_to_img:"\Group F; Group G; f \ gHom F G \ \ gker\<^bsub>F,(Img\<^bsub>F,G\<^esub> f)\<^esub> f = gker\<^bsub>F,G\<^esub> f" by (simp add:gkernel_def Gimage_def, frule sg_gimg[of "F" "G" "f"], assumption+, simp add:Group.one_Gp_one[of "G" "f ` (carrier F)"]) lemma Pj_im_subg:"\Group G; G \ H; G \ K; K \ H\ \ Pj G K ` H = carrier ((Gp G H) / K)" apply (simp add: Qg_def [of "Gp G H" "K"]) apply (rule equalityI) apply (simp add: Pj_def set_rcs_def Group.sg_subset_elem cong: image_cong_simp) using Group.Gp_rcs Group.carrier_Gp Group.nsg_sg apply fastforce apply (rule subsetI) apply (simp add:image_def Pj_def) apply (simp add:set_rcs_def) apply (simp add:Group.Gp_carrier, erule bexE) apply (frule Group.nsg_sg[of "G" "K"], assumption+) apply (frule_tac x = a in Group.Gp_rcs[of "G" "K" "H"], assumption+, simp) apply (frule_tac h = a in Group.sg_subset_elem[of "G" "H"], assumption+) apply blast done lemma (in Group) subg_Qsubg:"\G \ H; G \ K; K \ H\ \ (G / K) \ carrier ((Gp G H) / K)" apply (frule Pj_ghom[of "K"]) apply (frule nsg_sg [of "K"]) apply (frule Group_Qg[of "K"]) apply (cut_tac imgTr11 [of "G" "G / K" "Pj G K" "H"]) apply (cut_tac Pj_im_subg [of "G" "H" "K"]) apply simp apply (rule Group_axioms | assumption)+ done section "Induced homomorphisms" lemma inducedhomTr:"\Group F; Group G; f \ gHom F G; S \ set_rcs F (gker\<^bsub>F,G\<^esub> f); s1 \ S; s2 \ S \ \ f s1 = f s2" apply (simp add:set_rcs_def, erule bexE) apply (frule_tac a1 = a in gkernTr11[THEN sym, of "F" "G" "f"], assumption+, simp, thin_tac "S = iim F G f {f a}", thin_tac "gker\<^bsub>F,G\<^esub> f \\<^bsub>F\<^esub> a = iim F G f {f a}") apply (simp add:iim_def) done definition induced_ghom :: "[('a, 'more) Group_scheme, ('b, 'more1) Group_scheme, ('a \ 'b)] \ ('a set \ 'b )" where "induced_ghom F G f = (\X\ (set_rcs F (gker\<^bsub>F,G\<^esub> f)). f (SOME x. x \ X))" abbreviation INDUCED_GHOM :: "['a \ 'b, ('a, 'm) Group_scheme, ('b, 'm1) Group_scheme] \ ('a set \ 'b )" ("(3_\\<^bsub>_,_\<^esub>)" [82,82,83]82) where "f\\<^bsub>F,G\<^esub> == induced_ghom F G f" lemma induced_ghom_someTr:"\Group F; Group G; f \ gHom F G; X \ set_rcs F (gker\<^bsub>F,G\<^esub> f)\ \ f (SOME xa. xa \ X) \ f `(carrier F)" apply (simp add:set_rcs_def, erule bexE, simp) apply (frule gkernTr7 [of "F" "G" "f"], assumption+) apply (rule someI2_ex) apply (frule_tac a = a in Group.a_in_rcs[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+) apply blast apply (frule_tac a = a and x = x in Group.rcs_subset_elem[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+) apply (simp add:image_def, blast) done lemma induced_ghom_someTr1:"\Group F; Group G; f \ gHom F G; a \ carrier F\ \ f (SOME xa. xa \ (gker\<^bsub>F,G\<^esub> f) \\<^bsub>F\<^esub> a) = f a" apply (rule someI2_ex) apply (frule gkernTr7 [of "F" "G" "f"], assumption+) apply (frule Group.a_in_rcs [of "F" "gker\<^bsub>F,G\<^esub> f" "a"], assumption+) apply blast apply (simp add:gkernTr11 [THEN sym]) apply (simp add:iim_def) done lemma inducedHOMTr0:"\Group F; Group G; f \ gHom F G; a \ carrier F\ \ (f\\<^bsub>F,G\<^esub>) ((gker\<^bsub>F,G\<^esub> f) \\<^bsub>F\<^esub> a) = f a" apply (simp add:induced_ghom_def) apply (frule gkernTr7[of "F" "G" "f"], assumption+) apply (simp add:Group.rcs_in_set_rcs[of "F"]) apply (simp add:induced_ghom_someTr1) done lemma inducedHOMTr0_1:"\Group F; Group G; f \ gHom F G\ \ (f\\<^bsub>F,G\<^esub>) \ set_rcs F (gker\<^bsub>F,G\<^esub> f) \ carrier G" apply (rule Pi_I) apply (simp add:set_rcs_def, erule bexE, simp) apply (subst inducedHOMTr0[of "F" "G" "f"], assumption+) apply (simp add:gHom_mem) done lemma inducedHOMTr0_2:"\Group F; Group G; f \ gHom F G\ \ (f\\<^bsub>F,G\<^esub>) \ set_rcs F (gker\<^bsub>F,G\<^esub> f) \ f ` (carrier F)" apply (rule Pi_I) apply (simp add:set_rcs_def, erule bexE, simp) apply (subst inducedHOMTr0[of "F" "G" "f"], assumption+) apply (simp add:image_def, blast) done lemma inducedHom:"\Group F; Group G; f \ gHom F G\ \ (f\\<^bsub>F,G\<^esub>) \ gHom (F/(gker\<^bsub>F,G\<^esub> f)) G" apply (simp add: gHom_def [of "F/ gker\<^bsub>F,G\<^esub> f" "G"]) apply (rule conjI) apply (simp add:induced_ghom_def extensional_def) apply (rule allI) apply (rule impI)+ apply (simp add:Qg_def) apply (rule conjI) apply (simp add:Qg_def inducedHOMTr0_1) apply (rule ballI)+ apply (simp add:Qg_def set_rcs_def, (erule bexE)+) apply simp apply (frule gker_normal[of "F" "G" "f"], assumption+) apply (simp add:Group.c_top_welldef [THEN sym, of "F" "gker\<^bsub>F,G\<^esub> f"]) apply (frule_tac a = a and b = aa in Group.mult_closed[of "F"], assumption+) apply (simp add:inducedHOMTr0) apply (simp add:gHom) done lemma induced_ghom_ginjec: "\Group F; Group G; f \ gHom F G \ \ ginj\<^bsub>(F/(gker\<^bsub>F,G\<^esub> f)),G\<^esub> (f\\<^bsub>F,G\<^esub>)" apply (simp add:ginjec_def) apply (simp add:inducedHom) apply (simp add:inj_on_def) apply (rule ballI)+ apply (simp add:Qg_def) apply (simp add:set_rcs_def) apply ((erule bexE)+, rule impI, simp) apply (simp add:inducedHOMTr0) apply (simp add: gkernTr11[THEN sym]) done lemma inducedhomgsurjec:"\Group F; Group G; gsurj\<^bsub>F,G\<^esub> f\ \ gsurj\<^bsub>(F/(gker\<^bsub>F,G\<^esub> f)),G\<^esub> (f\\<^bsub>F,G\<^esub>)" apply (simp add:gsurjec_def) apply (simp add:inducedHom) apply (erule conjE) apply (frule gHom_func[of "F" "G" "f"], assumption+) apply (rule surj_to_test) apply (simp add:Qg_def) apply (frule inducedHOMTr0_2[of "F" "G" "f"], assumption+) apply (simp add:surj_to_def[of "f" "carrier F" "carrier G"]) apply (rule ballI) apply (simp add:surj_to_def[of "f" "carrier F" "carrier G"], frule sym, thin_tac "f ` carrier F = carrier G", simp, thin_tac "carrier G = f ` carrier F") apply (simp add:image_def, erule bexE, simp, thin_tac "b = f x") apply (simp add:Qg_def) apply (frule_tac a = x in inducedHOMTr0[of "F" "G" "f"], assumption+) apply (frule gkernTr7[of "F" "G" "f"], assumption+) apply (frule_tac a = x in Group.rcs_in_set_rcs[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+) apply blast done lemma homomtr: "\Group F; Group G; f \ gHom F G\ \ (f\\<^bsub>F,G\<^esub>) \ gHom (F / (gker\<^bsub>F,G\<^esub> f)) (Img\<^bsub>F,G\<^esub> f)" apply (simp add:gHom_def [of "F / (gker\<^bsub>F,G\<^esub> f)" _]) apply (rule conjI) apply (simp add:induced_ghom_def extensional_def) apply (rule allI, (rule impI)+, simp add:Qg_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:Gimage_def Qg_def Gp_def, simp add:set_rcs_def, erule bexE) apply simp apply (simp add:inducedHOMTr0) apply (rule ballI)+ apply (simp add:Qg_def set_rcs_def, (erule bexE)+, simp) apply (frule gker_normal[of "F" "G" "f"], assumption+) apply (simp add:Group.c_top_welldef[THEN sym, of "F" "gker\<^bsub>F,G\<^esub> f"]) apply (frule_tac a = a and b = aa in Group.mult_closed[of "F"], assumption+) apply (simp add:inducedHOMTr0) apply (simp add:Gimage_def) apply (subst Group.Gp_mult_induced[of "G" "f ` carrier F"], assumption+) apply (simp add:sg_gimg) apply (simp add:image_def, blast) apply (simp add:image_def, blast) apply (simp add:gHom) done lemma homom2img: "\Group F; Group G; f \ gHom F G \ \ (f\\<^bsub>F,(Img\<^bsub>F,G\<^esub> f)\<^esub>) \ gHom (F / (gker\<^bsub>F,G\<^esub> f)) (Img\<^bsub>F,G\<^esub> f)" apply (frule hom_to_Img[of "F" "G" "f"], assumption+) apply (frule inducedHom[of "F" "Img\<^bsub>F,G\<^esub> f" "f"]) apply (simp add:Group_Img) apply assumption apply (simp add:gker_hom_to_img[of "F" "G" "f"]) done lemma homom2img1:"\Group F; Group G; f \ gHom F G; X \ set_rcs F (gker\<^bsub>F,G\<^esub> f)\ \ (f\\<^bsub>F,(Img\<^bsub>F,G\<^esub> f)\<^esub>) X = (f\\<^bsub>F,G\<^esub>) X" apply (frule gker_hom_to_img [of "F" "G" "f"], assumption+) apply (simp add:induced_ghom_def) done subsection "Homomorphism therems" definition iota :: "('a, 'm) Group_scheme \ ('a \ 'a)" (* should be used exclusively as an inclusion map *) ("(\\<^bsub>_\<^esub>)" [1000]999) where "\\<^bsub>F\<^esub> = (\x \ carrier F. x)" lemma iotahomTr0:"\Group G; G \ H; h \ H \ \ (\\<^bsub>(Gp G H)\<^esub>) h = h" apply (simp add:iota_def) apply (simp add:Gp_def) done lemma iotahom:"\Group G; G \ H; G \ N\ \ \\<^bsub>(Gp G H)\<^esub> \ gHom (Gp G H) (Gp G (H \\<^bsub>G\<^esub> N))" apply (simp add:gHom_def) apply (rule conjI) apply (simp add:iota_def extensional_def) apply (rule conjI) apply (simp add:Pi_def restrict_def iota_def) apply (rule allI, rule impI) apply (simp add:Gp_def) apply (frule Group.nsg_sg[of "G" "N"], assumption+) apply (frule Group.l_sub_smult[of "G" "H" "N"], assumption+, simp add: subsetD) apply (rule ballI)+ apply (simp add:iota_def, simp add:Group.Gp_carrier) apply (frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+, frule Group.l_sub_smult[of "G" "H" "N"], assumption+, simp add:Group.nsg_sg, frule_tac c = x in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+, frule_tac c = y in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+) apply (simp add:Group.Gp_mult_induced[of "G"]) apply (simp add:Group.sg_mult_closed) done lemma iotaTr0: "\Group G; G \ H; G \ N\ \ ginj\<^bsub>(Gp G H),(Gp G (H \\<^bsub>G\<^esub> N))\<^esub> (\\<^bsub>(Gp G H)\<^esub>)" apply (simp add:ginjec_def) apply (simp add:iotahom) apply (simp add:inj_on_def iota_def Gp_def) done theorem homomthm1:"\Group F; Group G; f \ gHom F G \ \ gbij\<^bsub>(F/ (gkernel F G f)), (Gimage F G f)\<^esub> (f\\<^bsub>F, (Gimage F G f)\<^esub>)" apply (frule homom2img [of "F" "G" "f"], assumption+) apply (simp add:gbijec_def) apply (frule hom_to_Img [of "F" "G" "f"], assumption+) apply (frule Group_coim[of "F" "G" "f"], assumption+, frule gHom_func[of "F / (gker\<^bsub>F,G\<^esub> f)" "Img\<^bsub>F,G\<^esub> f" "f\\<^bsub>F,Img\<^bsub>F,G\<^esub> f\<^esub>"], simp add:Group_Img, assumption) apply (rule conjI) (** gsurjec **) apply (simp add:gsurjec_def, rule surj_to_test[of "f\\<^bsub>F,Img\<^bsub>F,G\<^esub> f\<^esub>" "carrier (F / gker\<^bsub>F,G\<^esub> f)" "carrier (Img\<^bsub>F,G\<^esub> f)"], assumption+, rule ballI) apply (thin_tac "f\\<^bsub>F,Img\<^bsub>F,G\<^esub> f\<^esub> \ gHom (F / gker\<^bsub>F,G\<^esub> f) (Img\<^bsub>F,G\<^esub> f)") apply (simp add:Img_carrier, simp add:image_def, erule bexE, simp, simp add:Qg_def, frule gkernTr7[of "F" "G" "f"], assumption+) apply (frule_tac a = x in Group.rcs_in_set_rcs[of "F" "gker\<^bsub>F,G\<^esub> f"], assumption+) apply (simp add:homom2img1) apply (frule_tac a = x in inducedHOMTr0[of "F" "G" "f"], assumption+) apply blast (** gsurjec done **) (** ginjec **) apply (frule induced_ghom_ginjec[of "F" "G" "f"], assumption+) apply (simp add:ginjec_def) apply (frule conjunct2) apply (thin_tac "f\\<^bsub>F,G\<^esub> \ gHom (F / gker\<^bsub>F,G\<^esub> f) G \ inj_on (f\\<^bsub>F,G\<^esub>) (carrier (F / gker\<^bsub>F,G\<^esub> f))") apply (simp add:inj_on_def) apply ((rule ballI)+, rule impI) apply (simp add:Qg_def, fold Qg_def) apply (simp add:homom2img1) done lemma isomTr0 [simp]:"Group F \ F \ F" by (frule IdTr2 [of "F"], simp add:isomorphic_def, blast) lemma isomTr1:"\Group F; Group G; F \ G \ \ G \ F" apply (simp add:isomorphic_def, erule exE) apply (frule_tac f = f in inv_gbijec_gbijec[of "F" "G"], assumption+) apply blast done lemma isomTr2:"\Group F; Group G; Group H; F \ G; G \ H \ \ F \ H" apply (simp add:isomorphic_def) apply (erule exE)+ apply (simp add:gbijec_def) apply (erule conjE)+ apply (frule gHom_comp_gsurjec [of "F" "G" "H" _ _], assumption+) apply (frule gHom_comp_ginjec [of "F" "G" "H" _ _], assumption+) apply blast done lemma gisom1: "\Group F; Group G; f \ gHom F G \ \ (F/ (gker\<^bsub>F,G\<^esub> f)) \ (Img\<^bsub>F,G\<^esub> f)" apply (simp add:isomorphic_def) apply (frule homomthm1 [of "F" "G" "f"], assumption+) apply blast done lemma homomth2Tr0: "\Group F; Group G; f \ gHom F G; G \ N\ \ F \ (iim F G f N)" apply (frule Group.cond_nsg[of "F" "iim F G f N"], frule Group.nsg_sg[of "G" "N"], assumption+, simp add:ghomTr4[of "F" "G" "f" "N"]) apply ((rule ballI)+, simp add:iim_def, erule conjE, frule_tac a = a in Group.i_closed[of "F"], assumption+, frule_tac a = a and b = h in Group.mult_closed[of "F"], assumption+) apply (simp add:gHom ghom_inv_inv Group.mult_closed) apply (frule_tac x = a in gHom_mem[of "F" "G" "f"], assumption+, simp add:Group.nsgPr1_1, assumption) done lemma kern_comp_gHom:"\Group F; Group G; gsurj\<^bsub>F,G\<^esub> f; G \ N\ \ gker\<^bsub>F, (G/N)\<^esub> ((Pj G N) \\<^bsub>F\<^esub> f) = iim F G f N" apply (simp add:gkernel_def iim_def) apply (simp add:Group.Qg_one[of "G" "N"] cmpghom_def compose_def) apply (rule equalityI) apply (rule subsetI, simp, erule conjE, simp) apply (simp add:gsurjec_def, frule conjunct1, fold gsurjec_def) apply (frule_tac x = x in gHom_mem[of "F" "G" "f"], assumption+) apply (simp add:Group.Pj_mem[of "G" "N"]) apply (frule Group.nsg_sg[of "G" "N"], assumption+) apply (frule_tac a = "f x" in Group.a_in_rcs[of "G" "N"], assumption+) apply simp apply (rule subsetI) apply (simp, erule conjE) apply (frule Group.nsg_sg[of "G" "N"], assumption, frule_tac h = "f x" in Group.sg_subset_elem[of "G" "N"], assumption+) apply (simp add:Group.Pj_mem[of "G" "N"]) apply (simp add:Group.rcs_fixed2[of "G" "N"]) done lemma QgrpUnit_1:"\Group G; Ugp E; G \ H; (G / H) \ E \ \ carrier G = H" apply (simp add:isomorphic_def, erule exE) apply (frule Group.Group_Qg[of "G" "H"], assumption+, simp add:gbijec_def, erule conjE) apply (frule_tac f = f in gkern2[of "G / H" "E"], simp add:Ugp_def, simp add:gsurjec_def, assumption, simp add:gsurjec_def, frule conjunct1, fold gsurjec_def, frule_tac f = f in gkern1[of "G/H" "E"], assumption+) apply (simp, thin_tac "gker\<^bsub>(G / H),E\<^esub> f = {\\<^bsub>G / H\<^esub>}", thin_tac "gsurj\<^bsub>(G / H),E\<^esub> f", thin_tac "ginj\<^bsub>(G / H),E\<^esub> f", thin_tac "Group (G / H)", thin_tac "f \ gHom (G / H) E", simp add:Group.Qg_carrier) apply (rule contrapos_pp, simp+, frule Group.nsg_sg[of "G" "H"], assumption+, frule Group.sg_subset[of "G" "H"], assumption+, frule sets_not_eq[of "carrier G" "H"], assumption, erule bexE, frule_tac a = a in Group.rcs_in_set_rcs[of "G" "H"], assumption+, simp) apply (thin_tac "set_rcs G H = {\\<^bsub>G / H\<^esub>}", simp add:Qg_def, frule_tac a = a in Group.a_in_rcs[of "G" "H"], assumption+, simp) done lemma QgrpUnit_2:"\Group G; Ugp E; G \ H; carrier G = H\ \ (G/H) \ E" apply (frule Group.Group_Qg [of "G" "H"], assumption+) apply (simp add:Group.Qg_unit_group[THEN sym, of "G" "H"]) apply (simp add:Ugp_def) apply (frule Group.Qg_carrier[of "G" "H"], simp) apply (thin_tac "set_rcs G H = {H}") apply (frule Group.Qg_one[of "G" "H"], assumption+, erule conjE) apply (rule Ugps_isomorphic[of "G / H" "E"]) apply (simp add:Ugp_def)+ done lemma QgrpUnit_3:"\Group G; Ugp E; G \ H; G \ H1; (Gp G H) \ H1; ((Gp G H) / H1) \ E \ \ H = H1" apply (frule Group.Group_Gp[of "G" "H"], assumption+) apply (frule QgrpUnit_1 [of "Gp G H" "E" "H1"], assumption+) apply (simp add:Group.Gp_carrier) done lemma QgrpUnit_4:"\Group G; Ugp E; G \ H; G \ H1; (Gp G H) \ H1; \ ((Gp G H) / H1) \ E \ \ H \ H1" apply (frule Group.Group_Gp[of "G" "H"], assumption+) apply (rule contrapos_pp, simp) apply simp apply (frule QgrpUnit_2 [of "Gp G H1" "E" "H1"], assumption+) apply (simp add:Group.Gp_carrier) apply simp done definition Qmp :: "[('a, 'm) Group_scheme, 'a set, 'a set] \ ('a set \ 'a set)" where "Qmp G H N = (\X\ set_rcs G H. {z. \ x \ X. \ y \ N. (y \\<^bsub>G\<^esub> x = z)})" abbreviation QP :: "[_, 'a set, 'a set] \ ('a set \ 'a set)" ("(3Qm\<^bsub>_ _,_\<^esub>)" [82,82,83]82) where "Qm\<^bsub>G H,N\<^esub> == Qmp G H N" (* "\ Group G; G \ H; G \ N; H \ N \ \ Qmp\<^bsub>G H,N\<^esub> \ gHom (G / H) (G / N)" *) (* show Qmp G H N is a welldefined map from G/H to G/N. step1 *) lemma (in Group) QmpTr0:"\G \ H; G \ N; H \ N ; a \ carrier G\ \ Qmp G H N (H \ a) = (N \ a)" apply (frule_tac a = a in rcs_in_set_rcs[of "H"], assumption, simp add:Qmp_def) apply (rule equalityI) apply (rule subsetI, simp, (erule bexE)+, thin_tac "H \ a \ set_rcs G H", simp add:rcs_def, erule bexE) apply (frule sym, thin_tac "y \ xa = x", frule sym, thin_tac "h \ a = xa", simp, frule_tac h = y in sg_subset_elem[of "N"], assumption+, frule_tac h = h in sg_subset_elem[of "H"], assumption+, frule_tac c = h in subsetD[of "H" "N"], assumption+, frule_tac x = y and y = h in sg_mult_closed[of "N"], assumption+, subst tassoc[THEN sym], assumption+, blast) apply (rule subsetI, thin_tac "H \ a \ set_rcs G H", simp add:rcs_def, erule bexE, frule sg_unit_closed[of "H"], frule l_unit[of "a"], blast) done (* show Qmp G H N is a welldefined map from G/H to G/N. step2 *) lemma (in Group) QmpTr1:"\G \ H; G \ N; H \ N; a \ carrier G; b \ carrier G; H \ a = H \ b\ \ N \ a = N \ b" apply (simp add:rcs_eq[THEN sym, of "H" "a" "b"], frule subsetD[of "H" "N" "b \ \ a"], assumption+, simp add:rcs_eq[of "N" "a" "b"]) done lemma (in Group) QmpTr2:"\G \ H; G \ N; H \ N ; X \ carrier (G/H)\ \ (Qmp G H N) X \ carrier (G/N)" by (simp add:Qg_carrier[of "H"] set_rcs_def, erule bexE, simp add: QmpTr0, simp add:Qg_carrier rcs_in_set_rcs) lemma (in Group) QmpTr2_1:"\G \ H; G \ N; H \ N \ \ Qmp G H N \ carrier (G/H) \ carrier (G/N)" by (simp add:QmpTr2 [of "H" "N"]) lemma (in Group) QmpTr3:"\G \ H; G \ N; H \ N; X \ carrier (G/H); Y \ carrier (G/H)\ \ (Qmp G H N) (c_top G H X Y) = c_top G N ((Qmp G H N) X) ((Qmp G H N) Y)" apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"]) apply (simp add:Qg_carrier) apply (simp add:set_rcs_def, (erule bexE)+, simp) apply (subst c_top_welldef [THEN sym], assumption+) apply (frule_tac a = a and b = aa in mult_closed, assumption+) apply (simp add:QmpTr0)+ apply (subst c_top_welldef [THEN sym], assumption+) apply simp done lemma (in Group) Gp_s_mult_nsg:"\G \ H; G \ N; H \ N; a \ N \ \ H \\<^bsub>(Gp G N)\<^esub> a = H \ a" by (frule nsg_sg[of "H"], frule nsg_sg[of "N"], rule Gp_rcs[of "H" "N" "a"], assumption+) lemma (in Group) QmpTr5:"\G \ H; G \ N; H \ N; X \ carrier (G/H); Y \ carrier (G/H) \ \ (Qmp G H N) ( X \\<^bsub>(G / H)\<^esub> Y) = ((Qmp G H N) X) \\<^bsub>(G / N)\<^esub> ((Qmp G H N) Y)" by (frule nsg_sg[of "H"], frule nsg_sg[of "N"], (subst Qg_def)+, simp, simp add:QmpTr3 [of "H" "N" "X" "Y"]) lemma (in Group) QmpTr:"\G \ H; G \ N; H \ N \ \ (Qm\<^bsub>G H,N\<^esub>) \ gHom (G / H) (G / N)" apply (simp add:gHom_def) apply (rule conjI) apply (simp add:Qmp_def extensional_def) apply (rule allI, (rule impI)+, simp add:Qg_def) apply (rule conjI) apply (rule QmpTr2_1[of "H" "N"]) apply (simp add:nsg_sg)+ apply (rule ballI)+ apply (simp add:QmpTr5) done lemma (in Group) Qmpgsurjec:"\G \ H; G \ N; H \ N \ \ gsurj\<^bsub>(G / H),(G / N)\<^esub> (Qm\<^bsub>G H,N\<^esub>)" apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"]) apply (frule QmpTr [of "H" "N"], assumption+) apply (simp add:gsurjec_def) apply (rule surj_to_test) apply (simp add:QmpTr2_1) apply (rule ballI) apply (simp add:Qg_carrier, simp add:set_rcs_def[of "G" "N"], erule bexE, frule_tac a = a in QmpTr0[of "H" "N"], assumption+, simp) apply (frule_tac a = a in rcs_in_set_rcs[of "H"], assumption+, blast) done lemma (in Group) gkerQmp:"\G \ H; G \ N; H \ N \ \ gker\<^bsub>(G / H),(G / N)\<^esub> (Qm\<^bsub>G H,N\<^esub>) = carrier ((Gp G N)/ H)" apply (frule nsg_sg[of "H"], frule nsg_sg[of "N"]) apply (simp add:gkernel_def) apply (rule equalityI) apply (rule subsetI, simp add:Qg_carrier set_rcs_def, erule conjE, erule bexE, simp, simp add:Qg_one) apply (simp add:QmpTr0, frule_tac a = a in a_in_rcs[of "N"], assumption+, simp, frule Group_Gp[of "N"]) apply (simp add:Group.Qg_carrier, simp add:set_rcs_def, simp add:Gp_carrier, simp add:Gp_rcs, blast) apply (rule subsetI) apply (frule Group_Gp[of "N"], simp add:Group.Qg_carrier Qg_one set_rcs_def, erule bexE, simp add:Qg_carrier, thin_tac "x = H \\<^bsub>\N\<^esub> a") apply (simp add:Gp_carrier, simp add:Gp_rcs, frule_tac h = a in sg_subset_elem[of "N"], assumption, simp add:rcs_in_set_rcs, simp add:QmpTr0, simp add:rcs_fixed2[of "N"]) done theorem (in Group) homom2:"\G \ H; G \ N; H \ N\ \ gbij\<^bsub>((G/H)/(carrier ((Gp G N)/H))),(G/N)\<^esub> ((Qm\<^bsub>G H,N\<^esub>)\\<^bsub>(G/H),(G/N)\<^esub>)" apply (frule QmpTr [of "H" "N"], assumption+) apply (simp add:gbijec_def) apply (rule conjI) apply (frule Group_Qg[of "H"], frule Group_Qg[of "N"]) apply (frule inducedHom [of "G/H" "G/N" " Qmp G H N"], assumption+) (** show surj_to **) apply (frule Qmpgsurjec [of "H" "N"], assumption+) apply (frule inducedhomgsurjec [of "G/H" "G/N" "Qm\<^bsub>G H,N\<^esub>"], assumption+) apply (simp add:gkerQmp [of "H" "N"]) (** show ginj **) apply (frule QmpTr [of "H" "N"], assumption+) apply (frule Group_Qg[of "H"], frule Group_Qg[of "N"]) apply (frule induced_ghom_ginjec [of "G/H" "G/N" "Qmp G H N"], assumption+) apply (simp add:gkerQmp [of "H" "N"]) done section "Isomorphims" theorem (in Group) isom2:"\G \ H; G \ N; H \ N\ \ ((G/H)/(carrier ((Gp G N)/H))) \ (G/N)" apply (frule homom2 [of "H" "N"], assumption+) apply (simp add:isomorphic_def) apply blast done theorem homom3:"\ Group F; Group G; G \ N; gsurj\<^bsub>F,G\<^esub> f; N1 = (iim F G f) N \ \ (F / N1) \ (G / N)" apply (frule Group.Pj_gsurjec [of "G" "N"], assumption+) apply (frule Group.Group_Qg[of "G" "N"], assumption) apply (frule gHom_comp_gsurjec [of "F" "G" "G / N" "f" "Pj G N"], assumption+) apply (frule inducedhomgsurjec [of "F" "G / N" "(Pj G N) \\<^bsub>F\<^esub> f"], assumption+) apply (frule induced_ghom_ginjec [of "F" "G / N" "(Pj G N) \\<^bsub>F\<^esub> f"], assumption+) apply (simp add:gsurjec_def [of "F" "G / N" "(Pj G N) \\<^bsub>F\<^esub> f"]) apply (simp add:kern_comp_gHom[of "F" "G" "f" "N"]) apply (frule sym, thin_tac "N1 = iim F G f N", simp) apply (simp add:isomorphic_def gbijec_def, blast) done lemma (in Group) homom3Tr1:"\G \ H; G \ N\ \ H \ N = gker\<^bsub>(Gp G H),((Gp G (H \\<^bsub>G\<^esub> N))/N)\<^esub> ((Pj (Gp G (H \\<^bsub>G\<^esub> N)) N) \\<^bsub>(Gp G H)\<^esub> (\\<^bsub>(Gp G H)\<^esub>))" apply (simp add:gkernel_def, frule nsg_sg, simp add:Gp_carrier[of "H"], frule smult_sg_nsg, assumption+, frule Gp_smult_nsg[of "H" "N"], assumption, frule Group_Gp [of "H \\<^bsub>G\<^esub> N"]) apply (simp add:Group.Qg_one[of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], simp add:iota_def Gp_carrier, simp add:cmpghom_def compose_def, simp add:Gp_carrier) apply (rule equalityI) apply (rule subsetI, simp, erule conjE) apply (frule_tac h = x in sg_subset_elem[of "H"], assumption+, subst Group.Pj_mem, assumption+, simp add:Gp_carrier, frule l_sub_smult[of "H" "N"], assumption+, rule_tac c = x in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+) apply (frule r_sub_smult[of "H" "N"], assumption+, frule_tac c = x in subsetD[of "N" "H \\<^bsub>G\<^esub> N"], assumption+, simp add:Gp_rcs[of "N" "H \\<^bsub>G\<^esub> N"]) apply (simp add:rcs_fixed2) apply (rule subsetI, simp, erule conjE, simp) apply (frule_tac h = x in sg_subset_elem[of "H"], assumption+) apply (frule l_sub_smult[of "H" "N"], assumption+, frule r_sub_smult[of "H" "N"], assumption+) apply (frule_tac x = x in Group.Pj_mem[of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], assumption+) apply (simp add:Gp_carrier) apply ( rule_tac c = x in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+) apply (frule_tac c = x in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+) apply (simp only:Group.Gp_rcs) apply (simp only:Gp_rcs[of "N" "H \\<^bsub>G\<^esub> N"]) apply (frule_tac a = x in a_in_rcs[of "N"], assumption+, simp) done subsection "An automorphism groups" definition automg :: "_ \ \ carrier :: ('a \ 'a) set, top :: ['a \ 'a,'a \ 'a] \ ('a \ 'a), iop :: ('a \ 'a) \ ('a \ 'a), one :: ('a \ 'a)\" where "automg G = \ carrier = {f. gbij\<^bsub>G,G\<^esub> f}, top = \g\{f. gbij\<^bsub>G,G\<^esub> f}. \f\{f. gbij\<^bsub>G,G\<^esub> f}. ( g \\<^bsub>G\<^esub> f), iop = \f\{f. gbij\<^bsub>G,G\<^esub> f}. (Ifn G G f), one = I\<^bsub>G\<^esub> \" lemma automgroupTr1:"\Group G; gbij\<^bsub>G,G\<^esub> f; gbij\<^bsub>G,G\<^esub> g; gbij\<^bsub>G,G\<^esub> h\ \ (h \\<^bsub>G\<^esub> g) \\<^bsub>G\<^esub> f = h \\<^bsub>G\<^esub> (g \\<^bsub>G\<^esub> f)" apply (simp add:cmpghom_def, unfold gbijec_def) apply (frule conjunct1, rotate_tac 2, frule conjunct1, rotate_tac 1, frule conjunct1, fold gbijec_def) apply (simp add:gsurjec_def, (erule conjE)+, frule gHom_func[of "G" "G" "f"], assumption+, frule gHom_func[of "G" "G" "g"], assumption+, frule gHom_func[of "G" "G" "h"], assumption+) apply (simp add:compose_assoc) done lemma automgroup:"Group G \ Group (automg G)" apply (unfold Group_def [of "automg G"]) apply(auto simp: automg_def Pi_def gbij_comp_bij automgroupTr1 IdTr2 Id_l_unit l_inv_gHom inv_gbijec_gbijec) done subsection "Complete system of representatives" definition gcsrp :: "_ \ 'a set \ 'a set \ bool" where "gcsrp G H S == \f. (bij_to f (set_rcs G H) S)" (** G_csrp is defined iff H is a nsg **) definition gcsrp_map::"_ \ 'a set \ 'a set \ 'a" where "gcsrp_map G H == \X\(set_rcs G H). SOME x. x \ X" lemma (in Group) gcsrp_func:"G \ H \ gcsrp_map G H \ set_rcs G H \ UNIV" by (simp add:set_rcs_def) lemma (in Group) gcsrp_func1:"G \ H \ gcsrp_map G H \ set_rcs G H \ (gcsrp_map G H) ` (set_rcs G H)" by (simp add:set_rcs_def) lemma (in Group) gcsrp_map_bij:"G \ H \ bij_to (gcsrp_map G H) (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H))" apply (simp add:bij_to_def, rule conjI) apply (rule surj_to_test) apply (rule Pi_I) apply (simp add:image_def, blast) apply (rule ballI, simp add:image_def, erule bexE, simp, blast) apply (simp add:inj_on_def) apply ((rule ballI)+, rule impI) apply (simp add:gcsrp_map_def) apply (frule_tac X = x in rcs_nonempty, assumption+, frule_tac X = y in rcs_nonempty, assumption+) apply (frule_tac A = x in nonempty_some, frule_tac A = y in nonempty_some, simp) apply (rule_tac X = x and Y = y in rcs_meet[of "H"], assumption+) apply blast done lemma (in Group) image_gcsrp:"G \ H \ gcsrp G H ((gcsrp_map G H) `(set_rcs G H))" apply (simp add:gcsrp_def) apply (frule gcsrp_map_bij[of "H"], blast) done lemma (in Group) gcsrp_exists:"G \ H \ \S. gcsrp G H S" by (frule image_gcsrp[of "H"], blast) definition gcsrp_top :: "[_ , 'a set] \ 'a \ 'a \ 'a" where "gcsrp_top G H == \x \ ((gcsrp_map G H) `(set_rcs G H)). \y \ ((gcsrp_map G H) `(set_rcs G H)). gcsrp_map G H (c_top G H ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) x) ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) y))" definition gcsrp_iop::"[_ , 'a set] \ 'a \ 'a" where "gcsrp_iop G H = (\x \ ((gcsrp_map G H) `(set_rcs G H)). gcsrp_map G H (c_iop G H ((invfun (set_rcs G H) ((gcsrp_map G H) `(set_rcs G H)) (gcsrp_map G H)) x)))" definition gcsrp_one::"[_ , 'a set] \ 'a" where "gcsrp_one G H = gcsrp_map G H H" definition Gcsrp :: "_ \ 'a set \ 'a Group" where "Gcsrp G N = \carrier = (gcsrp_map G N) `(set_rcs G N), top = gcsrp_top G N, iop = gcsrp_iop G N, one = gcsrp_one G N\" lemma (in Group) gcsrp_top_closed:"\Group G; G \ N; a \ ((gcsrp_map G N) `(set_rcs G N)); b \ ((gcsrp_map G N) `(set_rcs G N))\ \ gcsrp_top G N a b \ (gcsrp_map G N) `(set_rcs G N)" apply (frule nsg_sg[of "N"], frule gcsrp_func1[of "N"], frule gcsrp_map_bij[of "N"]) apply (frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+, frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "b"], assumption+) apply (frule Qg_top_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b"], assumption+) apply (simp add:gcsrp_top_def) done lemma (in Group) gcsrp_tassoc:"\Group G; G \ N; a \ ((gcsrp_map G N) `(set_rcs G N)); b \ ((gcsrp_map G N) `(set_rcs G N)); c \ ((gcsrp_map G N) `(set_rcs G N))\ \ (gcsrp_top G N (gcsrp_top G N a b) c) = (gcsrp_top G N a (gcsrp_top G N b c))" apply (frule nsg_sg[of "N"], frule gcsrp_func1[of "N"], frule gcsrp_map_bij[of "N"]) apply (frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+, frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "b"], assumption+, frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "c"], assumption+) apply (frule Qg_top_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b"], assumption+, frule Qg_top_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) b" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) c"], assumption+) apply (simp add:gcsrp_top_def) apply (simp add:invfun_l1) apply (simp add:Qg_tassoc[of "N"]) done lemma (in Group) gcsrp_l_one:"\Group G; G \ N; a \ ((gcsrp_map G N) `(set_rcs G N))\ \ (gcsrp_top G N (gcsrp_one G N) a) = a" apply (frule nsg_sg[of "N"], frule gcsrp_func1[of "N"], frule gcsrp_map_bij[of "N"], frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+) apply (simp add:gcsrp_top_def gcsrp_one_def) apply (frule Qg_unit_closed[of "N"]) apply (simp add:Pi_def invfun_l1 Qg_unit invfun_r1) done lemma (in Group) gcsrp_l_i:"\G \ N; a \ ((gcsrp_map G N) `(set_rcs G N))\ \ gcsrp_top G N (gcsrp_iop G N a) a = gcsrp_one G N" apply (frule nsg_sg[of "N"], frule gcsrp_func1[of "N"], frule gcsrp_map_bij[of "N"], frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+) apply (frule Qg_iop_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a"], assumption+) apply (simp add:gcsrp_top_def gcsrp_one_def gcsrp_iop_def) apply (simp add:invfun_l1 Qg_i) done lemma (in Group) gcsrp_i_closed:"\G \ N; a \ ((gcsrp_map G N) `(set_rcs G N))\ \ gcsrp_iop G N a \ ((gcsrp_map G N) `(set_rcs G N))" apply (frule nsg_sg[of "N"], frule gcsrp_func1[of "N"], frule gcsrp_map_bij[of "N"], frule invfun_mem1[of "gcsrp_map G N" "set_rcs G N" "(gcsrp_map G N) ` (set_rcs G N)" "a"], assumption+) apply (frule Qg_iop_closed[of "N" "invfun (set_rcs G N) (gcsrp_map G N ` set_rcs G N) (gcsrp_map G N) a"], assumption+) apply (simp add:gcsrp_iop_def) done lemma (in Group) Group_Gcsrp:"G \ N \ Group (Gcsrp G N)" apply (simp add:Group_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:Gcsrp_def) apply (rule Pi_I) apply (rule_tac a = x and b = xa in gcsrp_top_closed[of "N"], rule Group_axioms, assumption+) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Gcsrp_def) apply (rule_tac a = a and b = b and c = c in gcsrp_tassoc[of "N"], rule Group_axioms, assumption+) apply (rule conjI) apply (rule Pi_I) apply (simp add:Gcsrp_def, rule gcsrp_i_closed[of "N"], assumption+) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:Gcsrp_def, rule gcsrp_l_i[of "N"], assumption+) apply (rule conjI) apply (frule Qg_unit_closed[of "N"], simp add:Gcsrp_def gcsrp_one_def) apply (rule allI, rule impI) apply (simp add:Gcsrp_def) apply (rule gcsrp_l_one[of "N"], rule Group_axioms, assumption+) done lemma (in Group) gcsrp_map_gbijec:"G \ N \ gbij\<^bsub>(G/N), (Gcsrp G N)\<^esub> (gcsrp_map G N)" apply (simp add:gbijec_def gsurjec_def ginjec_def Qg_carrier Gcsrp_def) apply (frule nsg_sg[of "N"], frule gcsrp_map_bij[of "N"], simp add:bij_to_def) apply (fold Gcsrp_def) apply (simp add:gHom_def) apply (rule conjI) apply (simp add:Qg_carrier gcsrp_map_def) apply (rule conjI) apply (simp add:Qg_carrier Gcsrp_def) apply (fold bij_to_def) apply (rule ballI)+ apply (simp add:Qg_def Gcsrp_def gcsrp_top_def) apply (frule gcsrp_func1[of "N"]) apply (simp add:invfun_l1[of "gcsrp_map G N" "set_rcs G N" "gcsrp_map G N ` set_rcs G N"]) done lemma (in Group) Qg_equiv_Gcsrp:"G \ N \ (G / N) \ Gcsrp G N" apply (simp add:isomorphic_def) apply (frule gcsrp_map_gbijec[of "N"], blast) done section "Zassenhaus" text\we show \H \ H N/N\ is gsurjective\ lemma (in Group) homom4Tr1:"\G \ N; G \ H\ \ Group ((Gp G (H \\<^bsub>G\<^esub> N)) / N)" apply (frule Gp_smult_sg_nsg[of "H" "N"], assumption+) apply (frule Gp_smult_nsg [of "H" "N"], assumption+) apply (simp add:Group.Group_Qg) done lemma homom3Tr2:"\Group G; G \ H; G \ N\ \ gsurj\<^bsub>(Gp G H),((Gp G (H \\<^bsub>G\<^esub> N))/N)\<^esub> ((Pj (Gp G (H \\<^bsub>G\<^esub> N)) N) \\<^bsub>(Gp G H)\<^esub> (\\<^bsub>(Gp G H)\<^esub>))" apply (frule iotahom[of "G" "H" "N"], assumption+, frule Group.Gp_smult_nsg[of "G" "H" "N"], assumption+, frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+, frule Group.Gp_smult_sg_nsg[of "G" "H" "N"], assumption+, frule Group.Pj_gsurjec [of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], assumption, frule Group.Group_Gp[of "G" "H"], assumption+, frule Group.Group_Qg[of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], assumption+, frule gHomcomp[of "Gp G H" "Gp G (H \\<^bsub>G\<^esub> N)" "(Gp G (H \\<^bsub>G\<^esub> N)) / N" "\\<^bsub>(\\<^bsub>G\<^esub>H)\<^esub>" "Pj (Gp G (H \\<^bsub>G\<^esub> N)) N"], assumption+) apply (simp add:gsurjec_def) apply (subst gsurjec_def, simp) apply (rule surj_to_test, simp add:gHom_def) apply (rule ballI) apply (simp add:Group.Qg_carrier[of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], simp add:set_rcs_def, erule bexE, frule Group.nsg_sg[of "G" "N"], assumption, frule Group.r_sub_smult[of "G" "H" "N"], assumption+, simp add:Group.Gp_carrier) apply (simp add:Group.Gp_rcs[of "G" "N" "H \\<^bsub>G\<^esub> N"]) apply (thin_tac "\\<^bsub>(\\<^bsub>G\<^esub>H)\<^esub> \ gHom (\\<^bsub>G\<^esub>H) (\\<^bsub>G\<^esub>(H \\<^bsub>G\<^esub> N))", thin_tac "gsurj\<^bsub>(Gp G (H \\<^bsub>G\<^esub> N)),((\\<^bsub>G\<^esub>(H \\<^bsub>G\<^esub> N)) / N)\<^esub> Pj (\\<^bsub>G\<^esub>(H \\<^bsub>G\<^esub> N)) N", thin_tac "Pj (\\<^bsub>G\<^esub>(H \\<^bsub>G\<^esub> N)) N \\<^bsub>(\\<^bsub>G\<^esub>H)\<^esub> \\<^bsub>(\\<^bsub>G\<^esub>H)\<^esub> \ gHom (\\<^bsub>G\<^esub>H) ((\\<^bsub>G\<^esub>(H \\<^bsub>G\<^esub> N)) / N)") apply (simp add:cmpghom_def compose_def, simp add:Group.Gp_carrier iota_def, frule Group.smult_commute_sg_nsg[of "G" "H" "N"], assumption+, frule_tac a = a in eq_set_inc[of _ "H \\<^bsub>G\<^esub> N" "N \\<^bsub>G\<^esub> H"], assumption+, thin_tac "H \\<^bsub>G\<^esub> N = N \\<^bsub>G\<^esub> H") apply (simp add:s_top_def[of "G" "N" "H"], (erule bexE)+, rotate_tac -1, frule sym, thin_tac "x \\<^bsub>G\<^esub> y = a", frule_tac h = y in Group.sg_subset_elem[of "G" "H"], assumption+, simp add:Group.rcs_fixed1[THEN sym]) apply (frule Group.l_sub_smult[of "G" "H" "N"], assumption+, frule_tac x1 = y in Group.Pj_mem[THEN sym, of "Gp G (H \\<^bsub>G\<^esub> N)" "N"], assumption+, simp add:Group.Gp_carrier, simp add: subsetD) apply (frule_tac c = y in subsetD[of "H" "H \\<^bsub>G\<^esub> N"], assumption+, simp add:Group.Gp_rcs[of "G" "N" "H \\<^bsub>G\<^esub> N"], blast) done theorem homom4:"\Group G; G \ N; G \ H\ \gbij\<^bsub>((Gp G H)/(H \ N)),((Gp G (H \\<^bsub>G\<^esub> N)) / N)\<^esub> (((Pj (Gp G (H \\<^bsub>G\<^esub> N)) N) \\<^bsub>(Gp G H)\<^esub> (\\<^bsub>(Gp G H)\<^esub>))\\<^bsub>(Gp G H),((Gp G (H \\<^bsub>G\<^esub> N)) / N)\<^esub>)" apply (frule homom3Tr2 [of "G" "H" "N"], assumption+) apply (frule Group.Gp_smult_sg_nsg, assumption+) apply (frule Group.homom4Tr1[of "G" "N" "H"], assumption+) apply (frule Group.Group_Gp [of "G" "H"], assumption+) apply (frule induced_ghom_ginjec [of "Gp G H" "(Gp G (H \\<^bsub>G\<^esub> N)/N)" "(Pj (Gp G (H \\<^bsub>G\<^esub> N)) N) \\<^bsub>(Gp G H)\<^esub> (\\<^bsub>(Gp G H)\<^esub>)"], assumption+) apply (simp add:gsurjec_def) apply (frule inducedhomgsurjec [of "Gp G H" "(Gp G (H \\<^bsub>G\<^esub> N))/N" "(Pj (Gp G (H \\<^bsub>G\<^esub> N)) N) \\<^bsub>(Gp G H)\<^esub> (\\<^bsub>(Gp G H)\<^esub>)"], assumption+) apply (frule Group.homom3Tr1[of "G" "H" "N"], assumption+) apply simp apply (simp add:gbijec_def) done lemma (in Group) homom4_2:"\G \ N; G \ H\ \ Group ((Gp G H) / (H \ N))" by (frule Group_Gp[of "H"], frule inter_Gp_nsg[of "N" "H"], assumption, rule Group.Group_Qg, assumption+) lemma isom4:"\Group G; G \ N; G \ H\ \ ((Gp G H)/(H \ N)) \ ((Gp G (N \\<^bsub>G\<^esub> H)) / N)" apply (frule homom4 [of "G" "N" "H"], assumption+, frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+, frule Group.smult_commute_sg_nsg[of "G" "H" "N"], assumption+) apply (simp add:isomorphic_def, blast) done lemma ZassenhausTr5:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ ((Gp G (H \ K))/((H1 \ K) \\<^bsub>G\<^esub> (H \ K1))) \ ((Gp G (H1 \\<^bsub>G\<^esub> (H \ K)))/(H1 \\<^bsub>G\<^esub> (H \ K1)))" apply (frule Group.ZassenhausTr2_1 [of "G" "H" "H1" "K"], assumption+, frule Group.Group_Gp [of "G" "H1 \\<^bsub>G\<^esub> (H \ K)"], assumption+, frule Group.ZassenhausTr3 [of "G" "H" "H1" "K" "K1"], assumption+, frule Group.inter_sgs [of "G" "H" "K"], assumption+, frule Group.r_sub_smult[of "G" "H1" "H \ K"], assumption+, frule Group.sg_sg[of "G" "H1 \\<^bsub>G\<^esub> H \ K" "H \ K"], assumption+, frule isom4 [of "Gp G (H1 \\<^bsub>G\<^esub> H \ K)" "H1 \\<^bsub>G\<^esub> H \ K1" "H \ K"], assumption+) apply (simp add:Int_commute[of "H \ K" "H1 \\<^bsub>G\<^esub> H \ K1"]) apply (frule Group.Group_Gp[of "G" "H"], assumption, frule Group.Group_Gp[of "G" "K"], assumption, frule Group.nsg_sg[of "Gp G H" "H1"], assumption+, frule Group.sg_subset[of "Gp G H" "H1"], assumption+, frule Group.nsg_sg[of "Gp G K" "K1"], assumption+, frule Group.sg_subset[of "Gp G K" "K1"], assumption+, simp add:Group.Gp_carrier, frule Group.inter_sgs[of "G" "H" "K1"], assumption+, cut_tac subset_self[of "H"], frule Int_mono[of "H" "H" "K1" "K"], assumption) apply (simp add:Group.s_topTr6_1[of "G" "H1" "H \ K1" "H \ K"], simp add:Int_assoc[THEN sym, of "H1" "H" "K"]) apply (simp add:Int_absorb2[of "H1" "H"], simp add:Group.Gp_inherited[of "G" "H \ K" "H1 \\<^bsub>G\<^esub> H \ K"]) apply (frule Group.s_top_mono[of "G" "H1" "H \ K" "H1" "H \ K1"], frule Group.sg_subset[of "G" "H"], assumption+, rule subset_trans[of "H1" "H" "carrier G"], assumption+) apply (rule Group.sg_subset[of "G" "H \ K"], assumption+, simp, simp, (frule Group.ZassenhausTr2_1[of "G" "H" "H1" "K"], assumption+, frule Group.subg_sg_sg[of "G" "H1 \\<^bsub>G\<^esub> H \ K" "H1 \\<^bsub>G\<^esub> H \ K1"], assumption+, simp add:Group.nsg_sg)) apply (simp add:Group.s_top_induced[of "G" "H1 \\<^bsub>G\<^esub> H \ K" "H1 \\<^bsub>G\<^esub> H \ K1" "H \ K"], simp add:Group.s_top_assoc[of "G" "H1" "H \ K1" "H \ K"], cut_tac subset_self[of "H"], frule Int_mono[of "H" "H" "K1" "K"], assumption) apply (simp add:Group.K_absorb_HK[of "G" "H \ K1" "H \ K"]) apply (cut_tac subset_self[of "H1 \\<^bsub>G\<^esub> H \ K"], simp add:Group.Gp_inherited[of "G" "H1 \\<^bsub>G\<^esub> H \ K" "H1 \\<^bsub>G\<^esub> H \ K"]) done lemma ZassenhausTr5_1:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ ((Gp G (K \ H))/((K1 \ H) \\<^bsub>G\<^esub> (K \ H1))) \ ((Gp G (K1 \\<^bsub>G\<^esub> (K \ H)))/(K1 \\<^bsub>G\<^esub> (K \ H1)))" (* thm ZassenhausTr5 [of "G" "K" "K1" "H" "H1"] *) apply (simp add:ZassenhausTr5 [of "G" "K" "K1" "H" "H1"]) done lemma ZassenhausTr5_2: "\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ ((Gp G (H \ K))/((H1 \ K) \\<^bsub>G\<^esub> (H \ K1))) = ((Gp G (K \ H))/((K1 \ H) \\<^bsub>G\<^esub> (K \ H1)))" by (simp add:Group.ZassenhausTr3_3[of "G" "H" "H1" "K" "K1"], simp add:Int_commute[of "H" "K"]) lemma ZassenhausTr6_1:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ Group (Gp G (H \ K) / (H1 \ K \\<^bsub>G\<^esub> H \ K1))" apply (frule Group.inter_sgs [of "G" "H" "K"], assumption+, frule Group.Group_Gp [of "G" "H \ K"], assumption+, frule Group.ZassenhausTr3_5 [of "G" "H" "H1" "K" "K1"], assumption+) apply (rule Group.Group_Qg, assumption+) done lemma ZassenhausTr6_2:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ Group (Gp G (H1 \\<^bsub>G\<^esub> H \ K) / (H1 \\<^bsub>G\<^esub> H \ K1))" apply (frule Group.ZassenhausTr2_1 [of "G" "H" "H1" "K"], assumption+, frule Group.Group_Gp [of "G" "H1 \\<^bsub>G\<^esub> H \ K"], assumption+, frule Group.ZassenhausTr3 [of "G" "H" "H1" "K" "K1"], assumption+) apply (simp add:Group.Group_Qg) done lemma ZassenhausTr6_3:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ Group (Gp G (K1 \\<^bsub>G\<^esub> K \ H) / (K1 \\<^bsub>G\<^esub> K \ H1))" apply (frule Group.ZassenhausTr2_1 [of "G" "K" "K1" "H"], assumption+, frule Group.Group_Gp [of "G" "K1 \\<^bsub>G\<^esub> K \ H"], assumption+, frule Group.ZassenhausTr3[of "G" "K" "K1" "H" "H1"], assumption+) apply (simp add:Group.Group_Qg) done theorem Zassenhaus:"\Group G; G \ H; G \ H1; G \ K; G \ K1; Gp G H \ H1; Gp G K \ K1\ \ (Gp G (H1 \\<^bsub>G\<^esub> H \ K) / (H1 \\<^bsub>G\<^esub> H \ K1)) \ (Gp G (K1 \\<^bsub>G\<^esub> K \ H) / (K1 \\<^bsub>G\<^esub> K \ H1))" apply (frule ZassenhausTr6_1[of "G" "K" "K1" "H" "H1"], assumption+) apply (frule ZassenhausTr6_3 [of "G" "H" "H1" "K" "K1"], assumption+) apply (frule ZassenhausTr6_2 [of "G" "H" "H1" "K" "K1"], assumption+) apply (rule isomTr2[of "(\\<^bsub>G\<^esub>(H1 \\<^bsub>G\<^esub> H \ K)) / (H1 \\<^bsub>G\<^esub> H \ K1)" "(\\<^bsub>G\<^esub>(K \ H)) / (K1 \ H \\<^bsub>G\<^esub> K \ H1)" "(\\<^bsub>G\<^esub>(K1 \\<^bsub>G\<^esub> K \ H)) / (K1 \\<^bsub>G\<^esub> K \ H1)"], assumption+) apply (frule ZassenhausTr5_1[of "G" "K" "K1" "H" "H1"], assumption+) apply (simp add:Int_commute[of "K" "H"]) apply (simp add:Group.ZassenhausTr3_3[THEN sym, of "G" "H" "H1" "K" "K1"]) apply (rule isomTr1[of "(\\<^bsub>G\<^esub>(H \ K)) / (H1 \ K \\<^bsub>G\<^esub> H \ K1)" "(\\<^bsub>G\<^esub>(H1 \\<^bsub>G\<^esub> H \ K)) / (H1 \\<^bsub>G\<^esub> H \ K1)"], assumption+) apply (rule ZassenhausTr5_1[of "G" "H" "H1" "K" "K1"], assumption+) done section "Chain of groups I" definition d_gchain :: "[_ , nat, (nat \ 'a set)] \ bool" where "d_gchain G n g = (if n=0 then G \ g 0 else (\l\ n. G \ (g l) \ (\l \ (n - Suc 0). g (Suc l) \ g l )))" (** g 0 \ \ \ g n **) definition D_gchain :: "[_ , nat, (nat \ 'a set)] \ bool" where "D_gchain G n g = (if n = 0 then G \ (g 0) else (d_gchain G n g) \ (\l \ (n - Suc 0). (g (Suc l)) \ (g l)))" (** g 0 \ \ \ g n **) definition td_gchain :: "[_ , nat, (nat \ 'a set)] \ bool" where "td_gchain G n g = (if n=0 then g 0 = carrier G \ g 0 = {\\<^bsub>G\<^esub>} else d_gchain G n g \ g 0 = carrier G \ g n = {\\<^bsub>G\<^esub>})" (** g 0 \ \ \ g n with g 0 = carrier G and g n = {e} **) definition tD_gchain :: "[_, nat, (nat \ 'a set)] \ bool" where "tD_gchain G n g = (if n=0 then g 0 = carrier G \ g 0 = {\\<^bsub>G\<^esub>} else D_gchain G n g \ (g 0 = carrier G) \ (g n = {\\<^bsub>G\<^esub>}))" (** g 0 \ \ \ g n with g 0 = carrier G and g n = {e} **) definition w_cmpser :: "[_ , nat, (nat \ 'a set)] \ bool" where "w_cmpser G n g = (if n = 0 then d_gchain G n g else d_gchain G n g \ (\l \ (n - 1). (Gp G (g l)) \ (g (Suc l))))" (** g 0 \ \ \ g n ** with g l \ g (Suc l) **) definition W_cmpser :: "[_ , nat, (nat \ 'a set)] \ bool" where "W_cmpser G n g = (if n = 0 then d_gchain G 0 g else D_gchain G n g \ (\l \ (n - 1). (Gp G (g l)) \ (g (Suc l))))" (** g 0 \ \ \ g n with g i \ g (Suc i) **) definition tw_cmpser :: "[_ , nat, (nat \ 'a set)] \ bool" where "tw_cmpser G n g = (if n = 0 then td_gchain G 0 g else td_gchain G n g \ (\l \ (n - 1). (Gp G (g l)) \ (g (Suc l))))" (** g 0 \ \ \ g n with g 0 = carrier G and g n = {e} **) definition tW_cmpser :: "[_ , nat, (nat \ 'a set)] \ bool" where "tW_cmpser G n g = (if n = 0 then td_gchain G 0 g else tD_gchain G n g \ (\l \ (n - 1). (Gp G (g l)) \ (g (Suc l))))" (* g 0 \ \ \ g n with g 0 = carrier G, g n = {e} and g (Suc i) \ g i*) definition Qw_cmpser :: "[_ , nat \ 'a set] \ (nat \ ('a set) Group)" where "Qw_cmpser G f l = ((Gp G (f l)) / (f (Suc l)))" (* f 0 \ \ \ f (n+1), Qw_cmpser G n f is (f 0)/(f 1),\,f n/f (n+1) *) definition red_chn :: "[_ , nat, (nat \ 'a set)] \ (nat \ 'a set)" where "red_chn G n f = (SOME g. g \ {h.(tW_cmpser G (card (f ` {i. i \ n}) - 1) h) \ h `{i. i \ (card (f ` {i. i \ n}) - 1)} = f `{i. i \ n}})" definition chain_cutout :: "[nat, (nat \ 'a set) ] \ (nat \ 'a set)" where "chain_cutout l f = (\j. f (slide l j))" lemma (in Group) d_gchainTr0:"\0 < n; d_gchain G n f; k \ (n - 1)\ \ f (Suc k) \ f k" apply (simp add:d_gchain_def) apply (frule_tac a = k in forall_spec) apply (rule Nat.le_trans, assumption+, simp) apply (erule conjE, rotate_tac 2, frule_tac a = k in forall_spec, assumption, thin_tac "\l\n - Suc 0. f (Suc l) \ f l", thin_tac "\l\n. G \ f l \ (\l\n - Suc 0. f (Suc l) \ f l)") apply assumption done lemma (in Group) d_gchain_mem_sg:"d_gchain G n f \ \i\ n. G \ (f i)" apply (rule allI) apply (rule impI, simp add:d_gchain_def) apply (case_tac "n = 0", simp) apply simp done lemma (in Group) d_gchain_pre:"d_gchain G (Suc n) f \ d_gchain G n f" apply (simp add:d_gchain_def, rule impI, rule impI) apply (rule allI, rule impI) apply (frule_tac a = l in forall_spec, arith) apply (erule conjE) apply (thin_tac "\l\Suc n. G \ f l \ (\l\n. f (Suc l) \ f l)", frule_tac a = l in forall_spec, arith, assumption) done lemma (in Group) d_gchainTr1:"0 < n \ (\f. d_gchain G n f \ (\l \ n. \j \ n. l < j \ f j \ f l))" apply (induct_tac n) apply (rule impI, simp) (** n **) apply (rule impI, rule allI, rule impI) apply ((rule allI, rule impI)+, rule impI) (** case n = 0 **) apply (case_tac "n = 0", simp) apply (case_tac "j = 0", simp, frule le_imp_less_or_eq, thin_tac "j \ Suc 0", simp, simp add:d_gchain_def) apply (frule_tac a = 0 in forall_spec, simp, simp) (** case 0 < n **) apply simp (** case j = n **) apply (case_tac "j = Suc n") apply (frule d_gchain_pre, frule_tac a = f in forall_spec, assumption, thin_tac "\f. d_gchain G n f \ (\l\n. \j\n. l < j \ f j \ f l)", thin_tac "d_gchain G n f", simp add:d_gchain_def) apply (frule_tac a = n in forall_spec, simp, thin_tac "\l\Suc n. G \ f l \ (\l\n. f (Suc l) \ f l)", erule conjE, rotate_tac -1, frule_tac a = n in forall_spec, simp, thin_tac "\l\n. f (Suc l) \ f l", frule_tac x = l and n = n in Suc_less_le) apply (case_tac "l = n", simp, thin_tac "l < Suc n", frule_tac x = l and y = n in le_imp_less_or_eq, thin_tac "l \ n", simp) apply (frule_tac a = l in forall_spec, simp, thin_tac "\l\n. \j\n. l < j \ f j \ f l") apply ( frule_tac a = n in forall_spec) apply (simp, thin_tac "\j\n. l < j \ f j \ f l", simp) apply (simp add: d_gchain_pre) done lemma (in Group) d_gchainTr2:"\0 < n; d_gchain G n f; l \ n; j \ n; l \ j \ \ f j \ f l" apply (case_tac "l = j", simp) -apply (metis Group.d_gchainTr1 [OF Group_axioms] linorder_antisym_conv2) +apply (metis Group.d_gchainTr1 [OF Group_axioms] antisym_conv2) done lemma (in Group) im_d_gchainTr1:"\d_gchain G n f; f l \ (f ` {i. i \ n}) - {f 0}\ \ f (LEAST j. f j \ (f ` {i. i \ n}) - {f 0}) \ (f ` {i. i \ n} - {f 0})" apply (rule LeastI) apply simp done lemma (in Group) im_d_gchainTr1_0:"\d_gchain G n f; f l \ (f ` {i. i \ n}) - {f 0}\ \ 0 < (LEAST j. f j \ (f ` {i. i \ n}) - {f 0})" apply (frule im_d_gchainTr1 [of "n" "f"], assumption+) apply (rule contrapos_pp, simp) apply simp done lemma (in Group) im_d_gchainTr1_1: "\d_gchain G n f; \ i. f i \ (f ` {i. i \ n}) - {f 0}\ \ f (LEAST j. f j \ ((f ` {i. i \ n}) - {f 0})) \ ((f` {i. i \ n}) - {f 0})" apply (subgoal_tac "\i. f i \ f ` {i. i \ n} - {f 0} \ f (LEAST j. f j \ f `{i. i \ n} - {f 0}) \ f ` {i. i\ n} - {f 0}") apply blast apply (thin_tac "\i. f i \ f `{i. i \ n} - {f 0}") apply (rule allI) apply (rule impI) apply (rule im_d_gchainTr1 [of "n" "f" _], assumption+) done lemma (in Group) im_d_gchainsTr1_2:" \d_gchain G n f; i \ n; f i \ f `{i. i \ n} - {f 0}\ \ (LEAST j. f j \ (f `{i. i \ n} - {f 0})) \ i" by (rule Least_le, assumption) lemma (in Group) im_d_gchainsTr1_3:"\d_gchain G n f; \i \ n. f i \ f` {i. i \ n} - {f 0}; k < (LEAST j. f j \ (f `{i. i \ n} - {f 0}))\ \ f k = f 0" apply (erule exE) apply (rule contrapos_pp, simp+) apply (frule_tac i = i in im_d_gchainsTr1_2 [of "n" "f" _ ], simp+) apply (erule conjE)+ apply (frule_tac x = k and y = "LEAST j. f j \ f ` {i. i \ n} \ f j \ f 0" and z = i in less_le_trans, assumption, frule_tac x = k and y = i and z = n in less_le_trans, assumption) apply (frule_tac i = k in im_d_gchainsTr1_2 [of "n" "f" _ ], simp+) done lemma (in Group) im_gdchainsTr1_4:"\d_gchain G n f; \v\f `{i. i \ n}. v \ {f 0}; i < (LEAST j. f j \ (f `{i. i \ n}) \ f j \ f 0) \ \ f i = f 0" apply (rule im_d_gchainsTr1_3 [of "n" "f" "i"], assumption, thin_tac "i < (LEAST j. f j \ f ` {i. i \ n} \ f j \ f 0)", simp add:image_def, blast) apply simp done lemma (in Group) im_d_gchainsTr1_5:"\0 < n; d_gchain G n f; i \ n; f i \ (f ` {i. i \ n} - {f 0}); (LEAST j. f j \ (f `{i. i \ n} - {f 0})) = j\ \ f `{i. i \ (j - (Suc 0))} = {f 0}" apply (frule im_d_gchainTr1_0 [of "n" "f" "i"], assumption+) apply (subst image_def) apply (rule equalityI) apply (rule subsetI, simp, erule exE, erule conjE) apply (frule_tac x = xa and y = "j - Suc 0" and z = "(LEAST j. f j \ f ` {i. i \ n} \ f j \ f 0)" in le_less_trans, simp, frule_tac k = xa in im_d_gchainsTr1_3[of "n" "f"]) apply (simp, blast, simp, simp) apply (rule subsetI, blast) done lemma (in Group) im_d_gchains1:"\0 < n; d_gchain G n f; i \ n; f i \ (f ` {i. i \ n} - {f 0}); (LEAST j. f j \ (f `{i. i \ n} - {f 0})) = j \ \ f `{i. i \ n} = {f 0} \ {f i |i. j \ i \ i \ n}" apply (frule im_d_gchainTr1_0 [of "n" "f" "i"], assumption+, frule im_d_gchainsTr1_2 [of "n" "f" "i"], assumption+, frule Nset_nset_1 [of "n" "j - Suc 0"]) apply simp apply (subst im_set_un2, simp) apply (subst im_d_gchainsTr1_5[of "n" "f" "i" "j"]) apply (simp, assumption, simp+) apply (rule equalityI) apply (rule subsetI, simp, erule disjE, simp, simp add:image_def nset_def, erule exE, blast) apply (rule subsetI) apply (simp, erule disjE, simp) apply (erule exE, simp add:nset_def) done lemma (in Group) im_d_gchains1_1:"\d_gchain G n f; f n \ f 0\ \ f `{i. i \ n} = {f 0} \ {f i |i. (LEAST j. f j \ (f `{i. i \ n} - {f 0})) \ i \ i \ n}" apply (case_tac "n = 0") apply simp apply simp apply (frule im_d_gchains1 [of "n" "f" "n" "(LEAST j. f j \ f ` {i. i \ n} - {f 0})"], assumption+, simp add:n_in_Nsetn) apply (cut_tac n_in_Nsetn[of "n"], simp, simp) apply (simp cong del: image_cong) done lemma (in Group) d_gchains_leastTr:"\d_gchain G n f; f n \ f 0\ \ (LEAST j. f j \ (f `{i. i \ n} - {f 0})) \ {i. i \ n} \ f (LEAST j. f j \ (f `{i. i \ n} - {f 0})) \ f 0" apply (frule im_d_gchainsTr1_2 [of "n" "f" "n"], simp add:n_in_Nsetn, simp add:image_def, blast, frule mem_of_Nset[of "LEAST j. f j \ f ` {i. i \ n} - {f 0}" "n"], simp) apply (frule im_d_gchainTr1[of "n" "f" "n"], simp add:image_def, cut_tac n_in_Nsetn[of "n"], blast) apply (simp add:image_def) done lemma (in Group) im_d_gchainTr2:"\d_gchain G n f; j \ n; f j \ f 0\ \ \i \ n. f 0 = f i \ \ j \ i" apply (case_tac "n = 0", simp, simp) apply (rule allI, rule impI) apply (rule contrapos_pp, simp+) apply (case_tac "j = i", simp) apply (frule d_gchainTr2 [of "n" "f" "0" "j"], assumption+, simp, (erule conjE)+, rule_tac i = j and j = i and k = n in le_trans, assumption+, simp, (erule conjE)+, frule_tac l = j and j = i in d_gchainTr2 [of "n" "f"], assumption+) apply simp+ done lemma (in Group) D_gchain_pre:"\D_gchain G (Suc n) f\ \ D_gchain G n f" apply (simp add:D_gchain_def, erule conjE) apply (case_tac "n = 0", rotate_tac -1) apply (simp, insert lessI [of "0::nat"], simp) apply (simp add:d_gchain_def, simp) apply (frule d_gchain_pre [of "n"]) apply simp done lemma (in Group) D_gchain0:"\D_gchain G n f; i \ n; j \ n; i < j\ \ f j \ f i" apply (case_tac "n = 0") apply (simp, simp) apply (cut_tac d_gchainTr1[of "n"], simp) apply (simp add:D_gchain_def, frule conjunct1) apply (frule_tac a = f in forall_spec, assumption, thin_tac "\f. d_gchain G n f \ (\l\n. \j\n. l < j \ f j \ f l)") apply (frule_tac a = i in forall_spec, frule_tac x = i and y = j and z = n in less_le_trans, assumption+, simp) apply ( thin_tac "\l\n. \j\n. l < j \ f j \ f l", frule_tac a = j in forall_spec, assumption, thin_tac "\j\n. i < j \ f j \ f i", simp) apply (frule Suc_leI[of i j], frule less_le_trans[of i j n], assumption+, frule less_le_diff[of i n]) apply (frule_tac a = i in forall_spec, assumption, thin_tac "\l\n - Suc 0. f (Suc l) \ f l") apply (cut_tac d_gchainTr2[of "n" "f" "Suc i" "j"]) apply blast apply simp+ done lemma (in Group) D_gchain1:"D_gchain G n f \ inj_on f {i. i \ n}" apply (case_tac "n = 0") apply (simp add:inj_on_def) apply (simp) (** case 0 < n **) apply (simp add:inj_on_def) apply ((rule allI)+, rule impI, rule contrapos_pp, simp+, erule exE) apply (cut_tac x = x and y = y in less_linear, simp) apply (erule disjE, (erule conjE)+) apply (frule_tac i = x and j = y in D_gchain0 [of "n" "f"], assumption+, simp, simp, frule_tac i = y and j = x in D_gchain0 [of "n" "f"], simp+) done lemma (in Group) card_im_D_gchain:"\0 < n; D_gchain G n f\ \ card (f `{i. i \ n}) = Suc n" apply (frule D_gchain1 [of "n"]) apply (subst card_image, assumption+, simp) done lemma (in Group) w_cmpser_gr:"\0 < r; w_cmpser G r f; i \ r\ \ G \ (f i)" by (simp add:w_cmpser_def, erule conjE, simp add:d_gchain_def) lemma (in Group) w_cmpser_ns:"\0 < r; w_cmpser G r f; i \ (r - 1)\ \ (Gp G (f i)) \ (f (Suc i))" apply (simp add:w_cmpser_def) done lemma (in Group) w_cmpser_pre:"w_cmpser G (Suc n) f \ w_cmpser G n f" apply (simp add:w_cmpser_def) apply (erule conjE) apply (case_tac "n = 0", rotate_tac -1, simp) apply (rule d_gchain_pre [of "0" "f"], assumption+) apply simp apply (frule d_gchain_pre [of "n" "f"]) apply simp done lemma (in Group) W_cmpser_pre:"W_cmpser G (Suc n) f \ W_cmpser G n f" apply (simp add:W_cmpser_def) apply (erule conjE) apply (case_tac "n = 0", simp, simp add:D_gchain_def, erule conjE, rule d_gchain_pre, assumption+, simp) apply (frule D_gchain_pre, simp) done lemma (in Group) td_gchain_n:"\td_gchain G n f; carrier G \ {\}\ \ 0 < n" apply (simp add:td_gchain_def) apply (rule contrapos_pp, simp+) apply (erule conjE, simp) done section "Existence of reduced chain" lemma (in Group) D_gchain_is_d_gchain:"D_gchain G n f \ d_gchain G n f" apply (simp add:D_gchain_def) apply (case_tac "n = 0") apply (rotate_tac -1) apply (simp add:d_gchain_def) apply (rotate_tac -1) apply simp done lemma (in Group) joint_d_gchains:"\d_gchain G n f; d_gchain G m g; g 0 \ f n \ \ d_gchain G (Suc (n + m)) (jointfun n f m g)" apply (case_tac "n = 0") apply (case_tac "m = 0") apply (simp add:d_gchain_def [of "G" "Suc (n + m)" _]) apply (simp add:jointfun_def sliden_def d_gchain_def) apply (simp add:jointfun_def sliden_def d_gchain_def) apply (rule allI) apply (rule conjI) apply (rule impI) apply (rule allI) apply (rule impI)+ apply simp apply (frule_tac a = la in forall_spec, assumption, thin_tac "\l\m. G \ g l \ (\l\m - Suc 0. g (Suc l) \ g l)", erule conjE) apply (frule_tac a = "la - Suc 0" in forall_spec, thin_tac "\l\m - Suc 0. g (Suc l) \ g l", rule diff_le_mono, assumption, simp) apply (rule impI, rule impI) apply (frule_tac m = l and n = "Suc m" and l = "Suc 0" in diff_le_mono) apply simp apply (rule allI, rule impI, rule impI) apply (frule_tac a = la in forall_spec,assumption, thin_tac "\l\m. G \ g l \ (\l\m - Suc 0. g (Suc l) \ g l)", erule conjE) apply (frule_tac a = "la - Suc 0" in forall_spec, simp) apply simp_all apply (simp add:d_gchain_def [of "G" _ "jointfun n f m g"]) apply (rule allI, rule impI) apply (rule conjI) apply (case_tac "l \ n", simp add:jointfun_def d_gchain_def[of _ n f]) apply (simp add:jointfun_def sliden_def, frule_tac m = l and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono, thin_tac "l \ Suc (n + m)", simp add:d_gchain_def[of _ m g]) apply (case_tac "m = 0", simp, simp) apply (rule allI, rule impI) apply (case_tac "Suc la \ n") apply (simp add:jointfun_def) apply (rule_tac l = la and j = "Suc la" in d_gchainTr2[of n f], simp+) apply (simp add:jointfun_def) apply (cut_tac y = "Suc la" and x = n in not_less [symmetric], simp) apply (frule_tac m = n and n = "Suc la" in Suc_leI, thin_tac "n < Suc la", simp) apply (case_tac "la = n", simp add:sliden_def) apply (frule not_sym, thin_tac "la \ n", frule_tac x = n and y = la in le_imp_less_or_eq, thin_tac "n \ la", simp, frule_tac m = n and n = la in Suc_leI, simp add:sliden_def) apply (simp add:d_gchain_def[of _ m g]) apply (cut_tac m = la and n = "n + m" and l = "Suc n" in diff_le_mono, assumption, simp) apply (frule_tac a = m in forall_spec, simp, thin_tac "\l\m. G \ g l \ (\l\m - Suc 0. g (Suc l) \ g l)", erule conjE) apply ( frule_tac a = "la - Suc n" in forall_spec, assumption, thin_tac "\l\m - Suc 0. g (Suc l) \ g l") apply (cut_tac n1 = n and i1 = la in jointgd_tool4[THEN sym], simp+) done lemma (in Group) joint_D_gchains:"\D_gchain G n f; D_gchain G m g; g 0 \ f n \ \ D_gchain G (Suc (n + m)) (jointfun n f m g)" apply (simp add:D_gchain_def [of "G" "Suc (n + m)" _]) apply (rule conjI) apply (rule joint_d_gchains[of n f m g], simp add:D_gchain_is_d_gchain, simp add:D_gchain_is_d_gchain, simp add:psubset_imp_subset) apply (rule allI, rule impI) apply (case_tac "Suc l \ n") apply (simp add:jointfun_def) apply (rule_tac i = l and j = "Suc l" in D_gchain0[of n f], assumption, cut_tac x = l and y = "Suc l" and z = n in less_le_trans) apply simp+ apply (simp add:nat_not_le_less, frule_tac m = n and n = "Suc l" in Suc_leI, thin_tac "n < Suc l", simp) apply (case_tac "l = n", simp add:jointfun_def sliden_def) apply (frule not_sym, thin_tac "l \ n", frule_tac x = n and y = l in le_imp_less_or_eq, thin_tac "n \ l", simp) apply (simp add:jointfun_def sliden_def) apply (frule_tac m = l and n = "n + m" and l = n in diff_le_mono) apply (simp add:diff_add_assoc) apply (rule_tac i = "l - Suc n" and j = "l - n" in D_gchain0[of m g], assumption) apply (arith, assumption, arith) done lemma (in Group) w_cmpser_is_d_gchain:"w_cmpser G n f \ d_gchain G n f" apply (simp add:w_cmpser_def) apply (case_tac "n=0") apply (rotate_tac -1) apply simp apply (rotate_tac -1) apply simp done lemma (in Group) joint_w_cmpser:"\w_cmpser G n f; w_cmpser G m g; Gp G (f n) \ (g 0)\ \ w_cmpser G (Suc (n + m)) (jointfun n f m g)" apply (simp add:w_cmpser_def [of _ "Suc (n + m)" _]) apply (rule conjI) apply (frule w_cmpser_is_d_gchain[of "n" "f"], frule w_cmpser_is_d_gchain[of "m" "g"]) apply (rule joint_d_gchains, assumption+) apply (frule d_gchain_mem_sg[of "n" "f"], cut_tac n_in_Nsetn[of "n"], frule_tac a = n in forall_spec, simp, thin_tac "\i \ n. G \ f i") apply (frule Group_Gp[of "f n"], frule Group.nsg_sg[of "Gp G (f n)" "g 0"], assumption, frule Group.sg_subset[of "Gp G (f n)" "g 0"], assumption, simp add:Gp_carrier) apply (rule allI, rule impI) apply (case_tac "n = 0") apply simp apply (simp add:jointfun_def) apply (case_tac "l = 0") apply simp apply (simp add:sliden_def) apply simp apply (simp add:w_cmpser_def [of _ "m" "g"]) apply (case_tac "m = 0") apply (simp add:sliden_def) apply (erule conjE) apply (simp add:sliden_def) apply (frule_tac x = 0 and y = l and z = m in less_le_trans, assumption+) apply (frule_tac m = l and n = m and l = "Suc 0" in diff_le_mono) apply (frule_tac a = "l - Suc 0" in forall_spec, assumption, thin_tac "\l\(m - Suc 0). (\(g l)) \ (g (Suc l))") apply simp apply (case_tac "l \ n - Suc 0", simp) apply (frule less_pre_n [of "n"]) apply (frule_tac x = l in le_less_trans[of _ "n - Suc 0" "n"], assumption+) apply (simp add:jointfun_def w_cmpser_def [of _ "n"]) apply (simp add:nat_not_le_less) apply (frule_tac n = l in Suc_leI [of "n - Suc 0" _], simp) apply (case_tac "n = l") apply (frule sym) apply (thin_tac "n = l") apply simp apply (simp add:jointfun_def sliden_def) apply (frule_tac m = n and n = l in noteq_le_less, assumption+) apply (frule_tac m = n and n = l in Suc_leI) apply (simp add:jointfun_def) apply (frule_tac m = l and n = "n + m" and l = "Suc n" in diff_le_mono) apply simp apply (simp add:sliden_def w_cmpser_def [of _ "m" _]) apply (erule conjE) apply (simp add:jointgd_tool4) done lemma (in Group) W_cmpser_is_D_gchain:"W_cmpser G n f \ D_gchain G n f" apply (simp add:W_cmpser_def) apply (case_tac "n = 0") apply (rotate_tac -1) apply simp apply (simp add:D_gchain_def d_gchain_def) apply (rotate_tac -1) apply simp done lemma (in Group) W_cmpser_is_w_cmpser:"W_cmpser G n f \ w_cmpser G n f" apply (simp add:W_cmpser_def) apply (case_tac "n = 0") apply (rotate_tac -1) apply simp apply (simp add:w_cmpser_def) apply (rotate_tac -1) apply simp apply (erule conjE) apply (frule D_gchain_is_d_gchain) apply (simp add:w_cmpser_def) done lemma (in Group) tw_cmpser_is_w_cmpser:"tw_cmpser G n f \ w_cmpser G n f" apply (simp add:tw_cmpser_def) apply (case_tac "n = 0") apply (rotate_tac -1) apply simp apply (simp add:td_gchain_def w_cmpser_def) apply (simp add:d_gchain_def) apply (simp add:special_sg_G) apply (rotate_tac -1) apply simp apply (erule conjE) apply (simp add:td_gchain_def) apply (erule conjE)+ apply (simp add:w_cmpser_def) done lemma (in Group) tW_cmpser_is_W_cmpser:"tW_cmpser G n f \ W_cmpser G n f" apply (simp add:tW_cmpser_def) apply (case_tac "n = 0") apply (rotate_tac -1) apply simp apply (simp add:td_gchain_def) apply (simp add:W_cmpser_def d_gchain_def) apply (simp add:special_sg_G) apply (rotate_tac -1) apply simp apply (erule conjE) apply (simp add:tD_gchain_def) apply (erule conjE)+ apply (simp add:W_cmpser_def) done lemma (in Group) joint_W_cmpser:"\W_cmpser G n f; W_cmpser G m g; (Gp G (f n)) \ (g 0); g 0 \ f n\ \ W_cmpser G (Suc (n + m)) (jointfun n f m g)" apply (simp add:W_cmpser_def [of _ "Suc (n + m)" _]) apply (frule W_cmpser_is_D_gchain [of "n" "f"], frule W_cmpser_is_D_gchain [of "m" "g"]) apply (simp add:joint_D_gchains) apply (frule W_cmpser_is_w_cmpser [of "n" _], frule W_cmpser_is_w_cmpser [of "m" _]) apply (frule joint_w_cmpser [of "n" "f" "m" "g"], assumption+) apply (simp add:w_cmpser_def [of _ "Suc (n + m)" _]) done lemma (in Group) joint_d_gchain_n0:"\d_gchain G n f; d_gchain G 0 g; g 0 \ f n \ \ d_gchain G (Suc n) (jointfun n f 0 g)" apply (frule joint_d_gchains [of "n" "f" "0" "g"], assumption+) apply simp done lemma (in Group) joint_D_gchain_n0:"\D_gchain G n f; D_gchain G 0 g; g 0 \ f n \ \ D_gchain G (Suc n) (jointfun n f 0 g)" apply (frule joint_D_gchains [of "n" "f" "0" "g"], assumption+) apply simp done lemma (in Group) joint_w_cmpser_n0:"\w_cmpser G n f; w_cmpser G 0 g; (Gp G (f n)) \ (g 0)\ \ w_cmpser G (Suc n) (jointfun n f 0 g)" apply (frule joint_w_cmpser [of "n" "f" "0" "g"], assumption+) apply simp done lemma (in Group) joint_W_cmpser_n0:"\W_cmpser G n f; W_cmpser G 0 g; (Gp G (f n)) \ (g 0); g 0 \ f n \ \ W_cmpser G (Suc n) (jointfun n f 0 g)" apply (frule joint_W_cmpser [of "n" "f" "0" "g"], assumption+) apply simp done definition simple_Group :: "_ \ bool" where "simple_Group G \ {N. G \ N} = {carrier G, {\\<^bsub>G\<^esub>}}" definition compseries:: "[_ , nat, nat \ 'a set] \ bool" where "compseries G n f \ tW_cmpser G n f \ (if n = 0 then f 0 = {\\<^bsub>G\<^esub>} else (\i \ (n - 1). (simple_Group ((Gp G (f i))/(f (Suc i))))))" definition length_twcmpser :: "[_ , nat, nat \ 'a set] \ nat" where "length_twcmpser G n f = card (f `{i. i \ n}) - Suc 0" lemma (in Group) compseriesTr0:"\compseries G n f; i \ n\ \ G \ (f i)" apply (simp add:compseries_def) apply (frule conjunct1) apply (fold compseries_def) apply (frule tW_cmpser_is_W_cmpser, frule W_cmpser_is_w_cmpser, frule w_cmpser_is_d_gchain) apply (frule d_gchain_mem_sg[of "n" "f"]) apply simp done lemma (in Group) compseriesTr1:"compseries G n f \ tW_cmpser G n f" apply (simp add:compseries_def) done lemma (in Group) compseriesTr2:"compseries G n f \ f 0 = carrier G" apply (frule compseriesTr1, simp add:tW_cmpser_def) apply (case_tac "n = 0") apply (simp add:td_gchain_def) apply simp apply (erule conjE, simp add:tD_gchain_def) done lemma (in Group) compseriesTr3:"compseries G n f \ f n = {\}" apply (frule compseriesTr1) apply (simp add:tW_cmpser_def) apply (case_tac "n = 0") apply (simp add:td_gchain_def) apply (auto del:equalityI) apply (simp add:tD_gchain_def) done lemma (in Group) compseriesTr4:"compseries G n f \ w_cmpser G n f" apply (frule compseriesTr1, frule tW_cmpser_is_W_cmpser) apply (rule W_cmpser_is_w_cmpser, assumption) done lemma (in Group) im_jointfun1Tr1:"\l \ n. G \ (f l) \ f \ {i. i \ n} \ Collect (sg G)" apply (simp add:Pi_def) done lemma (in Group) Nset_Suc_im:"\l \ (Suc n). G \ (f l) \ insert (f (Suc n)) (f ` {i. i \ n}) = f ` {i. i \ (Suc n)}" apply (rule equalityI) apply (rule subsetI) apply (simp add:image_def) apply (erule disjE) apply blast apply (erule exE, erule conjE) apply (frule_tac x = xa and y = n and z = "Suc n" in le_less_trans, simp, frule_tac x = xa and y = "Suc n" in less_imp_le, blast) apply (cut_tac Nset_Suc [of "n"], simp) done definition NfuncPair_neq_at::"[nat \ 'a set, nat \ 'a set, nat] \ bool" where "NfuncPair_neq_at f g i \ f i \ g i" lemma LeastTr0:"\ (i::nat) < (LEAST l. P (l))\ \ \ P (i)" apply (rule not_less_Least) apply simp done lemma (in Group) funeq_LeastTr1:"\\l\ n. G \ f l; \l \ n. G \ g l; (l :: nat) < (LEAST k. (NfuncPair_neq_at f g k)) \ \ f l = g l" apply (rule contrapos_pp, simp+) apply (frule LeastTr0 [of "l" "NfuncPair_neq_at f g"]) apply (simp add:NfuncPair_neq_at_def) done lemma (in Group) funeq_LeastTr1_1:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; (l :: nat) < (LEAST k. (f k \ g k)) \ \ f l = g l" apply (rule funeq_LeastTr1[of "n" "f" "g" "l"], assumption+) apply (simp add:NfuncPair_neq_at_def) done lemma (in Group) Nfunc_LeastTr2_1:"\i \ n; \l \ n. G \ f l; \l \ n. G \ g l; NfuncPair_neq_at f g i\ \ NfuncPair_neq_at f g (LEAST k. (NfuncPair_neq_at f g k))" apply (simp add: LeastI [of "NfuncPair_neq_at f g" "i"]) done lemma (in Group) Nfunc_LeastTr2_2:"\i \ n; \l \ n. G \ f l; \l \ n. G \ g l; NfuncPair_neq_at f g i\ \ (LEAST k. (NfuncPair_neq_at f g k)) \ i" apply (simp add: Least_le [of "NfuncPair_neq_at f g" "i"]) done lemma (in Group) Nfunc_LeastTr2_2_1:"\i \ (n::nat); \l \ n. G \ f l; \l \ n. G \ g l; f i \ g i\ \ (LEAST k. (f k \ g k)) \ i" apply (rule contrapos_pp, simp+) apply (simp add:nat_not_le_less) apply (frule Nfunc_LeastTr2_2 [of "i" "n" "f" "g"], assumption+) apply (simp add:NfuncPair_neq_at_def)+ done lemma (in Group) Nfunc_LeastTr2_3:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; i \ n; f i \ g i\ \ f (LEAST k. (f k \ g k)) \ g (LEAST k. (f k \ g k))" apply (frule Nfunc_LeastTr2_1 [of "i" "n" "f" "g"], assumption+) apply (simp add:NfuncPair_neq_at_def)+ done lemma (in Group) Nfunc_LeastTr2_4:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; i \ n; f i \ g i\ \(LEAST k. (f k \ g k)) \ n" apply (frule_tac i = i in Nfunc_LeastTr2_2 [of _ "n" "f" "g"], assumption+) apply (simp add:NfuncPair_neq_at_def) apply (frule le_trans [of "(LEAST k. NfuncPair_neq_at f g k)" "i" "n"], assumption+) apply (simp add:NfuncPair_neq_at_def) done lemma (in Group) Nfunc_LeastTr2_5:"\\l\ (n::nat). G \ f l; \l \ n. G \ g l; \i \ n. (f i \ g i)\ \ f (LEAST k. (f k \ g k)) \ g ((LEAST k. f k \ g k))" apply (erule exE) apply (rule_tac i = i in Nfunc_LeastTr2_3[of n f g], assumption+, simp+) done lemma (in Group) Nfunc_LeastTr2_6:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; \i \ n. (f i \ g i)\ \ (LEAST k. (f k \ g k)) \ n" apply (erule exE) apply (rule_tac i = i in Nfunc_LeastTr2_4, assumption+, simp+) done lemma (in Group) Nfunc_Least_sym:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; \i \ n. (f i \ g i)\ \ (LEAST k. (f k \ g k)) = (LEAST k. (g k \ f k))" apply (erule exE) apply (frule_tac i = i in Nfunc_LeastTr2_4 [of n f g], assumption+, simp+, frule_tac i = i in Nfunc_LeastTr2_3 [of n f g _], assumption+, simp+, frule_tac i = i in Nfunc_LeastTr2_4 [of n g f], assumption+, simp+, rule not_sym, simp, frule_tac i = i in Nfunc_LeastTr2_3 [of n g f _], assumption+, simp+, rule not_sym, simp) apply (frule_tac i = "(LEAST k. f k \ g k)" in Nfunc_LeastTr2_2_1 [of _ "n" "g" "f"], assumption+, rule not_sym, simp) apply ( frule_tac i = "(LEAST k. g k \ f k)" in Nfunc_LeastTr2_2_1 [of _ n f g], assumption+, rule not_sym, simp) apply (rule le_antisym, assumption+) done lemma Nfunc_iNJTr:"\inj_on g {i. i \ (n::nat)}; i \ n; j \ n; i < j \ \ g i \ g j" apply (unfold inj_on_def) apply (simp add:CollectI) apply (rule contrapos_pp, simp+) apply (frule_tac a = i in forall_spec) apply (frule_tac x = i and y = j and z = n in less_le_trans, assumption+, simp add:less_imp_le, thin_tac "\x\n. \y\n. g x = g y \ x = y", rotate_tac -1, frule_tac a = j in forall_spec, assumption, thin_tac "\y\n. g i = g y \ i = y") apply simp done lemma (in Group) Nfunc_LeastTr2_7:"\\l \ (n::nat). G \ f l; \l \ n. G \ g l; inj_on g {i. i \ n}; \i \ n. (f i \ g i); f k = g (LEAST k.(f k \ g k))\ \(LEAST k.(f k \ g k)) < k" apply (rule contrapos_pp, simp+) apply (simp add:nat_not_le_less[THEN sym, of "LEAST k. f k \ g k" "k"]) apply (frule le_imp_less_or_eq) apply (case_tac "k = (LEAST k. f k \ g k)") apply simp apply (frule Nfunc_LeastTr2_5 [of "n" "f" "g"], assumption+) apply simp apply (frule funeq_LeastTr1_1 [of "n" "f" "g" "k"], assumption+) apply simp apply (frule Nfunc_LeastTr2_6 [of "n" "f" "g"], assumption+) apply simp apply (frule_tac le_trans[of "k" "LEAST k. f k \ g k" "n"], assumption+) apply (frule mem_of_Nset[of "k" "n"]) apply (simp add:inj_on_def[of "g"]) done lemma (in Group) Nfunc_LeastTr2_8:"\\l \ n. G \ f l; \l \ n. G \ g l; inj_on g {i. i \ n}; \i \ n. f i \ g i; f `{i. i \ n} = g `{i. i \ n}\ \ \ k \(nset (Suc (LEAST i. (f i \ g i))) n). f k = g (LEAST i. (f i \ g i))" apply (frule_tac Nfunc_LeastTr2_6 [of "n" "f" "g"], assumption+) apply (cut_tac mem_in_image2[of "LEAST k. f k \ g k" "{i. i \ n}" "g"]) apply (frule sym, thin_tac "f ` {i. i \ n} = g ` {i. i \ n}", simp) apply (thin_tac "g ` {i. i \ n} = f ` {i. i \ n}") apply (simp add:image_def) apply (rotate_tac -1, erule exE) apply (frule_tac k = x in Nfunc_LeastTr2_7[of "n" "f" "g"], assumption+) apply (erule conjE) apply (rule sym, assumption) apply (frule_tac m = "(LEAST k. f k \ g k)" and n = x in Suc_leI) apply (simp add:nset_def) apply blast apply simp done lemma (in Group) ex_redchainTr1:"\d_gchain G n f; D_gchain G (card (f ` {i. i \ n}) - Suc 0) g; g ` {i. i \ (card (f ` {i. i \ n}) - Suc 0)} = f ` {i. i \ n}\ \ g (card (f ` {i. i \ n}) - Suc 0) = f n" apply (case_tac "n = 0", simp, simp) apply (rule contrapos_pp, simp+) apply (cut_tac n_in_Nsetn[of "card (f ` {i. i \ n}) - Suc 0"]) apply (frule mem_in_image2[of "card (f ` {i. i \ n}) - Suc 0" "{i. i \ (card (f ` {i. i \ n}) - Suc 0)}" "g"]) apply (cut_tac n_in_Nsetn[of "n"]) apply (frule mem_in_image2[of "n" "{i. i \ n}" "f"]) apply simp apply (simp add:image_def[of "f" "{i. i \ n}"]) apply (erule exE) apply (frule_tac l = x in d_gchainTr2[of "n" "f" _ "n"], assumption+) apply simp+ apply (erule conjE) apply (rotate_tac -1, frule sym, thin_tac "g (card {y. \x\n. y = f x} - Suc 0) = f x", simp, thin_tac "f x = g (card {y. \x\n. y = f x} - Suc 0)") apply (cut_tac mem_in_image2[of "n" "{i. i \ n}" "f"], unfold image_def) apply (frule sym, thin_tac "{y. \x\{i. i \ card {y. \x\n. y = f x} - Suc 0}. y = g x} = {y. \x\n. y = f x}") apply (cut_tac eq_set_inc[of "f n" "{y. \x \ n. y = f x}" "{y. \x\{i. i \ card {y. \x\n. y = f x} - Suc 0}. y = g x}"]) apply (thin_tac "{y. \x\n. y = f x} = {y. \x\{i. i \ card {y. \x\n. y = f x} - Suc 0}. y = g x}", thin_tac "f n \ {y. \x\{i. i \ n}. y = f x}") apply (simp, erule exE, erule conjE) apply (case_tac "xa \ card {y. \x\n. y = f x} - Suc 0", simp) apply (frule D_gchain_is_d_gchain[of "card {y. \x\n. y = f x} - Suc 0" g]) apply (case_tac "card {y. \x \ n. y = f x} - Suc 0 = 0", simp) apply (frule nat_nonzero_pos[of "card {y. \x \ n. y = f x} - Suc 0"], thin_tac "card {y. \x\n. y = f x} - Suc 0 \ 0") apply (frule_tac l = xa in d_gchainTr2[of "card {y. \x \ n. y = f x} - Suc 0" "g" _ "card {y. \x \ n. y = f x} - Suc 0"], assumption+) apply simp apply simp apply (frule_tac A = "g xa" and B = "g (card {y. \x \ n. y = f x} - Suc 0)" in equalityI, assumption+, simp) apply simp+ done lemma (in Group) ex_redchainTr1_1:"\d_gchain G (n::nat) f; D_gchain G (card (f ` {i. i \ n}) - Suc 0) g; g ` {i. i \ (card (f ` {i. i \ n}) - Suc 0)} = f ` {i. i \ n}\ \ g 0 = f 0" apply (cut_tac Nset_inc_0[of "n"], frule mem_in_image2[of "0" "{i. i \ n}" "f"]) apply ( frule sym) apply ( thin_tac "g ` {i. i \ card (f ` {i. i \ n}) - Suc 0} = f ` {i. i \ n}") apply ( frule eq_set_inc[of "f 0" "f ` {i. i \ n}" "g ` {i. i \ card (f ` {i. i \ n}) - Suc 0}"], assumption) apply ( thin_tac "f 0 \ f ` {i. i \ n}", thin_tac "0 \ {i. i \ n}") apply (cut_tac Nset_inc_0[of "card (f ` {i. i \ n}) - Suc 0"], frule mem_in_image2[of "0" "{i. i \ (card (f ` {i. i \ n}) - Suc 0)}" "g"], frule sym) apply ( frule eq_set_inc[of "g 0" "g ` {i. i \ card (f ` {i. i \ n}) - Suc 0}" "f ` {i. i \ n}"], assumption) apply ( thin_tac "f ` {i. i \ n} = g ` {i. i \ card (f ` {i. i \ n}) - Suc 0}") apply ( thin_tac "0 \ {i. i \ card (f ` {i. i \ n}) - Suc 0}") apply ( thin_tac "g ` {i. i \ card (f ` {i. i \ n}) - Suc 0} = f ` {i. i \ n}") apply ( thin_tac "g 0 \ g ` {i. i \ card (f ` {i. i \ n}) - Suc 0}") apply (case_tac "n = 0", simp) apply (simp) apply (cut_tac mem_in_image3[of "f 0" "g" "{i. i \ card (f ` {i. i \ n}) - Suc 0}"], frule mem_in_image3[of "g 0" "f" "{i. i \ n}"]) apply ( thin_tac "f 0 \ g ` {i. i \ card (f ` {i. i \ n}) - Suc 0}", thin_tac "g 0 \ f ` {i. i \ n}") apply (erule bexE)+ apply ( frule_tac j = aa in d_gchainTr2[of "n" "f" "0"], assumption+) apply simp+ apply (rotate_tac -2, frule sym, thin_tac "g 0 = f aa", simp) apply (case_tac "a = 0", simp) apply (simp, frule_tac j = a in D_gchain0[of "card (f ` {i. i \ n}) - Suc 0" g 0], simp add:Nset_inc_0, assumption+, simp add:psubset_contr) apply simp done lemma (in Group) ex_redchainTr2:"d_gchain G (Suc n) f \ D_gchain G 0 (constmap {0::nat} {f (Suc n)})" apply (simp add:D_gchain_def constmap_def) apply (simp add:d_gchain_def) done lemma (in Group) last_mem_excluded:"\d_gchain G (Suc n) f; f n \ f (Suc n)\ \ f (Suc n) \ f ` {i. i \ n}" apply (rule contrapos_pp, simp+) apply (frule mem_in_image3[of "f (Suc n)" "f" "{i. i \ n}"], erule bexE) apply (cut_tac zero_less_Suc[of "n"]) apply (frule_tac l = a in d_gchainTr2[of "Suc n" "f" _ "n"], assumption+) apply simp+ apply (frule sym, thin_tac "f (Suc n) = f a", simp) apply (cut_tac l = n and j = "Suc n" in d_gchainTr2[of "Suc n" "f"]) apply simp+ done lemma (in Group) ex_redchainTr4:"\d_gchain G (Suc n) f; f n \ f (Suc n)\ \ card (f ` {i. i \ (Suc n)}) = Suc (card (f ` {i. i \ n}))" apply (cut_tac image_Nset_Suc [of "f" "n"]) apply simp apply (rule card_insert_disjoint) apply (simp) apply (simp add:last_mem_excluded) done lemma (in Group) ex_redchainTr5:"d_gchain G n f \ 0 < card (f ` {i. i\ n})" apply (simp add:image_Nsetn_card_pos) done lemma (in Group) ex_redchainTr6:"\f. d_gchain G n f \ (\g. D_gchain G (card (f `{i. i \ n}) - 1) g \ (g ` {i. i \ (card (f `{i. i \ n}) - 1)} = f `{i. i \ n}))" apply (induct_tac n) apply (rule allI, rule impI) apply (simp add:image_def) apply (simp add:D_gchain_def d_gchain_def) apply blast (** n **) apply (rule allI, rule impI) apply (case_tac "f (Suc n) = f n") apply (cut_tac n = n in Nset_Suc) apply (cut_tac n = n in n_in_Nsetn, frule_tac a = n and A = "{i. i \ n}" and f = f in mem_in_image2) apply (frule sym) apply (thin_tac "f (Suc n) = f n", simp) apply (subst insert_absorb, assumption)+ apply (frule_tac n = n and f = f in d_gchain_pre, blast) apply (frule_tac n = n and f = f in d_gchain_pre) apply (frule_tac a = f in forall_spec, assumption, thin_tac "\f. d_gchain G n f \ (\g. D_gchain G (card (f ` {i. i \ n}) - 1) g \ g ` {i. i \ card (f ` {i. i \ n}) - 1} = f ` {i. i \ n})") apply (erule exE, erule conjE) apply (simp add:image_Nset_Suc) apply (frule_tac n = n and f = f in ex_redchainTr2) apply (frule_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g and g = "constmap {0} {f (Suc n)}" in joint_D_gchain_n0, assumption+) apply (simp add: ex_redchainTr1) apply (simp add: constmap_def Nset_inc_0) apply (cut_tac n = n in zero_less_Suc) apply (frule_tac n = "Suc n" and f = f and l = n and j = "Suc n" in d_gchainTr2, assumption+) apply simp apply simp apply simp apply (simp add:psubset_eq) apply (cut_tac f = f and n = n in image_Nsetn_card_pos, cut_tac k = n in finite_Collect_le_nat, frule_tac F = "{i. i \ n}" and h = f in finite_imageI, frule_tac n = n and f = f in last_mem_excluded, rule not_sym, assumption) apply simp+ apply (cut_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g and m = 0 and g = "constmap {0} {f (Suc n)}" in im_jointfun1) apply simp apply (simp add:Nset_0 constmap_def) apply blast done lemma (in Group) ex_redchain:"d_gchain G n f \ (\g. D_gchain G (card (f ` {i. i \ n}) - 1) g \ g ` {i. i \ (card (f ` {i. i \ n}) - 1)} = f ` {i. i \ n})" apply (cut_tac ex_redchainTr6 [of "n"]) apply simp done lemma (in Group) const_W_cmpser:"d_gchain G (Suc n) f \ W_cmpser G 0 (constmap {0::nat} {f (Suc n)})" apply (simp add:W_cmpser_def d_gchain_def constmap_def) done lemma (in Group) ex_W_cmpserTr0m:"\f. w_cmpser G m f \ (\g. (W_cmpser G (card (f `{i. i \ m}) - 1) g \ g `{i. i \ (card (f `{i. i \ m}) - 1)} = f `{i. i \ m}))" apply (induct_tac m) (********* induct_step m = 0 ***************) apply (rule allI, rule impI) apply simp apply (simp add:w_cmpser_def W_cmpser_def) apply blast (********** induct_step m ****************) apply (rule allI, rule impI) apply (case_tac "f (Suc n) = f n") apply (cut_tac n = n in Nset_Suc) apply (cut_tac n = n in n_in_Nsetn, frule_tac a = n and A = "{i. i \ n}" and f = f in mem_in_image2) apply (frule sym) apply (thin_tac "f (Suc n) = f n", simp) apply (subst insert_absorb, assumption)+ apply (frule_tac n = n and f = f in w_cmpser_pre, blast) apply (frule_tac n = n and f = f in w_cmpser_pre) apply (frule_tac a = f in forall_spec, assumption, thin_tac "\f. w_cmpser G n f \ (\g. W_cmpser G (card (f ` {i. i \ n}) - 1) g \ g ` {i. i \ card (f ` {i. i \ n}) - 1} = f ` {i. i \ n})") apply (erule exE, erule conjE) apply (simp add:image_Nset_Suc, frule_tac n = "Suc n" and f = f in w_cmpser_is_d_gchain, frule_tac n = n and f = f in const_W_cmpser) apply (frule_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g and g = "constmap {0::nat} {f (Suc n)}" in joint_W_cmpser_n0, assumption+) apply (frule_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g in W_cmpser_is_D_gchain) apply (frule d_gchain_pre) apply (subst ex_redchainTr1, assumption+) apply (simp add:constmap_def Nset_inc_0) apply (simp add:w_cmpser_def) apply (frule d_gchain_pre) apply (frule_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g in W_cmpser_is_D_gchain) apply (frule_tac n = n and f = f and g = g in ex_redchainTr1, assumption+) apply simp apply (simp add:constmap_def Nset_inc_0, thin_tac "d_gchain G n f", simp add:d_gchain_def) apply (cut_tac n = "Suc n" in n_in_Nsetn, frule_tac x = "Suc n" in spec, simp, simp add:psubset_eq) apply (cut_tac f = f and n = n in image_Nsetn_card_pos, cut_tac k = n in finite_Collect_le_nat, frule_tac F = "{i. i \ n}" and h = f in finite_imageI, frule_tac n = n and f = f in last_mem_excluded, rule not_sym, assumption) apply simp+ apply (cut_tac n = "card (f ` {i. i \ n}) - Suc 0" and f = g and m = 0 and g = "constmap {0::nat} {f (Suc n)}" in im_jointfun1) apply simp apply (simp add:Nset_0 constmap_def) apply blast done lemma (in Group) ex_W_cmpser:"w_cmpser G m f \ \g. W_cmpser G (card (f ` {i. i \ m}) - 1) g \ g ` {i. i \ (card (f ` {i. i \ m}) - 1)} = f ` {i. i \ m}" apply (cut_tac ex_W_cmpserTr0m [of "m"]) apply simp done section "Existence of reduced chain and composition series" lemma (in Group) ex_W_cmpserTr3m1:"\tw_cmpser G (m::nat) f; W_cmpser G ((card (f ` {i. i \ m})) - 1) g; g ` {i. i \ ((card (f ` {i. i \ m})) - 1)} = f `{i. i \ m}\ \ tW_cmpser G ((card (f ` {i. i \ m})) - 1) g" apply (frule_tac tw_cmpser_is_w_cmpser [of "m" "f"], frule_tac w_cmpser_is_d_gchain [of "m" "f"], frule_tac W_cmpser_is_D_gchain [of "(card (f ` {i. i \ m}) - 1)" "g"]) apply (frule ex_redchainTr1 [of "m" "f" "g"]) apply simp+ apply (frule_tac ex_redchainTr1_1 [of "m" "f" "g"]) apply (simp add:tW_cmpser_def tw_cmpser_def) apply (case_tac "m = 0") apply simp apply (cut_tac card_image_le [of "{0::nat}" "f"]) apply (simp, simp) apply (simp add:tW_cmpser_def) apply (case_tac "card (f ` {i. i \ m}) \ Suc 0") apply simp apply (simp add:td_gchain_def tw_cmpser_def) apply (case_tac "m = 0") apply (thin_tac "f 0 = f m", thin_tac "g 0 = f m") apply simp apply ( simp add:td_gchain_def) apply ( erule conjE, simp) apply simp apply (simp add:td_gchain_def[of "G" "m" "f"], (erule conjE)+, simp) apply simp apply (simp add:tD_gchain_def tw_cmpser_def td_gchain_def [of _ "m" "f"]) apply (case_tac "m = 0", simp add:card1) apply (simp, erule conjE, simp add:td_gchain_def) apply (simp add:W_cmpser_def) done lemma (in Group) ex_W_cmpserTr3m:"tw_cmpser G m f \ \g. tW_cmpser G ((card (f ` {i. i \ m})) - 1) g \ g `{ i. i \ (card (f `{i. i \ m}) - 1)} = f ` {i. i \ m}" apply (frule tw_cmpser_is_w_cmpser[of "m" "f"]) apply (frule ex_W_cmpser [of "m" "f"]) apply (auto del:equalityI) apply (frule_tac g = g in ex_W_cmpserTr3m1 [of "m" "f"]) apply simp+ apply blast done definition red_ch_cd :: "[_ , nat \ 'a set, nat, nat \ 'a set ] \ bool" where "red_ch_cd G f m g \ tW_cmpser G (card (f ` {i. i \ m}) - 1) g \ (g `{i . i \ (card (f ` {i. i \ m}) - 1)} = f` {i. i \ m})" definition red_chain :: "[_ , nat, nat \ 'a set] \ (nat \ 'a set)" where "red_chain G m f = (SOME g. g \ {h. red_ch_cd G f m h})" lemma (in Group) red_chainTr0m1_1:"tw_cmpser G m f \ (SOME g. g \ {h. red_ch_cd G f m h}) \ {h. red_ch_cd G f m h}" apply (rule nonempty_some [of "{h. red_ch_cd G f m h}"]) apply (frule ex_W_cmpserTr3m [of "m" "f"]) apply simp apply (simp add:red_ch_cd_def) done lemma (in Group) red_chain_m:"tw_cmpser G m f \ tW_cmpser G (card (f ` {i. i \ m}) - 1) (red_chain G m f) \ (red_chain G m f) `{i. i \ (card (f `{i. i \ m}) - 1)} = f ` {i. i \ m}" apply (frule red_chainTr0m1_1 [of "m" "f"]) apply simp apply (simp add:red_ch_cd_def) apply (simp add:red_chain_def) done section "Chain of groups II" definition Gchain :: "[nat, nat \ (('a set), 'more) Group_scheme] \ bool" where "Gchain n g \ (\l \ n. Group (g l))" definition isom_Gchains :: "[nat, nat \ nat, nat \ (('a set), 'more) Group_scheme, nat \ (('a set), 'more) Group_scheme] \ bool" where "isom_Gchains n f g h \ (\i \ n. (g i) \ (h (f i)))" (* g, h are sequences of groups and f is gbijection of Nset to Nset *) definition Gch_bridge :: "[nat, nat \ (('a set), 'more) Group_scheme, nat \ (('a set), 'more) Group_scheme, nat \ nat] \ bool" where "Gch_bridge n g h f \ (\l \ n. f l \ n) \ inj_on f {i. i \ n} \ isom_Gchains n f g h" lemma Gchain_pre:"Gchain (Suc n) g \ Gchain n g" apply (simp add:Gchain_def) done lemma (in Group) isom_unit:"\G \ H; G \ K; H = {\}\ \ Gp G H \ Gp G K \ K = {\}" apply (simp add:isomorphic_def) apply (rule impI) apply (erule exE) apply (simp add:gbijec_def) apply (erule conjE) apply (simp add:gsurjec_def ginjec_def) apply (erule conjE) apply (simp add:Gp_carrier) apply (simp add:surj_to_def) apply (cut_tac a = "f \" in finite1) apply (frule sg_unit_closed [of "K"]) apply (frule singleton_sub[of "\" "K"]) apply (rotate_tac 4, frule sym, thin_tac "{f \} = K") apply (rule card_seteq[THEN sym, of "K" "{\}"]) apply (simp add:finite1) apply assumption apply (simp add:card1) done lemma isom_gch_unitsTr4:"\Group F; Group G; Ugp E; F \ G; F \ E\ \ G \ E" apply (simp add:Ugp_def) apply (erule conjE) apply (frule isomTr1 [of "F" "G"], assumption+) apply (rule isomTr2 [of "G" "F" "E"], assumption+) done lemma isom_gch_cmp:"\Gchain n g; Gchain n h; f1 \ {i. i \ n} \ {i. i \ n}; f2 \ {i. i \ n} \ {i. i \ n}; isom_Gchains n (cmp f2 f1) g h\ \ isom_Gchains n f1 g (cmp h f2)" apply (simp add:isom_Gchains_def) apply (simp add:cmp_def) done lemma isom_gch_transp:"\Gchain n f; i \ n; j \ n; i < j\ \ isom_Gchains n (transpos i j) f (cmp f (transpos i j))" apply (rule isom_gch_cmp [of "n" "f" _ "transpos i j" "transpos i j"], assumption+) apply (rule transpos_hom, assumption+) apply simp apply (rule transpos_hom, assumption+) apply simp apply (simp add:isom_Gchains_def) apply (rule allI, rule impI) apply (frule less_le_trans [of "i" "j" "n"], assumption+) apply (frule less_imp_le [of "i" "n"]) apply (frule_tac k = ia in cmp_transpos1 [of "i" "n" "j"], assumption+) apply simp+ apply (simp add:Gchain_def) done lemma isom_gch_units_transpTr0:"\Ugp E; Gchain n g; Gchain n h; i \ n; j \ n; i < j; isom_Gchains n (transpos i j) g h\ \ {i. i \ n \ g i \ E} - {i, j} ={i. i \ n \ h i \ E} - {i, j}" apply (simp add:isom_Gchains_def) apply (rule equalityI) apply (rule subsetI, simp add:CollectI) apply (erule conjE)+ apply (cut_tac x = x in transpos_id_1 [of "i" "n" "j"], simp+) apply (subgoal_tac "g x \ h (transpos i j x)", thin_tac "\ia\n. g ia \ h (transpos i j ia)", simp) apply (subgoal_tac "Group (g x)", subgoal_tac "Group (h x)") apply (simp add:Ugp_def) apply (erule conjE) apply (frule_tac F = "g x" and G = "h x" in isomTr1, assumption+) apply (rule_tac F = "h x" and G = "g x" and H = E in isomTr2, assumption+) apply (simp add:Gchain_def, simp add:Gchain_def) apply (thin_tac "transpos i j x = x") apply simp apply (rule subsetI, simp add:CollectI) apply (erule conjE)+ apply (cut_tac x = x in transpos_id_1 [of "i" "n" "j"], simp+) apply (subgoal_tac "g x \ h (transpos i j x)", thin_tac "\ia\n. g ia \ h (transpos i j ia)") apply simp apply (subgoal_tac "Group (g x)", subgoal_tac "Group (h x)") apply (simp add:Ugp_def) apply (erule conjE) apply (rule_tac F = "g x" and G = "h x" and H = E in isomTr2, assumption+) apply (simp add:Gchain_def, simp add:Gchain_def) apply (thin_tac "transpos i j x = x") apply simp done lemma isom_gch_units_transpTr1:"\Ugp E; Gchain n g; i \ n; j \ n; g j \ E; i \ j\ \ insert j ({i. i \ n \ g i \ E} - {i, j}) = {i. i \ n \ g i \ E} - {i}" apply (rule equalityI) apply (rule subsetI) apply (simp add:CollectI) apply blast apply (rule subsetI) apply (simp add:CollectI) done lemma isom_gch_units_transpTr2:"\Ugp E; Gchain n g; i \ n; j \ n; i < j; g i \ E\ \ {i. i \ n \ g i \ E} = insert i ({i. i \ n \ g i \ E} - {i})" apply (rule equalityI) apply (rule subsetI, simp add:CollectI) apply (rule subsetI, simp add:CollectI) apply (erule disjE, simp) apply simp done lemma isom_gch_units_transpTr3:"\Ugp E; Gchain n g; i \ n\ \ finite ({i. i \ n \ g i \ E} - {i})" apply (rule finite_subset[of "{i. i \ n \ g i \ E} - {i}" "{i. i \ n}"]) apply (rule subsetI, simp+) done lemma isom_gch_units_transpTr4:"\Ugp E; Gchain n g; i \ n\ \ finite ({i. i \ n \ g i \ E} - {i, j})" apply (rule finite_subset[of "{i. i \ n \ g i \ E} - {i, j}" "{i. i \ n}"]) apply (rule subsetI, simp+) done lemma isom_gch_units_transpTr5_1:"\Ugp E; Gchain n g; Gchain n h; i \ (n::nat); j \ n; i < j; isom_Gchains n (transpos i j) g h\ \ g i \ h j" apply (simp add:isom_Gchains_def) apply (frule_tac a = i in forall_spec, frule_tac x = i and y = j and z = n in less_le_trans, assumption+, simp, thin_tac "\ia \ n. g ia \ h (transpos i j ia)") apply (simp add:transpos_ij_1 [of "i" "n" "j"]) done lemma isom_gch_units_transpTr5_2:"\Ugp E; Gchain n g; Gchain n h; i \ n; j \ n; i < j; isom_Gchains n (transpos i j) g h\ \ g j \ h i" apply (simp add:isom_Gchains_def) apply (frule_tac x = j in spec, thin_tac "\ia\ n. g ia \ h (transpos i j ia)") apply (simp add:transpos_ij_2 [of "i" "n" "j"]) done lemma isom_gch_units_transpTr6:"\Gchain n g; i \ n\ \ Group (g i)" apply (simp add:Gchain_def) done lemma isom_gch_units_transpTr7:"\Ugp E; i \ n; j \ n; g j \ h i; Group (h i); Group (g j); \ g j \ E\ \ \ h i \ E" apply (rule contrapos_pp, simp+) apply (frule isomTr2 [of "g j" "h i" "E"], assumption+) apply (simp add:Ugp_def) apply assumption+ apply simp done lemma isom_gch_units_transpTr8_1:"\Ugp E; Gchain n g; i \ n; j \ n; g i \ E; \ g j \ E\ \ {i. i \ n \ g i \ E} = {i. i \ n \ g i \ E} - { j }" by auto lemma isom_gch_units_transpTr8_2:"\Ugp E; Gchain n g; i \ n; j \ n; \ g i \ E; \ g j \ E\ \ {i. i \ n \ g i \ E} = {i. i \ n \ g i \ E} - {i, j }" by auto lemma isom_gch_units_transp:"\Ugp E; Gchain n g; Gchain n h; i \ n; j \ n; i < j; isom_Gchains n (transpos i j) g h\ \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}" apply (frule isom_gch_units_transpTr0 [of "E" "n" "g" "h" "i" "j"], assumption+) apply (frule isom_gch_units_transpTr6 [of "n" "g" "i"], assumption+) apply (frule isom_gch_units_transpTr6 [of "n" "h" "i"], assumption+) apply (frule isom_gch_units_transpTr6 [of "n" "g" "j"], assumption+) apply (frule isom_gch_units_transpTr6 [of "n" "h" "j"], assumption+) apply (unfold Ugp_def) apply (frule conjunct1) apply (fold Ugp_def) apply (frule isom_gch_units_transpTr5_1 [of "E" "n" "g" "h" "i" "j"], assumption+) apply (frule isom_gch_units_transpTr5_2 [of "E" "n" "g" "h" "i" "j"], assumption+) apply (case_tac "g i \ E") apply (case_tac "g j \ E") (** g i \ E and g j \ E **) apply (subst isom_gch_units_transpTr2 [of "E" "n" "g" "i" "j"], assumption+) apply (subst isom_gch_units_transpTr2 [of "E" "n" "h" "i" "j"], assumption+) apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+) apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr3, assumption+) apply simp apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr3, assumption+) apply simp apply (subst isom_gch_units_transpTr1[THEN sym, of "E" "n" "g" "i" "j"], assumption+) apply simp apply (subst isom_gch_units_transpTr1[THEN sym, of "E" "n" "h" "i" "j"], assumption+) apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+) apply simp apply simp apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "g" "i" "j"], assumption+) apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "h" "j" "i"], assumption+) apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+) apply (rule isom_gch_units_transpTr7 [of "E" "i" "n" "j" "g" "h"], assumption+) apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "g" "j" "i"], assumption+) apply simp apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr4, assumption+) apply simp apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "h" "i" "j"], assumption+) apply (rule isom_gch_unitsTr4 [of "g i" "h j" "E"], assumption+) apply simp apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr4, assumption+) apply simp apply (insert Nset_2 [of "j" "i"]) apply simp apply (case_tac "g j \ E") apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "g" "j" "i"], assumption+) apply (subst isom_gch_units_transpTr8_1 [of "E" "n" "h" "i" "j"], assumption+) apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+) apply (rule isom_gch_units_transpTr7 [of "E" "j" "n" "i" "g" "h"], assumption+) apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "g" "i" "j"], assumption+) apply simp apply (subst isom_gch_units_transpTr1 [THEN sym, of "E" "n" "h" "j" "i"], assumption+) apply (rule isom_gch_unitsTr4 [of "g j" "h i" "E"], assumption+) apply simp apply simp apply (subst isom_gch_units_transpTr8_2 [of "E" "n" "g" "i" "j"], assumption+) apply (subst isom_gch_units_transpTr8_2 [of "E" "n" "h" "i" "j"], assumption+) apply (rule isom_gch_units_transpTr7[of "E" "i" "n" "j" "g" "h"], assumption+) apply (rule isom_gch_units_transpTr7[of "E" "j" "n" "i" "g" "h"], assumption+) apply simp done lemma TR_isom_gch_units:"\Ugp E; Gchain n f; i \ n; j \ n; i < j\ \ card {k. k \ n \ f k \ E} = card {k. k \ n \ (cmp f (transpos i j)) k \ E}" apply (frule isom_gch_transp [of "n" "f" "i" "j"], assumption+) (* apply (subgoal_tac "Gchain n (cmp f (transpos i j))") *) apply (rule isom_gch_units_transp [of "E" "n" "f" _ "i" "j"], assumption+) apply (simp add:Gchain_def) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac l = l in transpos_mem [of "i" "n" "j"], frule_tac x = i and y = j and z = n in less_le_trans, assumption+, simp) apply simp+ done lemma TR_isom_gch_units_1:"\Ugp E; Gchain n f; i \ n; j \ n; i < j\ \ card {k. k \ n \ f k \ E} = card {k. k \ n \ f (transpos i j k) \ E}" apply (frule TR_isom_gch_units [of "E" "n" "f" "i" "j"], assumption+) apply (simp add:cmp_def) done lemma isom_tgch_unitsTr0_1:"\Ugp E; Gchain (Suc n) g; g (Suc n) \ E\ \ {i. i \ (Suc n) \ g i \ E} = insert (Suc n) {i. i \ n \ g i \ E}" apply (rule equalityI) apply (rule subsetI) apply (simp add:CollectI) apply (case_tac "x = Suc n") apply simp apply (erule conjE) apply simp apply (rule subsetI) apply (simp add:CollectI) apply (erule disjE) apply simp apply simp done lemma isom_tgch_unitsTr0_2:"Ugp E \ finite ({i. i \ (n::nat) \ g i \ E})" apply (rule finite_subset[of "{i. i \ n \ g i \ E}" "{i. i \ n}"]) apply (rule subsetI, simp+) done lemma isom_tgch_unitsTr0_3:"\Ugp E; Gchain (Suc n) g; \ g (Suc n) \ E\ \ {i. i \ (Suc n) \ g i \ E} = {i. i \ n \ g i \ E}" apply (rule equalityI) apply (rule subsetI, simp add:CollectI) apply (case_tac "x = Suc n", simp, erule conjE) apply arith apply (rule subsetI, simp add:CollectI) done lemma isom_tgch_unitsTr0:"\Ugp E; card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}; Gchain (Suc n) g \ Gchain (Suc n) h \ Gch_bridge (Suc n) g h f; f (Suc n) = Suc n\ \ card {i. i \ (Suc n) \ g i \ E} = card {i. i \ (Suc n) \ h i \ E}" apply (erule conjE)+ apply (frule isom_gch_units_transpTr6 [of "Suc n" "g" "Suc n"]) apply (simp add:n_in_Nsetn) apply (frule isom_gch_units_transpTr6 [of "Suc n" "h" "Suc n"]) apply (simp add:n_in_Nsetn) apply (simp add:Gch_bridge_def isom_Gchains_def) apply (erule conjE)+ apply (rotate_tac -1, frule_tac a = "Suc n" in forall_spec, thin_tac "\i\Suc n. g i \ h (f i)", simp+, thin_tac "\i\Suc n. g i \ h (f i)") apply (case_tac "g (Suc n) \ E") apply (subst isom_tgch_unitsTr0_1 [of "E" "n" "g"], assumption+) apply (subst isom_tgch_unitsTr0_1 [of "E" "n" "h"], assumption+) apply (frule isom_gch_unitsTr4 [of "g (Suc n)" "h (Suc n)" "E"], assumption+) apply (subst card_insert_disjoint) apply (rule finite_subset[of "{i. i \ n \ g i \ E}" "{i. i \ n}"]) apply (rule subsetI, simp) apply (simp) apply simp apply (subst card_insert_disjoint) apply (rule finite_subset[of "{i. i \ n \ h i \ E}" "{i. i \ n}"]) apply (rule subsetI, simp) apply simp apply simp apply simp apply (cut_tac isom_gch_units_transpTr7[of E "Suc n" "Suc n" "Suc n" g h]) apply (subgoal_tac "{i. i \ Suc n \ g i \ E} = {i. i \ n \ g i \ E}", subgoal_tac "{i. i \ Suc n \ h i \ E} = {i. i \ n \ h i \ E}", simp) apply (rule equalityI, rule subsetI, simp, erule conjE, case_tac "x = Suc n", simp, frule_tac x = x and y = "Suc n" in le_imp_less_or_eq, thin_tac "x \ Suc n", simp, rule subsetI, simp) apply (rule equalityI, rule subsetI, simp, erule conjE, case_tac "x = Suc n", simp, frule_tac m = x and n = "Suc n" in noteq_le_less, assumption+, simp, rule subsetI, simp) apply simp+ done lemma isom_gch_unitsTr1_1:" \Ugp E; Gchain (Suc n) g \ Gchain (Suc n) h \ Gch_bridge (Suc n) g h f; f (Suc n) = Suc n\ \ Gchain n g \ Gchain n h \ Gch_bridge n g h f" apply (erule conjE)+ apply (frule Gchain_pre [of "n" "g"]) apply (frule Gchain_pre [of "n" "h"]) apply simp apply (simp add:Gch_bridge_def) apply (erule conjE)+ apply (rule conjI) apply (rule Nset_injTr2, assumption+) apply (rule conjI) apply (rule Nset_injTr1, assumption+) apply (simp add:isom_Gchains_def) done lemma isom_gch_unitsTr1_2:"\Ugp E; f (Suc n) \ Suc n; inj_on f {i. i\(Suc n)}; \l \ (Suc n). f l \ (Suc n)\ \ (cmp (transpos (f (Suc n)) (Suc n)) f) (Suc n) = Suc n" apply (simp add:cmp_def) apply (cut_tac n_in_Nsetn[of "Suc n"], simp) apply (frule_tac a = "Suc n" in forall_spec, simp, thin_tac "\l\Suc n. f l \ Suc n") apply (rule transpos_ij_1, assumption+, simp+) done lemma isom_gch_unitsTr1_3:"\Ugp E; f (Suc n) \ Suc n; \l \ (Suc n). f l \ (Suc n); inj_on f {i. i \ (Suc n)}\ \ inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i \ (Suc n)}" apply (cut_tac n_in_Nsetn[of "Suc n"], simp) apply (frule_tac a = "Suc n" in forall_spec, simp) apply (frule transpos_hom [of "f (Suc n)" "Suc n" "Suc n"], simp+) thm transpos_inj apply (cut_tac transpos_inj [of "f (Suc n)" "Suc n" "Suc n"]) apply (cut_tac cmp_inj_1 [of f "{i. i \ (Suc n)}" "{i. i \ (Suc n)}" "transpos (f (Suc n)) (Suc n)" "{i. i \ (Suc n)}"]) apply simp+ done lemma isom_gch_unitsTr1_4:"\Ugp E; f (Suc n) \ Suc n; inj_on f {i. i\(Suc n)}; \l \ (Suc n). f l \ (Suc n)\ \ inj_on (cmp (transpos (f (Suc n)) (Suc n)) f) {i. i \ n}" apply (frule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+) apply (frule isom_gch_unitsTr1_2 [of "E" "f" "n"], assumption+) apply (rule Nset_injTr1 [of "n" "(cmp (transpos (f (Suc n)) (Suc n)) f)"]) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac l = "f l" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"], simp+) done lemma isom_gch_unitsTr1_5:"\Ugp E; Gchain (Suc n) g \ Gchain (Suc n) h \ Gch_bridge (Suc n) g h f; f (Suc n) \ Suc n \ \ Gchain n g \ Gchain n (cmp h (transpos (f (Suc n)) (Suc n))) \ Gch_bridge n g (cmp h (transpos (f (Suc n)) (Suc n))) (cmp (transpos (f (Suc n)) (Suc n)) f)" apply (erule conjE)+ apply (simp add:Gchain_pre [of "n" "g"]) apply (rule conjI) apply (simp add:Gchain_def) apply (rule allI, rule impI) apply (simp add:Gch_bridge_def) apply (frule conjunct1) apply (fold Gch_bridge_def) apply (cut_tac n_in_Nsetn[of "Suc n"]) apply (cut_tac l = l in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"]) apply simp+ apply (simp add:cmp_def) apply (simp add:Gch_bridge_def) apply (erule conjE)+ apply (rule conjI) apply (cut_tac n_in_Nsetn[of "Suc n"]) apply (rule allI, rule impI, simp add:cmp_def) apply (frule isom_gch_unitsTr1_2 [of "E" "f" "n"], assumption+) apply (frule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+) apply (cut_tac l = "f l" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"]) apply simp+ apply (simp add:inj_on_def[of "cmp (transpos (f (Suc n)) (Suc n)) f"]) apply (rotate_tac -2, frule_tac a = "Suc n" in forall_spec, simp) apply ( thin_tac "\x\Suc n. \y\Suc n. cmp (transpos (f (Suc n)) (Suc n)) f x = cmp (transpos (f (Suc n)) (Suc n)) f y \ x = y") apply ( rotate_tac -1, frule_tac x = l in spec) apply ( thin_tac "\y\Suc n. cmp (transpos (f (Suc n)) (Suc n)) f (Suc n) = cmp (transpos (f (Suc n)) (Suc n)) f y \ Suc n = y") apply (metis Nfunc_iNJTr comp_transpos_1 le_SucE le_SucI le_refl less_Suc_eq_le transpos_ij_2) apply (simp add:isom_gch_unitsTr1_4) apply (simp add:isom_Gchains_def[of "n"]) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac l = "f i" in transpos_mem[of "f (Suc n)" "Suc n" "Suc n"]) apply simp+ apply (cut_tac k = "f i" in cmp_transpos1 [of "f (Suc n)" "Suc n" "Suc n"]) apply simp+ apply (simp add:cmp_def) apply (thin_tac "transpos (f (Suc n)) (Suc n) (transpos (f (Suc n)) (Suc n) (f i)) = f i") apply (simp add:isom_Gchains_def) done lemma isom_gch_unitsTr1_6:"\Ugp E; f (Suc n) \ Suc n; Gchain (Suc n) g \ Gchain (Suc n) h \ Gch_bridge (Suc n) g h f\ \ Gchain (Suc n) g \ Gchain (Suc n) (cmp h (transpos (f (Suc n)) (Suc n))) \ Gch_bridge (Suc n) g (cmp h (transpos (f (Suc n)) (Suc n))) (cmp (transpos (f (Suc n)) (Suc n)) f)" apply (erule conjE)+ apply simp apply (simp add:Gch_bridge_def, frule conjunct1) apply (frule conjunct2, fold Gch_bridge_def, erule conjE) apply (rule conjI) apply (thin_tac "Gchain (Suc n) g") apply (simp add:Gchain_def, rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac n_in_Nsetn[of "Suc n"]) apply (cut_tac l = l in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"], simp+) apply (simp add:Gch_bridge_def) apply (cut_tac n_in_Nsetn[of "Suc n"]) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (rule_tac l = "f l" in transpos_mem [of "f (Suc n)" "Suc n" "Suc n"], simp+) apply (rule conjI) apply (rule isom_gch_unitsTr1_3 [of "E" "f" "n"], assumption+) apply (simp add:isom_Gchains_def, rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac k = "f i" in cmp_transpos1 [of "f (Suc n)" "Suc n" "Suc n"], simp+) apply (simp add:cmp_def) done lemma isom_gch_unitsTr1_7_0:"\Gchain (Suc n) h; k \ Suc n; k \ (Suc n)\ \ Gchain (Suc n) (cmp h (transpos k (Suc n)))" apply (simp add:Gchain_def) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (insert n_in_Nsetn [of "Suc n"]) apply (cut_tac l = l in transpos_mem [of "k" "Suc n" "Suc n"]) apply simp+ done lemma isom_gch_unitsTr1_7_1:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n)\ \ {i. i \ (Suc n) \ cmp h (transpos k (Suc n)) i \ E} - {k , Suc n} = {i. i \ (Suc n) \ h i \ E} - {k , Suc n}" apply (cut_tac n_in_Nsetn[of "Suc n"]) apply auto apply (frule_tac x = x in transpos_id_1 [of "k" "Suc n" "Suc n"], simp+) apply (simp add:cmp_def) apply (simp add:cmp_def) apply (cut_tac x = x in transpos_id_1 [of "k" "Suc n" "Suc n"], simp+) done lemma isom_gch_unitsTr1_7_2:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); h (Suc n) \ E\ \ cmp h (transpos k (Suc n)) k \ E" apply (simp add:cmp_def) apply (subst transpos_ij_1 [of "k" "Suc n" "Suc n"], simp+) done lemma isom_gch_unitsTr1_7_3:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); h k \ E\ \ cmp h (transpos k (Suc n)) (Suc n) \ E" apply (simp add:cmp_def) apply (subst transpos_ij_2 [of "k" "Suc n" "Suc n"], assumption+) apply simp apply assumption+ done lemma isom_gch_unitsTr1_7_4:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); \ h (Suc n) \ E\ \ \ cmp h (transpos k (Suc n)) k \ E" apply (rule contrapos_pp, simp+) apply (simp add:cmp_def) apply (insert n_in_Nsetn [of "Suc n"]) apply (simp add: transpos_ij_1 [of "k" "Suc n" "Suc n"]) done lemma isom_gch_unitsTr1_7_5:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); \ h k \ E\ \ \ cmp h (transpos k (Suc n)) (Suc n) \ E" apply (rule contrapos_pp, simp+) apply (simp add:cmp_def) apply (insert n_in_Nsetn [of "Suc n"]) apply (simp add:transpos_ij_2 [of "k" "Suc n" "Suc n"]) done lemma isom_gch_unitsTr1_7_6:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); h (Suc n) \ E; h k \ E\ \ {i. i \ (Suc n) \ h i \ E} = insert k (insert (Suc n) ({i. i \ (Suc n) \ h i \ E} - {k, Suc n}))" apply auto done lemma isom_gch_unitsTr1_7_7:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); h (Suc n) \ E; \ h k \ E\ \ {i. i \ (Suc n) \ h i \ E} = insert (Suc n) ({i. i \ (Suc n) \ h i \ E} - {k, Suc n})" apply auto done lemma isom_gch_unitsTr1_7_8:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); \ h (Suc n) \ E; h k \ E\ \ {i. i \ (Suc n) \ h i \ E} = insert k ({i. i \ (Suc n) \ h i \ E} - {k, Suc n})" apply auto done lemma isom_gch_unitsTr1_7_9:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n); \ h (Suc n) \ E; \ h k \ E\ \ {i. i \ (Suc n) \ h i \ E} = {i. i \ (Suc n) \ h i \ E} - {k, Suc n}" apply auto done lemma isom_gch_unitsTr1_7:"\Ugp E; Gchain (Suc n) h; k \ Suc n; k \ (Suc n)\ \ card {i. i \ (Suc n) \ cmp h (transpos k (Suc n)) i \ E} = card {i. i \ (Suc n) \ h i \ E}" apply (cut_tac finite_Collect_le_nat[of "Suc n"]) apply (frule isom_gch_unitsTr1_7_1 [of "E" "n" "h" "k"], assumption+) apply (cut_tac n_in_Nsetn[of "Suc n"]) apply (case_tac "h (Suc n) \ E") apply (case_tac "h k \ E") apply (subst isom_gch_unitsTr1_7_6 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_2 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_3 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+) apply (subst isom_gch_unitsTr1_7_6 [of "E" "n" "cmp h (transpos k (Suc n))" "k"], assumption+) apply (subst card_insert_disjoint) apply (rule finite_subset[of "insert (Suc n) ({i. i \ (Suc n) \ cmp h (transpos k (Suc n)) i \ E} - {k, Suc n})" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply (erule disjE) apply simp apply simp apply assumption apply simp apply (subst card_insert_disjoint)+ apply (rule finite_subset[of "{i. i \ (Suc n) \ cmp h (transpos k (Suc n)) i \ E} - {k, Suc n}" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply assumption apply simp apply (subst card_insert_disjoint)+ apply (rule finite_subset[of "insert (Suc n) ({i. i \ (Suc n) \ h i \ E} - {k, Suc n})" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply (erule disjE, simp add:n_in_Nsetn) apply simp apply assumption apply simp apply (subst card_insert_disjoint) apply (rule finite_subset[of "{i. i \ (Suc n) \ h i \ E} - {k, Suc n}" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply assumption apply simp apply simp apply (subst isom_gch_unitsTr1_7_7[of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_0[of "n" "h" "k"], assumption+) apply (subst isom_gch_unitsTr1_7_8[of "E" "n" "cmp h (transpos k (Suc n))" "k"], assumption+) apply (subst cmp_def) apply (subst transpos_ij_2[of "k" "Suc n" "Suc n"], assumption+, simp+) apply (simp add:cmp_def, simp add:transpos_ij_1) apply (subst card_insert_disjoint) apply (rule finite_subset[of "{i. i \ (Suc n) \ cmp h (transpos k (Suc n)) i \ E} - {k, Suc n}" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply assumption apply simp apply (subst card_insert_disjoint) apply (rule finite_subset[of "{i. i \ (Suc n) \ h i \ E} - {k, Suc n}" "{i. i \ (Suc n)}"]) apply (rule subsetI, simp) apply assumption apply simp apply simp apply (case_tac "h k \ E") apply (subst isom_gch_unitsTr1_7_8 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_3 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_4 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+) apply (subst isom_gch_unitsTr1_7_7 [of "E" "n" "cmp h (transpos k (Suc n))" "k"], assumption+) apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr4, assumption+) apply simp apply (subst card_insert_disjoint) apply (rule isom_gch_units_transpTr4, assumption+) apply simp apply simp apply (subst isom_gch_unitsTr1_7_9 [of "E" "n" "h" "k"], assumption+) apply (frule_tac isom_gch_unitsTr1_7_4 [of "E" "n" "h" "k"], assumption+) apply (frule_tac isom_gch_unitsTr1_7_5 [of "E" "n" "h" "k"], assumption+) apply (frule isom_gch_unitsTr1_7_0 [of "n" "h" "k" ], assumption+) apply (subst isom_gch_unitsTr1_7_9 [of "E" "n" " cmp h (transpos k (Suc n))" "k"], assumption+) apply simp done lemma isom_gch_unitsTr1:"Ugp E \ \g. \h. \f. Gchain n g \ Gchain n h \ Gch_bridge n g h f \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}" apply (induct_tac n) apply (clarify) apply (simp add:Gch_bridge_def isom_Gchains_def Collect_conv_if) apply rule apply (simp add:Gchain_def) apply(metis isom_gch_unitsTr4) apply (simp add:Gchain_def) apply (metis Ugp_def isomTr2) (***** n ******) apply (rule allI)+ apply (rule impI) (** n ** case f (Suc n) = Suc n **) apply (case_tac "f (Suc n) = Suc n") apply (subgoal_tac "card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}") apply (thin_tac " \g h f. Gchain n g \ Gchain n h \ Gch_bridge n g h f \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}") apply (rule isom_tgch_unitsTr0, assumption+) apply (frule_tac n = n and g = g and h = h and f = f in isom_gch_unitsTr1_1 [of "E"], assumption+) apply (rotate_tac -1) apply (thin_tac "Gchain (Suc n) g \ Gchain (Suc n) h \ Gch_bridge (Suc n) g h f") apply blast (**** n **** f (Suc n) \ Suc n ***) apply (frule_tac n = n and g = g and h = h and f = f in isom_gch_unitsTr1_5 [of "E"], assumption+) apply (subgoal_tac "card {i. i \ n \ g i \ E} = card {i. i \ n \ (cmp h (transpos (f (Suc n)) (Suc n))) i \ E}") prefer 2 apply blast apply (thin_tac "\g h f. Gchain n g \ Gchain n h \ Gch_bridge n g h f \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}") apply (thin_tac "Gchain n g \ Gchain n (cmp h (transpos (f (Suc n)) (Suc n))) \ Gch_bridge n g (cmp h (transpos (f (Suc n)) (Suc n))) (cmp (transpos (f (Suc n)) (Suc n)) f)") apply (subgoal_tac "cmp (transpos (f (Suc n)) (Suc n)) f (Suc n) = Suc n") apply (frule_tac n = n and g = g and h = "cmp h (transpos (f (Suc n)) (Suc n))" and f = "cmp (transpos (f (Suc n)) (Suc n)) f" in isom_tgch_unitsTr0 [of "E"], assumption+) apply (rule isom_gch_unitsTr1_6, assumption+) apply (thin_tac "card {i. i \ n \ g i \ E} = card {i. i \ n \ cmp h (transpos (f (Suc n)) (Suc n)) i \ E}") prefer 2 apply (erule conjE)+ apply (simp add:Gch_bridge_def) apply (erule conjE)+ apply (rule isom_gch_unitsTr1_2, assumption+) apply simp apply (erule conjE)+ apply (rule isom_gch_unitsTr1_7, assumption+) apply (simp add:Gch_bridge_def) done lemma isom_gch_units:"\Ugp E; Gchain n g; Gchain n h; Gch_bridge n g h f\ \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}" apply (simp add:isom_gch_unitsTr1) done lemma isom_gch_units_1:"\Ugp E; Gchain n g; Gchain n h; \f. Gch_bridge n g h f\ \ card {i. i \ n \ g i \ E} = card {i. i \ n \ h i \ E}" apply (auto del:equalityI) apply (simp add:isom_gch_units) done section "Jordan Hoelder theorem" subsection \\Rfn_tools\. Tools to treat refinement of a cmpser, rtos.\ lemma rfn_tool1:"\ 0 < (r::nat); (k::nat) = i * r + j; j < r \ \ (k div r) = i" proof - assume p1: "0 < r" and p2: "k = i * r + j" and p3: "j < r" from p1 and p2 have q1: "(j + i * r) div r = i + j div r" apply (simp add:div_mult_self1 [of "r" "j" "i" ]) done from p1 and p3 have q2: "j div r = 0" apply (simp add:div_if) done from q1 and q2 have q3:"(j + i * r) div r = i" apply simp done from q3 have q4: "(i * r + j) div r = i" apply (simp add:add.commute) done from p2 and q4 show ?thesis apply simp done qed lemma pos_mult_pos:"\ 0 < (r::nat); 0 < s\ \ 0 < r * s" by simp lemma rfn_tool1_1:"\ 0 < (r::nat); j < r \ \ (i * r + j) div r = i" apply (rule rfn_tool1 [of "r" "i * r + j" "i" "j"], assumption+) apply simp+ done lemma rfn_tool2:"(a::nat) < s \ a \ s - Suc 0" apply (rule less_le_diff) apply simp+ done lemma rfn_tool3:"(0::nat) \ m \ (m + n) - n = m" apply auto done lemma rfn_tool11:"\0 < b; (a::nat) \ b - Suc 0\ \ a < b" apply (frule le_less_trans [of "a" "b - Suc 0" "b"]) apply simp+ done lemma rfn_tool12:"\0 < (s::nat); (i::nat) mod s = s - 1 \ \ Suc (i div s) = (Suc i) div s " proof - assume p1:"0 < s" and p2:"i mod s = s - 1" have q1:"i div s * s + i mod s = i" apply (insert div_mult_mod_eq [of "i" "s"]) apply simp done have q2:"Suc (i div s * s + i mod s) = i div s * s + Suc (i mod s)" apply (insert add_Suc_right [THEN sym, of "i div s * s" "i mod s"]) apply assumption done from p1 and p2 and q2 have q3:"Suc (i div s * s + i mod s) = i div s * s + s" apply simp done from q3 have q4:"Suc (i div s * s + i mod s) = Suc (i div s) * s " apply simp done from p1 and q1 and q4 show ?thesis apply auto done qed lemma rfn_tool12_1:"\0 < (s::nat); (l::nat) mod s < s - 1 \ \ Suc (l mod s) = (Suc l) mod s " apply (insert div_mult_mod_eq [of "l" "s"]) apply (insert add_Suc_right [THEN sym, of "l div s * s" "l mod s"]) apply (insert mod_mult_self1 [of "Suc (l mod s)" "l div s" "s"]) apply (frule Suc_mono [of "l mod s" "s - 1"]) apply simp done lemma rfn_tool12_2:"\0 < (s::nat); (i::nat) mod s = s - Suc 0\ \ (Suc i) mod s = 0" apply (insert div_mult_mod_eq [THEN sym, of "i" "s"]) apply (insert add_Suc_right [THEN sym, of "i div s * s" "i mod s"]) apply simp done lemma rfn_tool13:"\ (0::nat) < r; a = b \ \ a mod r = b mod r" apply simp done lemma rfn_tool13_1:"\ (0::nat) < r; a = b \ \ a div r = b div r" apply simp done lemma div_Tr1:"\ (0::nat) < r; 0 < s; l \ s * r\ \ l div s \ r" apply (frule_tac m = l and n = "s * r" and k = s in div_le_mono) apply simp done lemma div_Tr2:"\(0::nat) < r; 0 < s; l < s * r\ \ l div s \ r - Suc 0" apply (rule contrapos_pp, simp+) apply (simp add: not_less [symmetric, of "l div s" "r - Suc 0"]) apply (frule Suc_leI [of "r - Suc 0" "l div s"]) apply simp apply (frule less_imp_le [of "l" "s * r"]) apply (frule div_le_mono [of "l" "s * r" "s"]) apply simp apply (insert div_mult_mod_eq [THEN sym, of "l" "s"]) apply (frule sym) apply (thin_tac "r = l div s") apply simp apply (simp add:mult.commute [of "r" "s"]) done lemma div_Tr3:"\(0::nat) < r; 0 < s; l < s * r\ \ Suc (l div s) \ r" apply (frule_tac div_Tr2[of "r" "s"], assumption+, cut_tac n1 = "l div s" and m1 = "r - Suc 0" in Suc_le_mono[THEN sym]) apply simp done lemma div_Tr3_1:"\(0::nat) < r; 0 < s; l mod s = s - 1\ \ Suc l div s = Suc (l div s)" apply (frule rfn_tool12 [of "s" "l"], assumption+) apply (rotate_tac -1) apply simp done lemma div_Tr3_2:"\(0::nat) < r; 0 < s; l mod s < s - 1\ \ l div s = Suc l div s" apply (frule Suc_mono [of "l mod s" "s - 1"]) apply simp apply (cut_tac div_mult_mod_eq [of "l" "s"]) apply (cut_tac add_Suc_right [THEN sym, of "l div s * s" "l mod s"]) apply (cut_tac zero_less_Suc[of "l mod s"], frule less_trans[of "0" "Suc (l mod s)" "s"], assumption+) apply (frule rfn_tool13_1 [of "s" "Suc (l div s * s + l mod s)" "l div s * s + Suc (l mod s)"], assumption+) apply (subgoal_tac "s \ 0") apply (frule div_mult_self1 [of "s" "Suc (l mod s)" "l div s"]) apply simp_all done lemma mod_div_injTr:"\(0::nat) < r; x mod r = y mod r; x div r = y div r\ \ x = y" apply (cut_tac div_mult_mod_eq[of "x" "r"]) apply simp done definition rtos :: "[nat, nat] \ (nat \ nat)" where "rtos r s i = (if i < r * s then (i mod s) * r + i div s else r * s)" lemma rtos_hom0:"\(0::nat) < r; (0::nat) < s; i \ (r * s - Suc 0)\ \ i div s < r" apply (frule div_Tr2 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute [of "r" "s"]) apply (rule le_less_trans, assumption+) apply simp apply (rule le_less_trans, assumption+) apply simp done lemma rtos_hom1:"\(0::nat) < r; 0 < s; l \ (r * s - Suc 0)\ \ (rtos r s) l \ (s * r - Suc 0)" apply (simp add:rtos_def) apply (frule le_less_trans [of "l" "r * s - Suc 0" "r * s"]) apply simp apply simp apply (frule mod_less_divisor [of "s" "l"]) apply (frule less_le_diff [of "l mod s" "s"]) apply (frule_tac i = "l mod s" and j = "s - Suc 0" and k = r and l = r in mult_le_mono, simp) apply (frule_tac i = "l mod s * r" and j = "(s - Suc 0) * r" and k = "l div s" and l = "r - Suc 0" in add_le_mono) apply (rule div_Tr2, assumption+, simp add:mult.commute) apply (simp add:diff_mult_distrib) done lemma rtos_hom2:"\(0::nat) < r; (0::nat) < s; l \ (r * s - Suc 0)\ \ rtos r s l \ (r * s - Suc 0)" apply (insert rtos_hom1 [of "r" "s"]) apply simp apply (simp add:mult.commute) done lemma rtos_hom3:"\(0::nat) < r; 0 < s; i \ (r * s - Suc 0) \ \ (rtos r s i div r) = i mod s" apply (simp add:rtos_def) apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"]) apply simp apply simp apply (auto simp add: div_mult2_eq [symmetric] mult.commute) done lemma rtos_hom3_1:"\(0::nat) < r; (0::nat) < s; i \ (r * s - Suc 0) \ \ (rtos r s i mod r) = i div s" apply (simp add:rtos_def) apply (simp add:rfn_tool11 [of "r * s" "i"]) apply (frule rtos_hom0 [of "r" "s" "i"], assumption+) apply (simp add:mem_of_Nset) done lemma rtos_hom5:"\(0::nat) < r; (0::nat) < s; i \ (r *s - Suc 0); i div s = r - Suc 0 \ \ Suc (rtos r s i) div r = Suc (i mod s)" apply (frule mult_less_mono2[of "0" "s" "r"], simp only:mult.commute, simp only:mult_0_right) apply (frule rfn_tool11 [of "r * s" "i"], assumption+) apply (simp add: rtos_def) done lemma rtos_hom7:"\(0::nat) < r; (0::nat) < s; i \ (r * s - Suc 0); i div s = r - Suc 0 \ \ Suc (rtos r s i) mod r = 0" apply (frule rtos_hom0 [of "r" "s" "i"], assumption+) apply (simp add: rtos_def) apply (frule mult_less_mono2[of "0" "s" "r"], simp only:mult.commute, simp only:mult_0_right) apply (simp add: rfn_tool11 [of "r * s" "i"]) done lemma rtos_inj:"\ (0::nat) < r; (0::nat) < s \ \ inj_on (rtos r s) {i. i \ (r * s - Suc 0)}" apply (simp add:inj_on_def) apply ((rule allI, rule impI)+, rule impI) apply (frule_tac i1 = x in rtos_hom3_1[THEN sym, of "r" "s"], assumption+, frule_tac i1 = x in rtos_hom3[THEN sym, of "r" "s"], assumption+, frule_tac i = y in rtos_hom3_1[of "r" "s"], assumption+, frule_tac i = y in rtos_hom3[of "r" "s"], assumption+) apply simp apply (rule_tac x = x and y = y in mod_div_injTr[of "s"], assumption+) done lemma rtos_rs_Tr1:"\(0::nat) < r; 0 < s \ \ rtos r s (r * s) = r * s" apply (simp add:rtos_def) done lemma rtos_rs_Tr2:"\(0::nat) < r; 0 < s \ \ \l \ (r * s). rtos r s l \ (r * s)" apply (rule allI, rule impI) apply (case_tac "l = r * s", simp add:rtos_rs_Tr1) apply (frule le_imp_less_or_eq) apply (thin_tac "l \ r * s", simp) apply (frule mult_less_mono2[of "0" "s" "r"], simp only:mult.commute, simp only:mult_0_right) apply (frule_tac r = r and s = s and l = l in rtos_hom2, assumption+) apply (rule less_le_diff) apply simp apply (metis le_pre_le) done lemma rtos_rs_Tr3:"\(0::nat) < r; 0 < s \ \ inj_on (rtos r s) {i. i \ (r * s)}" apply (frule rtos_inj [of "r" "s"], assumption+) apply (simp add:inj_on_def [of _ "{i. i \ (r * s)}"]) apply ((rule allI, rule impI)+, rule impI) apply (case_tac "x = r * s") apply (rule contrapos_pp, simp+) apply (frule not_sym) apply (frule mult_less_mono2[of "0" "s" "r"], simp only:mult.commute, simp only:mult_0_right) apply (cut_tac x = y and n = "r * s - Suc 0" in Nset_pre, simp+) apply (frule_tac l = y in rtos_hom1[of "r" "s"], assumption+) apply (simp only: rtos_rs_Tr1) apply (frule sym, thin_tac "r * s = rtos r s y", simp) apply (simp add:mult.commute[of "s" "r"]) apply (simp add: not_less [symmetric, of "r * s" "r * s - Suc 0"]) apply (frule mult_less_mono2[of "0" "s" "r"], simp only:mult.commute, simp only:mult_0_right, cut_tac x = x in Nset_pre[of _ "r * s - Suc 0"], simp+) apply (case_tac "y = r * s") apply (simp add: rtos_rs_Tr1) apply (frule_tac l = x in rtos_hom1[of "r" "s"], assumption+) apply (simp add:mult.commute[of "s" "r"]) apply (simp add: not_less [symmetric, of "r * s" "r * s - Suc 0"]) apply (cut_tac x = y in Nset_pre[of _ "r * s - Suc 0"], simp+) apply (frule rtos_inj[of "r" "s"], assumption+) apply (simp add:inj_on_def) done lemma Qw_cmpser:"\Group G; w_cmpser G (Suc n) f \ \ Gchain n (Qw_cmpser G f)" apply (simp add:Gchain_def) apply (rule allI, rule impI) apply (simp add:Qw_cmpser_def) apply (simp add:w_cmpser_def) apply (erule conjE) apply (simp add:d_gchain_def) apply (cut_tac c = l in subsetD[of "{i. i \ n}" "{i. i \ (Suc n)}"], rule subsetI, simp, simp) apply (frule_tac H = "f l" in Group.Group_Gp[of "G"], frule_tac a = l in forall_spec, simp+) apply (frule_tac G = "Gp G (f l)" and N = "f (Suc l)" in Group.Group_Qg) apply blast apply assumption done definition wcsr_rfns :: "[_ , nat, nat \ 'a set, nat] \ (nat \ 'a set) set" where "wcsr_rfns G r f s = {h. tw_cmpser G (s * r) h \ (\i \ r. h (i * s) = f i)}" (** where f \ tw_cmpser G r **) (** 0-+-+-2-+-+-4-+-+-6 h f 0 f 1 f 2 f 3 **) definition trivial_rfn :: "[_ , nat, nat \ 'a set, nat] \ (nat \ 'a set)" where "trivial_rfn G r f s k == if k < (s * r) then f (k div s) else f r" lemma (in Group) rfn_tool8:"\compseries G r f; 0 < r\ \ d_gchain G r f" apply (simp add:compseries_def tW_cmpser_def) apply (erule conjE)+ apply (simp add:tD_gchain_def) apply (erule conjE)+ apply (simp add:D_gchain_is_d_gchain) done lemma (in Group) rfn_tool16:"\0 < r; 0 < s; i \ (s * r - Suc 0); G \ f (i div s); (Gp G (f (i div s))) \ f (Suc (i div s)); (Gp G (f (i div s))) \ (f (i div s) \ g (s - Suc 0))\ \ (Gp G ((f (Suc (i div s)) \\<^bsub>G\<^esub> (f (i div s) \ g (s - Suc 0))))) \ (f (Suc (i div s)))" apply (rule ZassenhausTr4_1 [of "f (i div s)" "f (Suc (i div s))" "g (s - Suc 0)"], assumption+) done text\Show existence of the trivial refinement. This is not necessary to prove JHS\ lemma rfn_tool30:"\0 < r; 0 < s; l div s * s + s < s * r\ \ Suc (l div s) < r" apply (simp add:mult.commute[of "s" "r"]) apply (cut_tac add_mult_distrib[THEN sym, of "l div s" "s" "1"]) apply (simp only:nat_mult_1) apply (thin_tac "l div s * s + s = (l div s + 1) * s") apply (cut_tac mult_less_cancel2[of "l div s + 1" "s" "r"]) apply simp done lemma (in Group) simple_grouptr0:"\G \ H; G \ K; K \ H; simple_Group (G / K)\ \ H = carrier G \ H = K" apply (simp add:simple_Group_def) apply (frule subg_Qsubg[of "H" "K"], assumption+) apply (rotate_tac -1) apply (frule in_set_with_P[of _ "carrier ((Gp G H) / K)"]) apply simp apply (thin_tac "{N. G / K \ N } = {carrier (G / K), {\\<^bsub>G / K\<^esub>}}") apply (erule disjE) apply (frule sg_subset[of "H"]) apply (frule equalityI[of "H" "carrier G"]) apply (rule subsetI) apply (simp add:Qg_carrier) apply (frule nsg_sg[of "K"]) apply (frule_tac a = x in rcs_in_set_rcs[of "K"], assumption+) apply (frule sym, thin_tac "carrier ((\H)/ K) = set_rcs G K", simp, thin_tac "set_rcs G K = carrier ((\H) / K)") apply (frule Group_Gp[of "H"], simp add:Group.Qg_carrier[of "Gp G H" "K"], simp add:set_rcs_def, erule bexE, simp add:Gp_carrier) apply (simp add:rcs_in_Gp[THEN sym]) apply (frule_tac a = x in a_in_rcs[of "K"], assumption+, simp, thin_tac "K \ x = K \ a", thin_tac "G / K \ {C. \a\H. C = K \ a}", simp add:rcs_def, erule bexE, frule_tac c = h in subsetD[of "K" "H"], assumption+) apply (frule_tac x = h and y = a in sg_mult_closed[of "H"], assumption+, simp) apply simp apply (frule equalityI[THEN sym, of "K" "H"]) apply (rule subsetI) apply (frule Group_Gp[of "H"], simp add:Group.Qg_carrier) apply (simp add:Qg_one) apply (frule nsg_sg[of "K"]) apply (frule_tac a = x in Group.rcs_in_set_rcs[of "Gp G H" "K"]) apply (simp add:sg_sg) apply (simp add:Gp_carrier) apply simp apply (frule_tac a = x in Group.rcs_fixed[of "Gp G H" "K"]) apply (simp add:sg_sg) apply (simp add:Gp_carrier) apply assumption+ apply simp done lemma (in Group) compser_nsg:"\0 < n; compseries G n f; i \ (n - 1)\ \ Gp G (f i) \ (f (Suc i))" apply (simp add:compseries_def tW_cmpser_def) done lemma (in Group) compseriesTr5:"\0 < n; compseries G n f; i \ (n - Suc 0)\ \ (f (Suc i)) \ (f i)" apply (frule compseriesTr4[of "n" "f"]) apply (frule w_cmpser_is_d_gchain[of "n" "f"]) apply (simp add:d_gchain_def, cut_tac n_in_Nsetn[of "n"], frule_tac a = n in forall_spec, simp, thin_tac "\l \ n. G \ f l \ (\l \ (n - Suc 0). f (Suc l) \ f l)", erule conjE, blast) done lemma (in Group) refine_cmpserTr0:"\0 < n; compseries G n f; i \ (n - 1); G \ H; f (Suc i) \ H \ H \ f i\ \ H = f (Suc i) \ H = f i" apply (frule compseriesTr0 [of "n" "f" "i"]) apply (cut_tac lessI[of "n - Suc 0"], simp only:Suc_pred, simp) apply (frule Group_Gp [of "f i"]) apply (erule conjE) apply (frule compseriesTr4[of "n" "f"]) apply (frule w_cmpser_ns[of "n" "f" "i"], simp+) apply (unfold compseries_def, frule conjunct2, fold compseries_def, simp) apply (frule_tac x = i in spec, thin_tac "\i\n - Suc 0. simple_Group ((\f i) / f (Suc i))", simp) apply (frule Group.simple_grouptr0 [of "Gp G (f i)" "H" "f (Suc i)"]) apply (simp add:sg_sg) apply assumption+ apply (simp add:Gp_carrier) apply blast done lemma div_Tr4:"\ (0::nat) < r; 0 < s; j < s * r \ \ j div s * s + s \ r * s" apply (frule div_Tr2 [of "r" "s" "j"], assumption+) apply (frule mult_le_mono [of "j div s" "r - Suc 0" "s" "s"]) apply simp apply (frule add_le_mono [of "j div s * s" "(r - Suc 0) * s" "s" "s"]) apply simp apply (thin_tac "j div s * s \ (r - Suc 0) * s") apply (cut_tac add_mult_distrib[THEN sym, of "r - Suc 0" "s" "1"]) apply (simp only:nat_mult_1) apply simp done lemma (in Group) compseries_is_tW_cmpser:"\0 < r; compseries G r f\ \ tW_cmpser G r f" by (simp add:compseries_def) lemma (in Group) compseries_is_td_gchain:"\0 < r; compseries G r f\ \ td_gchain G r f" apply (frule compseries_is_tW_cmpser, assumption+) apply (simp add:tW_cmpser_def, erule conjE) apply (thin_tac "\l\(r - Suc 0). (\f l) \ (f (Suc l))") apply (simp add:tD_gchain_def, (erule conjE)+) apply (frule D_gchain_is_d_gchain) apply (simp add:td_gchain_def) done lemma (in Group) compseries_is_D_gchain:"\0 < r; compseries G r f\ \ D_gchain G r f" apply (frule compseriesTr1) apply (frule tW_cmpser_is_W_cmpser) apply (rule W_cmpser_is_D_gchain, assumption) done lemma divTr5:"\0 < r; 0 < s; l < (r * s)\ \ l div s * s \ l \ l \ (Suc (l div s)) * s" apply (insert div_mult_mod_eq [THEN sym, of "l" "s"]) apply (rule conjI) apply (insert le_add1 [of "l div s * s" "l mod s"]) apply simp apply (frule mod_less_divisor [of "s" "l"]) apply (frule less_imp_le [of "l mod s" "s"]) apply (insert self_le [of "l div s * s"]) apply (frule add_le_mono [of "l div s * s" "l div s * s" "l mod s" "s"]) apply assumption+ apply (thin_tac "l div s * s \ l div s * s + l mod s") apply (thin_tac "l div s * s \ l div s * s") apply simp done lemma (in Group) rfn_compseries_iMTr1:"\0 < r; 0 < s; compseries G r f; h \ wcsr_rfns G r f s\ \ f ` {i. i \ r} \ h ` {i. i \ (s * r)}" apply (simp add:wcsr_rfns_def) apply (rule subsetI) apply (simp add:image_def) apply (auto del:equalityI) apply (frule_tac i = xa in mult_le_mono [of _ "r" "s" "s"]) apply simp apply (simp add:mult.commute [of "r" "s"]) apply (frule_tac a = xa in forall_spec, assumption, thin_tac "\i\r. h (i * s) = f i") apply (frule sym, thin_tac "h (xa * s) = f xa") apply (cut_tac a = xa in mult_mono[of _ r s s], simp, simp, simp, simp) apply (simp only:mult.commute[of r s], blast) done lemma rfn_compseries_iMTr2:"\0 < r; 0 < s; xa < s * r \ \ xa div s * s \ r * s \ Suc (xa div s) * s \ r * s" apply (frule div_Tr1 [of "r" "s" "xa"], assumption+) apply (simp add:less_imp_le) apply (rule conjI) apply (simp add:mult_le_mono [of "xa div s" "r" "s" "s"]) apply (thin_tac "xa div s \ r") apply (frule div_Tr2[of "r" "s" "xa"], assumption+) apply (thin_tac "xa < s * r") apply (frule le_less_trans [of "xa div s" "r - Suc 0" "r"]) apply simp apply (frule Suc_leI [of "xa div s" "r"]) apply (rule mult_le_mono [of "Suc (xa div s)" "r" "s" "s"], assumption+) apply simp done lemma (in Group) rfn_compseries_iMTr3:"\0 < r; 0 < s; compseries G r f; j \ r; \i \ r. h (i * s) = f i\ \ h (j * s) = f j" apply blast done lemma (in Group) rfn_compseries_iM:"\0 < r; 0 < s; compseries G r f; h \ wcsr_rfns G r f s\ \ card (h `{i. i \ (s * r)}) = r + 1" apply (frule compseries_is_D_gchain, assumption+) apply (frule D_gchain1) apply simp apply (subst card_Collect_le_nat[THEN sym, of "r"]) apply (subst card_image[THEN sym, of "f" "{i. i \ r}"], assumption+) apply (rule card_eq[of "h ` {i. i \ (s * r)}" "f ` {i. i \ r}"]) apply (frule rfn_compseries_iMTr1[of "r" "s" "f" "h"], assumption+) apply (rule equalityI[of "h ` {i. i \ (s * r)}" "f ` {i. i \ r}"]) prefer 2 apply simp apply (rule subsetI, thin_tac "f ` {i. i \ r} \ h ` {i. i \ s * r}") apply (frule_tac b = x and f = h and A = "{i. i \ (s * r)}" in mem_in_image3, erule bexE) apply (simp add:mult.commute[of "s" "r"]) apply (simp add:wcsr_rfns_def, (erule conjE)+) apply (frule rfn_compseries_iMTr3[of "r" "s" "f" "r" "h"], assumption+, simp add:n_in_Nsetn, assumption+, subst image_def, simp) apply (case_tac "a = s * r", simp add:mult.commute[of "s" "r"], cut_tac n_in_Nsetn[of "r"], blast) apply (simp add:mult.commute[of "s" "r"]) apply (frule_tac m = a and n = "r * s" in noteq_le_less, assumption+) apply (frule tw_cmpser_is_w_cmpser, frule w_cmpser_is_d_gchain) apply (frule_tac xa = a in rfn_compseries_iMTr2[of "r" "s"], assumption+) apply (simp add:mult.commute) apply (erule conjE) apply (frule_tac l = a in divTr5[of "r" "s"], assumption+) apply (frule pos_mult_pos[of "r" "s"], assumption+) apply (erule conjE, frule_tac l = "a div s * s" and j = a in d_gchainTr2[of "r * s" "h"], assumption+) apply (frule_tac l = a and j = "Suc (a div s) * s" in d_gchainTr2[of "r * s" "h"], assumption+) apply (frule_tac i = a in rtos_hom0[of "r" "s"], assumption+) apply (rule less_le_diff) apply simp apply (frule_tac x = "a div s" and y = r in less_imp_le, frule_tac a = "a div s" in forall_spec, assumption, frule_tac a = "Suc (a div s)" in forall_spec) apply (cut_tac m = "Suc (a div s)" and k = s and n = r in mult_le_cancel2) apply simp apply (thin_tac "\i \ r. h (i * s) = f i") apply simp apply (cut_tac i = "a div s" and H = "h a"in refine_cmpserTr0[of "r" "f"], simp, assumption+, cut_tac x = "a div s" and n = r in less_le_diff, assumption, simp) apply (simp add:d_gchain_mem_sg[of "r * s" "h"]) apply blast apply (erule disjE) apply (frule_tac m = "a div s" and n = r in Suc_leI, blast) apply (frule_tac x = "a div s" and y = r in less_imp_le, blast) done definition cmp_rfn :: "[_ , nat, nat \ 'a set, nat, nat \ 'a set] \ (nat \ 'a set)" where "cmp_rfn G r f s g = (\i. (if i < s * r then f (Suc (i div s)) \\<^bsub>G\<^esub> (f (i div s) \ g (i mod s)) else {\\<^bsub>G\<^esub>}))" (** refinement of compseries G r f by a compseries G s g **) lemma (in Group) cmp_rfn0:"\0 < r; 0 < s; compseries G r f; compseries G s g; i \ (r - 1); j \ (s - 1)\ \ G \ f (Suc i) \\<^bsub>G\<^esub> ((f i ) \ (g j))" apply (rule ZassenhausTr2_1[of "f i" "f (Suc i)" "g j"], simp, rule compseriesTr0[of "r" "f" "i"], assumption+, frule_tac le_less_trans[of i "r - Suc 0" r], simp+) apply (rule compseriesTr0[of "r" "f" "Suc i"], assumption+, arith) apply(rule compseriesTr0[of "s" "g" "j"], assumption+, simp) apply (frule compseries_is_tW_cmpser[of "r" "f"], assumption+, simp add:tW_cmpser_def) done lemma (in Group) cmp_rfn1:"\0 < r; 0 < s; compseries G r f; compseries G s g\ \ f (Suc 0) \\<^bsub>G\<^esub> ((f 0 ) \ (g 0)) = carrier G" apply (frule compseriesTr2 [of "r" "f"]) apply (frule compseriesTr2 [of "s" "g"]) apply (frule compseriesTr0 [of _ "f" "Suc 0"]) apply simp apply simp apply (rule K_absorb_HK[of "f (Suc 0)" "carrier G"], assumption+, simp add:special_sg_G) apply (rule sg_subset, assumption) done lemma (in Group) cmp_rfn2:"\0 < r; 0 < s; compseries G r f; compseries G s g; l \ (s * r)\ \ G \ cmp_rfn G r f s g l" apply (simp add:cmp_rfn_def) apply (case_tac "l < s * r") apply simp apply (frule_tac i = "l div s" and j = "l mod s" in cmp_rfn0 [of "r" "s"], assumption+) apply (simp add:div_Tr2) apply (frule_tac m = l in mod_less_divisor [of "s"]) apply (frule_tac x = "l mod s" and n = s in less_le_diff) apply simp apply assumption apply simp apply (rule special_sg_e) done lemma (in Group) cmp_rfn3:"\0 < r; 0 < s; compseries G r f; compseries G s g\ \ cmp_rfn G r f s g 0 = carrier G \ cmp_rfn G r f s g (s * r) = {\}" apply (rule conjI) apply (simp add:cmp_rfn_def) apply (rule cmp_rfn1 [of "r" "s" "f" "g"], assumption+) apply (simp add:cmp_rfn_def) done lemma rfn_tool20:"\(0::nat) < m; a = b * m + c; c < m \ \ a mod m = c" apply (simp add:add.commute) done lemma Suci_mod_s_2:"\0 < r; 0 < s; i \ r * s - Suc 0; i mod s < s - Suc 0\ \ (Suc i) mod s = Suc (i mod s)" apply (insert div_mult_mod_eq [of "i" "s", THEN sym]) apply (subgoal_tac "Suc i = i div s * s + Suc (i mod s)") apply (subgoal_tac "Suc i mod s = (i div s * s + Suc (i mod s)) mod s") apply (subgoal_tac "Suc (i mod s) < s") apply (frule_tac m = s and a = "Suc i" and b = "i div s" and c = "Suc (i mod s)" in rfn_tool20, assumption+) apply (subgoal_tac "Suc (i mod s) < Suc (s - Suc 0)") apply simp apply (simp del:Suc_pred) apply simp+ done lemma (in Group) inter_sgsTr1:"\0 < r; 0 < s; compseries G r f; compseries G s g; i < r * s\ \ G \ f (i div s) \ g (s - Suc 0)" apply (rule inter_sgs) apply (rule compseriesTr0, assumption+) apply (frule less_imp_le [of "i" "r * s"]) apply (frule div_Tr1 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute, simp) apply (rule compseriesTr0, simp+) done lemma (in Group) JHS_Tr0_2:"\0 < r; 0 < s; compseries G r f; compseries G s g\ \ \i \ (s * r - Suc 0). Gp G (cmp_rfn G r f s g i) \ cmp_rfn G r f s g (Suc i)" apply (frule compseriesTr4 [of "r" "f"], frule compseriesTr4 [of "s" "g"]) apply (rule allI, rule impI) apply (frule pos_mult_pos [of "s" "r"], assumption+, frule_tac a = i in rfn_tool11 [of "s * r"], assumption+, frule_tac l = i in div_Tr1 [of "r" "s"], assumption+, simp add:less_imp_le) apply (frule_tac x = "i div s" in mem_of_Nset [of _ "r"], thin_tac "i div s \ r", thin_tac "i \ s * r - Suc 0", frule_tac l = i in div_Tr2 [of "r" "s"], assumption+, frule_tac x = "i div s" in mem_of_Nset [of _ "r - Suc 0"], frule_tac a = "i div s" in rfn_tool11 [of "r"], assumption+, frule_tac m = "i div s" in Suc_leI [of _ "r"], frule_tac x = "Suc (i div s)" in mem_of_Nset [of _ "r"], thin_tac "i div s < r", thin_tac "Suc (i div s) \ r") apply (simp add:cmp_rfn_def) apply (case_tac "Suc i < s * r", simp, case_tac "i mod s = s - 1", cut_tac l = i in div_Tr3_1 [of "r" "s"], simp+, cut_tac l = "Suc i" in div_Tr2 [of "r" "s"], simp+, cut_tac x = "Suc i div s" in mem_of_Nset [of _ "r - Suc 0"], simp, cut_tac a = "Suc i div s" in rfn_tool11 [of "r"], simp+, cut_tac x = "Suc i div s" in less_mem_of_Nset [of _ "r"], simp, cut_tac m = "Suc i div s" in Suc_leI [of _ "r"], simp, frule_tac x = "Suc (Suc i div s)" in mem_of_Nset [of _ "r"], frule w_cmpser_is_d_gchain [of"r" "f"], frule_tac rfn_tool12_2 [of "s"], assumption+, thin_tac "i mod s = s - Suc 0", thin_tac "Suc i div s \ {i. i \ r}", thin_tac "Suc (Suc i div s) \ r") apply (simp, cut_tac l = "i div s" and j = "Suc (i div s)" in d_gchainTr2 [of "r" "f"], simp, assumption+, cut_tac x = "i div s" and y = "Suc (i div s)" and z = r in less_trans, simp, assumption, simp add:less_imp_le, simp add:less_imp_le, simp, cut_tac l = "Suc (i div s)" and j = "Suc (Suc (i div s))" in d_gchainTr2 [of "r" "f"], simp+, thin_tac "Suc i div s = Suc (i div s)", thin_tac "Suc i mod s = 0", cut_tac i = "i div s" in compseriesTr0 [of "r" "f"], assumption, simp, cut_tac i = "Suc (i div s)" in compseriesTr0 [of "r" "f"], assumption, simp add:less_imp_le, cut_tac i = "Suc (Suc (i div s))" in compseriesTr0 [of "r" "f"], assumption, simp) apply (subst compseriesTr2 [of "s" "g"], assumption, frule_tac H = "f (i div s)" in sg_subset, frule_tac H = "f (Suc (i div s))" in sg_subset, frule_tac A = "f (Suc (i div s))" in Int_absorb2 [of _ "carrier G"], simp, thin_tac "f (Suc (i div s)) \ carrier G = f (Suc (i div s))", frule_tac H = "f (Suc (Suc (i div s)))" and K = "f (Suc (i div s))" in K_absorb_HK, assumption+, simp, thin_tac "f (Suc (Suc (i div s))) \\<^bsub>G\<^esub> (f (Suc (i div s))) = f (Suc (i div s))") apply (rule rfn_tool16 [of "r" "s" _], simp+, cut_tac x = "Suc i" and y = "s * r" and z = "Suc (s *r)" in less_trans, assumption, simp, thin_tac "Suc i < s * r", simp) apply (rule compser_nsg[of r f], simp+) apply (rule_tac H = "f (i div s) \ g (s - Suc 0)" and K = "f (i div s)" in sg_sg, assumption+, cut_tac i = i in inter_sgsTr1 [of r s f g], simp+, cut_tac x = i and y = "Suc i" and z = "s * r" in less_trans, simp+, cut_tac x = i and y = "Suc i" and z = "r * s" in less_trans, simp+, simp add:mult.commute, assumption+, simp add:Int_lower1) apply (frule_tac m = i in mod_less_divisor [of "s"], frule_tac x = "i mod s" in less_le_diff [of _ "s"], simp, frule_tac m = "i mod s" in noteq_le_less [of _ "s - Suc 0"], assumption+, thin_tac "i mod s \ s - Suc 0", thin_tac "i mod s \ s - Suc 0") apply (frule_tac x = "i mod s" and y = "s - Suc 0" and z = s in less_trans, simp, frule_tac k = "i mod s" in nat_pos2 [of _ s], cut_tac l1 = i in div_Tr3_2 [THEN sym, of "r" "s"], simp+, frule_tac i = "i mod s" in compser_nsg [of "s" "g"], assumption+, simp, frule_tac i = "Suc (i div s)" in compseriesTr0 [of "r" "f"], assumption+, frule_tac m = "i mod s" in Suc_mono [of _ "s - Suc 0"], simp only:Suc_pred) apply (frule_tac x = "Suc (i mod s)" in less_mem_of_Nset [of _ "s"], cut_tac i = "Suc (i mod s)" in compseriesTr0 [of "s" "g"], simp+, cut_tac H = "f (i div s)" and ?H1.0 = "f (Suc (i div s))" and K = "g (i mod s)" and ?K1.0 = "g (Suc (i mod s))" in ZassenhausTr3, rule_tac i = "i div s" in compseriesTr0 [of "r" "f"], assumption+, frule_tac x = "i div s" and y = "r - Suc 0" and z = r in le_less_trans, simp, simp, assumption+, frule_tac i = "i mod s" in compseriesTr0 [of "s" "g"], simp, simp, simp) apply (rule_tac i = "i div s" in compser_nsg[of r f], simp+, cut_tac i = i in Suci_mod_s_2[of r s], simp+, cut_tac x = "Suc i" and y = "s * r" and z = "Suc (s *r)" in less_trans, assumption, simp, thin_tac "Suc i < s * r", simp, frule_tac x = i and n = "s * r" in less_le_diff, simp add:mult.commute[of s r], simp+) apply (cut_tac a = "i mod s" in rfn_tool11 [of "s"], simp+, frule_tac m = i in mod_less_divisor [of "s"], frule_tac x = "i mod s" in less_le_diff [of _ "s"], simp, rule special_nsg_e, cut_tac i = "i div s" and j = "i mod s" in cmp_rfn0[of r s f g], simp+) done lemma (in Group) cmp_rfn4:"\0 < r; 0 < s; compseries G r f; compseries G s g; l \ (s * r - Suc 0)\ \ cmp_rfn G r f s g (Suc l) \ cmp_rfn G r f s g l" apply (frule JHS_Tr0_2 [of "r" "s" "f" "g"], assumption+) apply (frule_tac a = l in forall_spec, simp, thin_tac "\i \ (s * r - Suc 0). (\(cmp_rfn G r f s g i)) \ (cmp_rfn G r f s g (Suc i))") apply (frule cmp_rfn2 [of "r" "s" "f" "g" "l"], assumption+, frule le_less_trans [of "l" "s * r - Suc 0" "s* r"], simp, simp add:less_imp_le) apply (frule Group_Gp [of "cmp_rfn G r f s g l"], frule Group.nsg_subset [of "Gp G (cmp_rfn G r f s g l)" ], assumption+, simp add:Gp_carrier) done lemma (in Group) cmp_rfn5:"\0 < r; 0 < s; compseries G r f; compseries G s g\ \ \i \ r. cmp_rfn G r f s g (i * s) = f i" apply (rule allI, rule impI) apply (simp add:cmp_rfn_def) apply (case_tac "i < r", simp, frule_tac x = i in less_imp_le [of _ "r"], frule_tac x = i in mem_of_Nset[of _ "r"], frule_tac i = i in compseriesTr0 [of "r" "f"], assumption+, thin_tac "i \ r", frule_tac m = i in Suc_leI [of _ "r"], frule_tac x = "Suc i" in mem_of_Nset[of _ "r"], frule_tac i = "Suc i" in compseriesTr0 [of "r" "f"], assumption+) apply (subst compseriesTr2 [of "s" "g"], assumption+, frule_tac H = "f i" in sg_subset, subst Int_absorb2, assumption+, frule rfn_tool8, assumption+, simp) apply (cut_tac n = i in zero_less_Suc, frule_tac x = 0 and y = "Suc i" and z = r in less_le_trans, assumption+, frule_tac l = i and j = "Suc i" in d_gchainTr2 [of "r" "f"], frule compseries_is_D_gchain[of "r" "f"], assumption, rule D_gchain_is_d_gchain, assumption+, cut_tac x = i and y = "Suc i" and z = r in less_le_trans, simp+, rule_tac K = "f i" and H = "f (Suc i)" in K_absorb_HK, assumption+, simp, frule compseries_is_td_gchain, assumption+, simp add:td_gchain_def) done lemma (in Group) JHS_Tr0:"\(0::nat) < r; 0 < s; compseries G r f; compseries G s g\ \ cmp_rfn G r f s g \ wcsr_rfns G r f s" apply (simp add:wcsr_rfns_def, frule cmp_rfn5 [of "r" "s" "f" "g"], assumption+, simp, thin_tac "\i\r. cmp_rfn G r f s g (i * s) = f i") apply (simp add:tw_cmpser_def, rule conjI) prefer 2 apply (rule JHS_Tr0_2 [of "r" "s" "f" "g"], assumption+) apply (simp add:td_gchain_def, rule conjI) prefer 2 apply (rule cmp_rfn3, assumption+) apply (simp add:d_gchain_def, rule allI, rule impI, rule conjI, rule cmp_rfn2, assumption+, rule allI, rule impI, rule cmp_rfn4, assumption+) done lemma rfn_tool17:"(a::nat) = b \ a - c = b - c" by simp lemma isom4b:"\Group G; G \ N; G \ H\ \ (Gp G (N \\<^bsub>G\<^esub> H) / N) \ (Gp G H / (H \ N))" apply (frule isom4 [of "G" "N" "H"], assumption+) apply (rule isomTr1) apply (simp add:Group.homom4_2) apply (frule Group.smult_sg_nsg[of "G" "H" "N"], assumption+) apply (simp add:Group.smult_commute_sg_nsg[THEN sym, of "G" "H" "N"]) apply (rule Group.homom4Tr1, assumption+) done lemma Suc_rtos_div_r_1:"\ 0 < r; 0 < s; i \ r * s - Suc 0; Suc (rtos r s i) < r * s; i mod s = s - Suc 0; i div s < r - Suc 0\ \ Suc (rtos r s i) div r = i mod s" apply (simp add:rtos_def, frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply simp apply (subgoal_tac "Suc ((s - Suc 0) * r + i div s) div r = ((s - Suc 0) * r + Suc (i div s)) div r", simp del: add_Suc add_Suc_right) apply (subgoal_tac "Suc (i div s) < Suc (r - Suc 0)") apply simp_all done lemma Suc_rtos_mod_r_1:"\0 < r; 0 < s; i \ r * s - Suc 0; Suc (rtos r s i) < r * s; i mod s = s - Suc 0; i div s < r - Suc 0\ \ Suc (rtos r s i) mod r = Suc (i div s)" apply (simp add:rtos_def) done lemma i_div_s_less:"\0 < r; 0 < s; i \ r * s - Suc 0; Suc (rtos r s i) < r * s; i mod s = s - Suc 0; Suc i < s * r \ \ i div s < r - Suc 0" apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply (frule_tac r = r and s = s and l = i in div_Tr2, assumption+) apply (simp add:mult.commute) apply (rule contrapos_pp, simp+, subgoal_tac "i div s = r - Suc 0", thin_tac "i div s = r - Suc 0") apply (simp add:rtos_def, subgoal_tac "(s - Suc 0) * r + r = r * s", simp) apply (thin_tac "(s - Suc 0) * r + r < r * s") apply (simp add:mult.commute, simp add:diff_mult_distrib2) apply simp done lemma rtos_mod_r_1:"\ 0 < r; 0 < s; i \ r * s - Suc 0; rtos r s i < r * s; i mod s = s - Suc 0 \ \ rtos r s i mod r = i div s" apply (simp add:rtos_def) apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply simp apply (rule rfn_tool20, assumption+, simp) apply (frule_tac r = r and s = s and l = i in div_Tr2, assumption+) apply (rule le_less_trans, assumption+, simp add:mult.commute) apply (rule le_less_trans, assumption+, simp) done lemma Suc_i_mod_s_0_1:"\0 < r; 0 < s; i \ r * s - Suc 0; i mod s = s - Suc 0\ \ Suc i mod s = 0" apply (insert div_mult_mod_eq [of "i" "s", THEN sym]) apply simp apply (thin_tac "i mod s = s - Suc 0") apply (subgoal_tac "Suc i mod s = Suc (i div s * s + s - Suc 0) mod s") apply (thin_tac "i = i div s * s + s - Suc 0", simp del: add_Suc) apply (subgoal_tac "Suc (i div s * s + s - Suc 0) mod s = (i div s * s + s) mod s") apply (simp del: add_Suc) apply (subgoal_tac "Suc (i div s * s + s - Suc 0) = i div s * s + s") apply (simp del: add_Suc) apply (rule Suc_pred [of "i div s * s + s"], simp) done (* same as div_Tr3_2 *) lemma Suci_div_s_2:"\0 < r; 0 < s; i \ r * s - Suc 0; i mod s < s - Suc 0\ \ Suc i div s = i div s" apply (rule div_Tr3_2 [THEN sym], assumption+) apply simp done lemma rtos_i_mod_r_2:"\0 < r; 0 < s; i \ r * s - Suc 0\ \ rtos r s i mod r = i div s" apply (simp add:rtos_def) apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply simp apply (frule_tac r = r and s = s and l = i in div_Tr2, assumption+) apply (simp add:mult.commute) apply (subgoal_tac "i div s < r") apply (rule rfn_tool20, assumption+, simp) apply assumption apply (rule le_less_trans, assumption+, simp) done lemma Suc_rtos_i_mod_r_2:"\0 < r; 0 < s; i \ r * s - Suc 0; i div s = r - Suc 0\ \ Suc (rtos r s i) mod r = 0" apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply (simp add:rtos_def) done lemma Suc_rtos_i_mod_r_3:"\0 < r; 0 < s; i \ r * s - Suc 0; i div s < r - Suc 0\ \ Suc (rtos r s i) mod r = Suc (i div s)" apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply (simp add:rtos_def) done lemma Suc_rtos_div_r_3:"\0 < r; 0 < s; i mod s < s - Suc 0; i \ r * s - Suc 0; Suc (rtos r s i) < r * s; i div s < r - Suc 0\ \ Suc (rtos r s i) div r = i mod s" apply (simp add:rtos_def) apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"]) apply simp apply simp apply (subst add_Suc_right[THEN sym, of "i mod s * r" "i div s"]) apply (subst add.commute[of "i mod s * r" "Suc (i div s)"]) apply (frule Suc_leI[of "i div s" "r - Suc 0"]) apply (frule le_less_trans[of "Suc (i div s)" "r - Suc 0" "r"], simp) apply (subst div_mult_self1 [of "r" "Suc (i div s)" "i mod s"]) apply auto done lemma r_s_div_s:"\0 < r; 0 < s\ \ (r * s - Suc 0) div s = r - Suc 0" apply (frule rfn_tool1_1 [of "s" "s - Suc 0" "r - Suc 0"]) apply simp apply (simp add:diff_mult_distrib) done lemma r_s_mod_s:"\0 < r; 0 < s\ \ (r * s - Suc 0) mod s = s - Suc 0" apply (frule rfn_tool20 [of "s" "(r - Suc 0) * s + (s - Suc 0)" "r - Suc 0" "s - Suc 0"], simp, simp) apply (simp add:diff_mult_distrib) done lemma rtos_r_s:"\0 < r; 0 < s\ \ rtos r s (r * s - Suc 0) = r * s - Suc 0" apply (simp add:rtos_def) apply (frule r_s_div_s [of "r" "s"], assumption+) apply (frule r_s_mod_s [of "r" "s"], assumption+) apply (simp, simp add:diff_mult_distrib, simp add:mult.commute) done lemma rtos_rs_1:"\ 0 < r; 0 < s; rtos r s i < r * s; \ Suc (rtos r s i) < r * s\ \ rtos r s i = r * s - Suc 0" apply (frule pos_mult_pos [of "r" "s"], assumption+) apply (frule less_le_diff [of "rtos r s i" "r * s"]) apply (simp add:nat_not_le_less[THEN sym, of "Suc (rtos r s i)" "r * s"]) done lemma rtos_rs_i_rs:"\ 0 < r; 0 < s; i \ r * s - Suc 0; rtos r s i = r * s - Suc 0\ \ i = r * s - Suc 0" apply (frule mem_of_Nset [of "i" "r * s - Suc 0"]) apply (frule rtos_r_s [of "r" "s"], assumption+) apply (frule rtos_inj[of "r" "s"], assumption+) apply (cut_tac n_in_Nsetn[of "r * s - Suc 0"]) apply (simp add:inj_on_def) done lemma JHS_Tr1_1:"\Group G; 0 < r; 0 < s; compseries G r f; compseries G s g\ \ f (Suc ((r * s - Suc 0) div s)) \\<^bsub>G\<^esub> (f ((r * s - Suc 0) div s) \ g ((r * s - Suc 0) mod s)) = f (r - Suc 0) \ g (s - Suc 0)" apply (frule r_s_div_s [of "r" "s"], assumption+) apply (frule r_s_mod_s [of "r" "s"], assumption+) apply simp apply (subst Group.compseriesTr3 [of "G" "r" "f"], assumption+) apply (frule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+) apply (simp add:less_mem_of_Nset) apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+) apply (simp add:less_mem_of_Nset) apply (frule Group.inter_sgs [of "G" "f (r - (Suc 0))" "g (s - Suc 0)"], assumption+) apply (rule Group.s_top_l_unit, assumption+) done lemma JHS_Tr1_2:"\Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; k < r - Suc 0\ \ ((Gp G (f (Suc k) \\<^bsub>G\<^esub> (f k \ g (s - Suc 0)))) / (f (Suc (Suc k)) \\<^bsub>G\<^esub> (f (Suc k) \ g 0))) \ ((Gp G (g s \\<^bsub>G\<^esub> (g (s - Suc 0) \ f k))) / (g s \\<^bsub>G\<^esub> (g (s - Suc 0) \ f (Suc k))))" apply (frule Group.compseriesTr0 [of "G" "r" "f" "k"], assumption+) apply (frule less_trans [of "k" "r - Suc 0" "r"]) apply simp apply simp apply (frule Suc_leI [of "k" "r - Suc 0"]) apply (frule le_less_trans [of "Suc k" "r - Suc 0" "r"]) apply simp apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc k"], assumption+) apply simp apply (thin_tac "Suc k \ r - Suc 0") apply (frule Suc_leI[of "Suc k" "r"]) apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc (Suc k)"], assumption+) apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+) apply simp apply (subst Group.compseriesTr2 [of "G" "s" "g"], assumption+) apply (subst Group.compseriesTr3 [of "G" "s" "g"], assumption+) apply (subst Int_absorb2 [of "f (Suc k)" "carrier G"]) apply (rule Group.sg_subset, assumption+) apply (subst Group.K_absorb_HK[of "G" "f (Suc (Suc k))" "f (Suc k)"], assumption+) apply (rule Group.compseriesTr5 [of "G" "r" "f" "Suc k"], assumption+) apply (frule Suc_leI [of "k" "r - Suc 0"]) apply simp apply (subst Group.s_top_l_unit[of "G" "g (s - Suc 0) \ f k"], assumption+) apply (simp add:Group.inter_sgs) apply (subst Group.compseriesTr3[of "G" "s" "g"], assumption+) apply (subst Group.s_top_l_unit[of "G" "g (s - Suc 0) \ f (Suc k)"], assumption+) apply (simp add:Group.inter_sgs) apply (frule Group.Group_Gp [of "G" "f k"], assumption+) apply (frule Group.compser_nsg [of "G" "r" "f" "k"], assumption+) apply (simp add:less_mem_of_Nset [of "k" "r - Suc 0"]) apply (frule isom4b [of "Gp G (f k)" "f (Suc k)" "f k \ g (s - Suc 0)"], assumption+) apply (rule Group.sg_sg, assumption+) apply (rule Group.inter_sgs, assumption+) apply (simp add:Int_lower1) apply (frule Group.compseriesTr5[of "G" "r" "f" "k"], assumption+) apply simp apply (frule Group.s_top_induced[of "G" "f k" "f (Suc k)" "f k \ g (s - Suc 0)"], assumption+) apply (simp add:Int_lower1) apply simp apply (thin_tac "f (Suc k) \\<^bsub>\\<^bsub>G\<^esub>(f k)\<^esub> (f k \ g (s - Suc 0)) = f (Suc k) \\<^bsub>G\<^esub> (f k \ g (s - Suc 0))") apply (frule Suc_pos [of "Suc k" "r"]) apply (frule Group.cmp_rfn0 [of "G" "r" "s" "f" "g" "k" "s - Suc 0"], assumption+) apply simp apply simp apply (frule Group.sg_inc_set_mult[of "G" "f k" "f (Suc k)" "f k \ g (s - Suc 0)"], assumption+) apply (simp add:Int_lower1) apply (simp add:Group.Gp_inherited [of "G" "f (Suc k) \\<^bsub>G\<^esub> (f k \ g (s - Suc 0))" "f k"]) apply (frule Group.inter_sgs [of "G" "f k" "g (s - Suc 0)"], assumption+) apply (frule Group.Gp_inherited [of "G" "f k \ g (s - Suc 0)" "f k"], assumption+) apply (rule Int_lower1) apply simp apply (thin_tac "(Gp (Gp G (f k)) ((f k) \ (g (s - Suc 0)))) = (Gp G ((f k) \ (g (s - Suc 0))))") apply (thin_tac "f (Suc k) \\<^bsub>G\<^esub> f k \ g (s - Suc 0) \ f k") apply (thin_tac "G \ f (Suc k) \\<^bsub>G\<^esub> f k \ g (s - Suc 0)") apply (simp add:Int_assoc [of "f k" "g (s - Suc 0)" "f (Suc k)"]) apply (simp add:Int_commute [of "g (s - Suc 0)" "f (Suc k)"]) apply (simp add:Int_assoc [of "f k" "f (Suc k)" "g (s - Suc 0)", THEN sym]) apply (simp add:Int_absorb1[of "f (Suc k)" "f k"]) apply (simp add:Int_commute) done lemma JHS_Tr1_3:"\Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i \ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0; Suc i div s \ r - Suc 0; i div s = r - Suc 0\ \ Group (Gp G (f r \\<^bsub>G\<^esub> (f (r - Suc 0) \ g (i mod s))) / (f r \\<^bsub>G\<^esub> (f (r - Suc 0) \ g (Suc (i mod s)))))" apply (frule nat_eq_le[of "i div s" "r - Suc 0"]) apply (frule Group.compser_nsg [of "G" "r" "f" "i div s"], assumption+) apply simp apply (frule Group.compser_nsg [of "G" "s" "g" "i mod s"], assumption+) apply simp apply (frule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+) apply simp apply (frule Group.compseriesTr0 [of "G" "r" "f" "r"], assumption+) apply simp apply (frule Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+) apply (simp add:less_imp_le) apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+) apply (frule Suc_mono [of "i mod s" "s - (Suc 0)"], simp add:less_mem_of_Nset) apply (frule Group.cmp_rfn0 [of "G" "r" "s" "f" "g" "i div s" "i mod s"], assumption+, simp, simp) apply (frule Group.ZassenhausTr3 [of "G" "f (r - Suc 0)" "f r" "g (i mod s)" "g (Suc (i mod s))"], assumption+, simp, simp) apply (cut_tac Group.Group_Gp [of "G" "f r \\<^bsub>G\<^esub> (f (r - Suc 0) \ g (i mod s))"]) apply (rule Group.Group_Qg, assumption+) apply simp done lemma JHS_Tr1_4:"\Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i \ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0; Suc i div s \ r - Suc 0; i div s = r - Suc 0\ \ Group (Gp G (g (Suc (i mod s)) \\<^bsub>G\<^esub> (g (i mod s) \ f (r - Suc 0))) / (g (Suc (Suc (i mod s))) \\<^bsub>G\<^esub> (g (Suc (i mod s)) \ f 0)))" apply (subst Group.compseriesTr2 [of "G" "r" "f"], assumption+) apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+) apply (frule Suc_mono [of "i mod s" "s - Suc 0"], simp) apply (frule Group.sg_subset [of "G" "g (Suc (i mod s))"], assumption+) apply (subst Int_absorb2, assumption+) apply (frule Suc_mono [of "i mod s" "s - Suc 0"]) apply (frule less_mem_of_Nset [of "i mod s" "s - Suc 0"], simp) apply (frule Suc_leI [of "Suc (i mod s)" "s"]) apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (Suc (i mod s))"], assumption+) apply (frule less_le_diff [of "Suc (i mod s)" "s"]) apply (frule Suc_pos [of "Suc (i mod s)" "s"]) apply (frule Group.compseriesTr5[of "G" "s" "g" "Suc (i mod s)"], assumption+) apply (subst Group.K_absorb_HK[of "G" "g (Suc (Suc (i mod s)))" "g (Suc (i mod s))"], assumption+) apply (frule Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+) apply (frule mod_less_divisor [of "s" "i"], simp) apply (frule Group.cmp_rfn0 [of "G" "s" "r" "g" "f" "i mod s" "r - Suc 0"], assumption+, simp, simp, simp) apply (cut_tac Group.compser_nsg [of "G" "s" "g" "i mod s"]) apply (frule Group.ZassenhausTr4_1 [of "G" "g (i mod s)" "g (Suc (i mod s))" "f (r - Suc 0)"], assumption+) apply (rule Group.sg_sg, assumption+) apply (rule Group.inter_sgs, assumption+) apply (rule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+) apply simp apply (rule Int_lower1) apply (rule Group.Group_Qg) apply (rule Group.Group_Gp, assumption+, simp+) done lemma JHS_Tr1_5:"\Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i \ s * r - Suc 0; Suc (rtos r s i) < s * r; Suc i < s * r; i mod s < s - Suc 0; i div s < r - Suc 0\ \ (Gp G (f (Suc (i div s)) \\<^bsub>G\<^esub> (f (i div s) \ g (i mod s))) / (f (Suc (i div s)) \\<^bsub>G\<^esub> (f (i div s) \ g (Suc (i mod s))))) \ (Gp G (g (Suc (i mod s)) \\<^bsub>G\<^esub> (g (i mod s) \ f (i div s))) / (g (Suc (Suc (rtos r s i) div r)) \\<^bsub>G\<^esub> (g (Suc (rtos r s i) div r) \ f (Suc (rtos r s i) mod r))))" apply (frule Group.compseriesTr0 [of "G" "s" "g" "i mod s"], assumption+) apply simp apply (frule Group.compseriesTr0 [of "G" "s" "g" "Suc (i mod s)"], assumption+) apply (frule Suc_mono [of "i mod s" "s - Suc 0"], simp) apply (frule Group.compseriesTr0 [of "G" "r" "f" "i div s"], assumption+) apply (frule less_trans [of "i div s" "r - Suc 0" "r"], simp) apply simp apply (frule Group.compseriesTr0 [of "G" "r" "f" "Suc (i div s)"], assumption+) apply (frule Suc_mono [of "i div s" "r - Suc 0"]) apply simp apply (frule Group.compser_nsg [of "G" "r" "f" "i div s"], assumption+) apply simp apply (frule Group.compser_nsg [of "G" "s" "g" "i mod s"], assumption+, simp) apply (subst Suc_rtos_i_mod_r_3 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute, assumption) apply (subst Suc_rtos_div_r_3 [of "r" "s" "i" ], assumption+)+ apply (simp add:mult.commute)+ apply (subst Suc_rtos_div_r_3[of "r" "s" "i"], assumption+) apply (rule Zassenhaus [of "G" "f (i div s)" "f (Suc (i div s))" "g (i mod s)" "g (Suc (i mod s))"], assumption+) done lemma JHS_Tr1_6:" \Group G; 0 < r; 0 < s; compseries G r f; compseries G s g; i \ r * s - Suc 0; Suc (rtos r s i) < r * s\ \ ((Gp G (cmp_rfn G r f s g i)) / (cmp_rfn G r f s g (Suc i))) \ ((Gp G (g (Suc (rtos r s i div r)) \\<^bsub>G\<^esub> (g (rtos r s i div r) \ f (rtos r s i mod r)))) / (g (Suc (Suc (rtos r s i) div r)) \\<^bsub>G\<^esub> (g (Suc (rtos r s i) div r) \ f (Suc (rtos r s i) mod r))))" apply (simp add:cmp_rfn_def) apply (frule le_less_trans [of "i" "r * s - Suc 0" "r * s"], simp) apply (simp add:mult.commute [of "r" "s"]) apply (frule Suc_leI [of "i" "s * r"], thin_tac "i < s * r") apply (case_tac "\ Suc i < s * r", simp) apply (frule rfn_tool17 [of "Suc i" "s * r" "Suc 0"]) apply (thin_tac " Suc i = s * r") apply simp apply (frule rtos_r_s [of "r" "s"], assumption+) apply (simp add:mult.commute [of "r" "s"]) (* !! ??? *) apply simp apply (frule mod_less_divisor [of "s" "i"]) apply (frule less_le_diff [of "i mod s" "s"], thin_tac "i mod s < s") apply (case_tac "i mod s = s - Suc 0", simp) apply (frule_tac div_Tr2 [of "r" "s" "Suc i"], assumption+) apply (simp add:le_imp_less_or_eq) apply (subst div_Tr3_1[of "r" "s" "i"], assumption+, simp) apply (subst rtos_hom3 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply (subst rtos_hom3_1 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply (frule div_Tr3_1 [of "r" "s" "i"], assumption+, simp) apply (simp, thin_tac "Suc i div s = Suc (i div s)") apply (insert n_less_Suc [of "i div s"]) apply (frule less_le_trans [of "i div s" "Suc (i div s)" "r - Suc 0"], assumption+) apply (subst Suc_rtos_div_r_1 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute[of "s" "r"])+ apply (subst Suc_rtos_mod_r_1 [of "r" "s" "i"], assumption+) apply (subst Suc_i_mod_s_0_1 [of "r" "s" "i"], assumption+) apply (simp only:Suc_rtos_div_r_1 [of "r" "s" "i"]) apply (subst rtos_hom3[of "r" "s" "i"], assumption+, simp) apply (frule JHS_Tr1_2 [of "G" "r" "s" "f" "g" "i div s"], assumption+, simp) apply (frule noteq_le_less [of "i mod s" "s - Suc 0"], assumption+) apply (thin_tac "i mod s \ s - Suc 0") apply (thin_tac "i mod s \ s - Suc 0") apply (frule div_Tr2 [of "r" "s" "Suc i"], assumption+, rule noteq_le_less, assumption+) apply (subst div_Tr3_2 [THEN sym, of "r" "s" "i"], assumption+) apply simp apply (subst rfn_tool12_1 [THEN sym, of "s" "i"], assumption+) apply simp apply (subst rtos_hom3 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply (subst rtos_hom3_1 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply (case_tac "i div s = r - Suc 0") apply (subst rtos_hom5 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply assumption apply (subst rtos_hom7 [of "r" "s" "i"], assumption+) apply (simp add:mult.commute) apply (assumption, simp) apply (frule JHS_Tr1_3 [of "G" "r" "s" "f" "g" "i"], assumption+, simp only:noteq_le_less, assumption+) apply (frule JHS_Tr1_4 [of "G" "r" "s" "f" "g" "i"], assumption+, simp only:noteq_le_less, assumption+) apply (subst rtos_hom3 [of "r" "s" "i"], assumption+, simp only:mult.commute[of "s" "r"]) apply (subst rtos_hom5 [of "r" "s" "i"], assumption+, simp only:mult.commute[of "s" "r"], simp add:mem_of_Nset) apply (rule isomTr1, assumption+) apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+, simp only:mult.commute, assumption) apply simp apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+, simp only:mult.commute, assumption, simp) apply (rule JHS_Tr1_2 [of "G" "s" "r" "g" "f" "i mod s"], assumption+) apply (frule div_Tr2 [of "r" "s" "i"], assumption+) apply (rule le_less_trans [of "i" "s * r - Suc 0" "s * r"], assumption+) apply simp apply (frule noteq_le_less [of "i div s" "r - Suc 0"], assumption+) apply (frule Suci_div_s_2[of "r" "s" "i"], assumption+, simp only:mult.commute[of "s" "r"], assumption, simp) apply (frule JHS_Tr1_5[of "G" "r" "s" "f" "g" "i"], assumption+) apply (simp add:noteq_le_less[of "Suc i"], assumption+) apply (frule mem_of_Nset[of "i" "s*r - Suc 0"], simp add:mult.commute[of "s" "r"]) apply (simp add:rtos_hom3 [of "r" "s" "i"]) done lemma JHS_Tr1:"\ Group G; 0 < r; 0 < s; compseries G r f; compseries G s g\ \ isom_Gchains (r * s - 1) (rtos r s) (Qw_cmpser G (cmp_rfn G r f s g)) (Qw_cmpser G (cmp_rfn G s g r f))" apply (simp add:isom_Gchains_def) apply (rule allI, rule impI) apply (frule pos_mult_pos [of "r" "s"], assumption+) apply (frule_tac b = "r * s" and a = i in rfn_tool11, assumption+) apply (simp add:Qw_cmpser_def) apply (simp only:cmp_rfn_def [of "G" "s" "g"]) apply (frule_tac l = i in rtos_hom1 [of "r" "s"], assumption+) apply (frule_tac x = "rtos r s i" and y = "s * r - Suc 0" and z = "s * r" in le_less_trans, simp) apply (simp add:mult.commute [of "s" "r"]) apply (case_tac "Suc (rtos r s i) < r * s", simp) prefer 2 apply simp apply (frule_tac i = i in rtos_rs_1 [of "r" "s"], assumption+) apply (frule_tac i = i in rtos_rs_i_rs [of "r" "s"], assumption+) apply (rule less_le_diff, assumption+) apply (simp add:cmp_rfn_def) apply (simp add:mult.commute) apply (subst JHS_Tr1_1 [of "G" "r" "s" "f" "g"], assumption+) apply (frule JHS_Tr1_1 [of "G" "s" "r" "g" "f"], assumption+) apply (simp add:mult.commute [of "r" "s"]) apply (simp add:Int_commute) apply (frule Group.compseriesTr0 [of "G" "r" "f" "r - Suc 0"], assumption+, simp) apply (frule Group.compseriesTr0 [of "G" "s" "g" "s - Suc 0"], assumption+, simp) apply (frule Group.inter_sgs [of "G" "f (r - (Suc 0))" "g (s - Suc 0)"], assumption+) apply (frule Group.special_sg_e [of "G"]) apply (frule Group.special_nsg_e [of "G" "f (r - Suc 0) \ g (s - Suc 0)"], assumption+) apply (frule Group.Group_Gp [of "G" "f (r - Suc 0) \ g (s - Suc 0)"], assumption+) apply (frule Group.Group_Qg[of "Gp G (f (r - Suc 0) \ g (s - Suc 0))" "{\\<^bsub>G\<^esub>}"], assumption+) apply (simp add:isomTr0[of "(\\<^bsub>G\<^esub>(f (r - Suc 0) \ g (s - Suc 0))) / {\\<^bsub>G\<^esub>}"]) apply (rule JHS_Tr1_6, assumption+) done lemma abc_SucTr0:"\(0::nat) < a; c \ b; a - Suc 0 = b - c\ \ a = (Suc b) - c" apply (subgoal_tac "Suc 0 \ a") apply (frule le_add_diff_inverse2 [of "Suc 0" "a", THEN sym]) apply auto done lemma length_wcmpser0_0:"\Group G; Ugp E; w_cmpser G (Suc 0) f\ \ f ` {i. i \ (Suc 0)} = {f 0, f (Suc 0)}" apply (simp add:Nset_1) done lemma length_wcmpser0_1:"\Group G; Ugp E; w_cmpser G (Suc n) f; i\{i. i \ n}; (Qw_cmpser G f) i \ E\ \ f i = f (Suc i)" apply (subgoal_tac "0 < Suc n") apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "i"], assumption+) prefer 2 apply simp apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "Suc i"], simp+) apply (frule Group.w_cmpser_ns [of "G" "Suc n" "f" "i"], simp+) apply (simp add:Qw_cmpser_def) apply (rule QgrpUnit_3 [of "G" "E" "f i" "f (Suc i)"], assumption+, simp+) done lemma length_wcmpser0_2:"\Group G; Ugp E; w_cmpser G (Suc n) f; i \ n; \ (Qw_cmpser G f) i \ E\ \ f i \ f (Suc i)" apply (cut_tac zero_less_Suc[of "n"]) apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "i"], assumption+) apply simp apply (frule Group.w_cmpser_gr [of "G" "Suc n" "f" "Suc i"], assumption+) apply simp apply (frule Group.w_cmpser_ns [of "G" "Suc n" "f" "i"], assumption+, simp) apply (simp add:Qw_cmpser_def) apply (rule QgrpUnit_4 [of "G" "E" "f i" "f (Suc i)"], assumption+) done lemma length_wcmpser0_3:"\Group G; Ugp E; w_cmpser G (Suc (Suc n)) f; f (Suc n) \ f (Suc (Suc n))\ \ f (Suc (Suc n)) \ f ` {i. i \ (Suc n)}" apply (frule Group.w_cmpser_is_d_gchain, assumption+) apply (thin_tac "w_cmpser G (Suc (Suc n)) f") apply (rule contrapos_pp, simp+) apply (frule Group.d_gchainTr2 [of "G" "Suc ((Suc n))" "f" "Suc n" "Suc (Suc n)"]) apply simp apply assumption+ apply simp+ apply (frule psubsetI [of "f (Suc (Suc n))" "f (Suc n)"]) apply (rule not_sym, assumption+) apply (thin_tac "f (Suc (Suc n)) \ f (Suc n)") apply (simp add:image_def) apply (erule exE) apply (cut_tac zero_less_Suc[of "Suc n"]) apply (frule_tac f = f and l = x in Group.d_gchainTr2 [of "G" "Suc ((Suc n))" _ _ "Suc (Suc n)"], assumption+) apply simp+ apply (frule_tac f = f and l = x in Group.d_gchainTr2 [of "G" "Suc (Suc n)" _ _ "Suc n"], simp+) done lemma length_wcmpser0_4:"\Group G; Ugp E; w_cmpser G (Suc 0) f\ \ card (f ` {i. i \ Suc 0}) - 1 = Suc 0 - card {i. i = 0 \ Qw_cmpser G f i \ E}" (* card (f ` Nset (Suc 0)) - 1 = Suc 0 - card {i. i \ Nset 0 \ Qw_cmpser G f i \ E}" *) apply (auto simp add: length_wcmpser0_0 Collect_conv_if) apply (frule_tac n = 0 and f = f and i = 0 in length_wcmpser0_1 [of "G" "E"], assumption+, simp+) apply (frule_tac f = f and i = 0 in length_wcmpser0_2 [of "G" "E" "0"], (assumption | simp)+) done lemma length_wcmpser0_5:" \Group G; Ugp E; w_cmpser G (Suc (Suc n)) f; w_cmpser G (Suc n) f; card (f ` {i. i \ (Suc n)}) - 1 = Suc n - card {i. i \ n \ Qw_cmpser G f i \ E}; Qw_cmpser G f (Suc n) \ E\ \ card (f ` {i . i \ (Suc (Suc n))}) - 1 = Suc (Suc n) - card {i. i \ (Suc n) \ Qw_cmpser G f i \ E}" apply (frule_tac n = "Suc n" and f = f and i = "Suc n" in length_wcmpser0_1 [of "G" "E"], assumption+) apply (simp, assumption) apply (subgoal_tac "f ` {i. i \ (Suc (Suc n))} = f ` {i. i \ (Suc n)}") apply simp prefer 2 apply (rule equalityI) apply (simp add:image_def) apply (auto del:equalityI) apply (case_tac "xa = Suc (Suc n)", simp) apply (thin_tac " xa = Suc (Suc n)", rotate_tac -2) apply (frule sym, thin_tac "f (Suc n) = f (Suc (Suc n))", simp, thin_tac "f (Suc (Suc n)) = f (Suc n)") apply (cut_tac n_in_Nsetn[of "Suc n"], blast) apply (frule_tac m = xa and n = "Suc (Suc n)" in noteq_le_less, assumption, frule_tac x = xa in Suc_less_le[of _ "Suc n"], blast) apply (subgoal_tac "{i. i \ (Suc n) \ Qw_cmpser G f i \ E} = insert (Suc n) {i. i \ n \ Qw_cmpser G f i \ E}") apply simp apply fastforce done lemma length_wcmpser0_6:"\Group G; w_cmpser G (Suc (Suc n)) f\ \ 0 < card (f ` {i. i \ (Suc n)})" apply (insert finite_Collect_le_nat [of "Suc n"]) apply (frule finite_imageI [of "{i. i \ (Suc n)}" "f"]) apply (subgoal_tac "{f 0} \ f ` {i. i \ (Suc n)}") apply (frule card_mono [of "f ` {i. i \ (Suc n)}" "{f 0}"], assumption+) apply (simp add:card1 [of "f 0"]) apply (rule subsetI, simp add:image_def) apply (subgoal_tac "0 \ {i. i \ (Suc n)}", blast) apply simp done lemma length_wcmpser0_7:"\Group G; w_cmpser G (Suc (Suc n)) f\ \ card {i. i \ n \ Qw_cmpser G f i \ E} \ Suc n" apply (insert finite_Collect_le_nat [of "n"]) apply (subgoal_tac "{i. i \ n \ Qw_cmpser G f i \ E} \ {i. i \ n}") apply (frule card_mono [of "{i. i \ n}" "{i. i \ n \ Qw_cmpser G f i \ E}"]) apply (assumption, simp) apply (rule subsetI, simp add:CollectI) done lemma length_wcmpser0:"\Group G; Ugp E\ \\f. w_cmpser G (Suc n) f \ card (f ` {i. i \ (Suc n)}) - 1 = (Suc n) - (card {i. i \ n \ ((Qw_cmpser G f) i \ E)})" apply (induct_tac n) (*** n = 0 ***) apply (rule allI) apply (rule impI) apply (frule_tac f = f in length_wcmpser0_4[of G E], assumption+, simp) (**** n ****) apply (rule allI) apply (rule impI) apply (frule_tac n = "Suc n" and f = f in Group.w_cmpser_pre [of "G"], assumption+) apply (subgoal_tac "card (f ` {i. i \ (Suc n)}) - 1 = Suc n - card {i. i \ n \ Qw_cmpser G f i \ E}") prefer 2 apply simp apply (thin_tac " \f. w_cmpser G (Suc n) f \ card (f ` {i. i \ (Suc n)}) - 1 = Suc n - card {i. i \ n \ Qw_cmpser G f i \ E}") apply (case_tac "Qw_cmpser G f (Suc n) \ E") apply (rule length_wcmpser0_5, assumption+) apply (frule_tac n = "Suc n" and f = f and i = "Suc n" in length_wcmpser0_2 [of "G" "E"], assumption+) apply simp apply assumption apply (subgoal_tac "f ` {i. i \ (Suc (Suc n))} = insert (f (Suc (Suc n))) (f ` {i. i \ (Suc n)})") apply simp apply (thin_tac "f ` {i. i \ (Suc (Suc n))} = insert (f (Suc (Suc n))) (f ` {i. i \ (Suc n)})") apply (subgoal_tac "finite (f ` {i. i \ (Suc n)})") apply (subgoal_tac "f (Suc (Suc n)) \ f ` {i. i \ (Suc n)}") apply (subst card_insert_disjoint, assumption) apply assumption prefer 2 apply (rule length_wcmpser0_3, assumption+) prefer 2 apply (subgoal_tac "finite {i. i \ (Suc n)}") apply (rule finite_imageI, assumption+, simp) prefer 2 apply (thin_tac " \ Qw_cmpser G f (Suc n) \ E", thin_tac " w_cmpser G (Suc n) f", thin_tac "f (Suc n) \ f (Suc (Suc n))") apply (subgoal_tac "{i. i \ (Suc (Suc n))} = {i. i\(Suc n)} \ {Suc (Suc n)}") prefer 2 apply (rule_tac n = "Suc n" in Nset_un, simp) apply (subgoal_tac "card {i. i \ (Suc n) \ Qw_cmpser G f i \ E} = card {i. i \ n \ Qw_cmpser G f i \ E}") apply simp apply (thin_tac " card {i. i \ (Suc n) \ Qw_cmpser G f i \ E} = card {i. i \ n \ Qw_cmpser G f i \ E}", thin_tac "\ Qw_cmpser G f (Suc n) \ E", thin_tac "f (Suc n) \ f (Suc (Suc n))", thin_tac "f (Suc (Suc n)) \ f ` {i. i \ (Suc n)}") apply (frule_tac n = n and f = f in length_wcmpser0_6 [of "G"], assumption+, frule_tac n = n and f = f in length_wcmpser0_7 [of "G"], assumption+) apply (rule abc_SucTr0, assumption+) apply (rule card_eq) apply (thin_tac "card (f ` {i. i \ (Suc n)}) - Suc 0 = Suc n - card {i. i \ n \ Qw_cmpser G f i \ E}", thin_tac "f (Suc n) \ f (Suc (Suc n))", thin_tac "f (Suc (Suc n)) \ f ` {i. i \ (Suc n)}") apply (rule equalityI) apply (rule subsetI, simp add:CollectI, erule conjE) apply (case_tac "x = Suc n", simp, simp) apply (rule subsetI, simp add:CollectI) done lemma length_of_twcmpser:"\Group G; Ugp E; tw_cmpser G (Suc n) f \ \ length_twcmpser G (Suc n) f = (Suc n) - (card {i. i \ n \ ((Qw_cmpser G f) i \ E)})" apply (unfold length_twcmpser_def) apply (insert length_wcmpser0 [of "G" "E" "n"]) apply (subgoal_tac "w_cmpser G (Suc n) f", rotate_tac -1, simp, simp, thin_tac "\f. w_cmpser G (Suc n) f \ card (f ` {i. i \ (Suc n)}) - Suc 0 = Suc n - card {i. i \ n \ Qw_cmpser G f i \ E}") apply (simp add:tw_cmpser_def w_cmpser_def, erule conjE) apply (thin_tac "\i\ n. Gp G (f i) \ f (Suc i)") apply (simp add:td_gchain_def) done lemma JHS_1:"\Group G; Ugp E; compseries G r f; compseries G s g; 0 \ r = r * s - card {i. i \ (r * s - Suc 0) \ Qw_cmpser G (cmp_rfn G r f s g) i \ E}" apply (frule_tac r = r and s = s and G = G and f = f and g = g in Group.JHS_Tr0, assumption+) apply (simp add:wcsr_rfns_def, erule conjE) apply (frule_tac length_of_twcmpser [of "G" "E" "r * s - Suc 0" "cmp_rfn G r f s g"], assumption+, simp add:mult.commute) apply (simp add:length_twcmpser_def) apply (frule Group.rfn_compseries_iM [of "G" "r" "s" "f" "cmp_rfn G r f s g"], assumption+, rule Group.JHS_Tr0 [of "G" "r" "s" "f" "g"], assumption+) apply (simp add:mult.commute [of "s" "r"]) done lemma J_H_S:"\Group G; Ugp E; compseries G r f; compseries G s g; 0 \ r = s" apply (frule JHS_1 [of "G" "E" "r" "f" "s" "g"], assumption+, frule JHS_1 [of "G" "E" "s" "g" "r" "f"], assumption+, frule JHS_Tr1 [of "G" "r" "s" "f" "g"], assumption+, frule Group.JHS_Tr0 [of "G" "r" "s" "f" "g"], assumption+, frule Group.JHS_Tr0 [of "G" "s" "r" "g" "f"], assumption+) apply (simp add:wcsr_rfns_def, (erule conjE)+, frule Group.tw_cmpser_is_w_cmpser [of "G" "s * r" "cmp_rfn G r f s g"], assumption+, frule Qw_cmpser [of "G" "s * r - Suc 0" "cmp_rfn G r f s g"], simp add:pos_mult_pos [of "s" "r"]) apply (frule Group.tw_cmpser_is_w_cmpser [of "G" "r * s" "cmp_rfn G s g r f"], assumption+, frule Qw_cmpser [of "G" "r * s - Suc 0" "cmp_rfn G s g r f"], simp add:pos_mult_pos [of "r" "s"], simp add:mult.commute [of "s" "r"]) apply (frule isom_gch_units [of "E" "r * s - Suc 0" "Qw_cmpser G (cmp_rfn G r f s g)" "Qw_cmpser G (cmp_rfn G s g r f)" "rtos r s"], assumption+) prefer 2 apply simp apply (simp add:Gch_bridge_def) apply (rule conjI) apply (rule allI, rule impI) apply (frule_tac l = l in rtos_hom1 [of "r" "s"], assumption+, simp add:mult.commute [of "s" "r"]) apply (rule rtos_inj, assumption+) done end diff --git a/thys/LinearQuantifierElim/Thys/QEdlo_inf.thy b/thys/LinearQuantifierElim/Thys/QEdlo_inf.thy --- a/thys/LinearQuantifierElim/Thys/QEdlo_inf.thy +++ b/thys/LinearQuantifierElim/Thys/QEdlo_inf.thy @@ -1,229 +1,229 @@ (* Author: Tobias Nipkow, 2007 *) theory QEdlo_inf imports DLO begin subsection "Quantifier elimination with infinitesimals" text\This section presents a new quantifier elimination procedure for dense linear orders based on (the simulation of) infinitesimals. It is a fairly straightforward adaptation of the analogous algorithm by Loos and Weispfenning for linear arithmetic described in \S\ref{sec:lin-inf}.\ fun asubst_peps :: "nat \ atom \ atom fm" ("asubst\<^sub>+") where "asubst_peps k (Less 0 0) = FalseF" | "asubst_peps k (Less 0 (Suc j)) = Atom(Less k j)" | "asubst_peps k (Less (Suc i) 0) = (if i=k then TrueF else Or (Atom(Less i k)) (Atom(Eq i k)))" | "asubst_peps k (Less (Suc i) (Suc j)) = Atom(Less i j)" | "asubst_peps k (Eq 0 0) = TrueF" | "asubst_peps k (Eq 0 _) = FalseF" | "asubst_peps k (Eq _ 0) = FalseF" | "asubst_peps k (Eq (Suc i) (Suc j)) = Atom(Eq i j)" abbreviation subst_peps :: "atom fm \ nat \ atom fm" ("subst\<^sub>+") where "subst\<^sub>+ \ k \ amap\<^sub>f\<^sub>m (asubst\<^sub>+ k) \" definition "nolb \ xs l x = (\y\{l<.. LB \ xs)" lemma nolb_And[simp]: "nolb (And \\<^sub>1 \\<^sub>2) xs l x = (nolb \\<^sub>1 xs l x \ nolb \\<^sub>2 xs l x)" apply(clarsimp simp:nolb_def) apply blast done lemma nolb_Or[simp]: "nolb (Or \\<^sub>1 \\<^sub>2) xs l x = (nolb \\<^sub>1 xs l x \ nolb \\<^sub>2 xs l x)" apply(clarsimp simp:nolb_def) apply blast done context notes [[simp_depth_limit=3]] begin lemma innermost_intvl: "\ nqfree \; nolb \ xs l x; l < x; x \ EQ \ xs; DLO.I \ (x#xs); l < y; y \ x\ \ DLO.I \ (y#xs)" proof(induct \) case (Atom a) show ?case proof (cases a) case (Less i j) then show ?thesis using Atom unfolding nolb_def by (clarsimp simp: nth.simps Ball_def split:if_split_asm nat.splits) (metis not_le_imp_less order_antisym order_less_trans)+ next case (Eq i j) thus ?thesis using Atom apply(clarsimp simp:EQ_def nolb_def nth_Cons') apply(case_tac "i=0 \ j=0") apply simp apply(case_tac "i\0 \ j\0") apply simp apply(case_tac "i=0 \ j\0") apply (fastforce split:if_split_asm) apply(case_tac "i\0 \ j=0") apply (fastforce split:if_split_asm) apply arith done qed next case And thus ?case by (fastforce) next case Or thus ?case by (fastforce) qed simp+ lemma I_subst_peps2: "nqfree \ \ xs!l < x \ nolb \ xs (xs!l) x \ x \ EQ \ xs \ \y \ {xs!l <.. x}. DLO.I \ (y#xs) \ DLO.I (subst\<^sub>+ \ l) xs" proof(induct \) case FalseF thus ?case - by simp (metis linorder_antisym_conv1 linorder_neq_iff) + by simp (metis antisym_conv1 linorder_neq_iff) next case (Atom a) show ?case proof(cases "(l,a)" rule:asubst_peps.cases) case 3 thus ?thesis using Atom by (auto simp: nolb_def EQ_def Ball_def) - (metis One_nat_def linorder_antisym_conv1 not_less_iff_gr_or_eq) + (metis One_nat_def antisym_conv1 not_less_iff_gr_or_eq) qed (insert Atom, auto simp: nolb_def EQ_def Ball_def) next case Or thus ?case by(simp add: Ball_def)(metis order_refl innermost_intvl) qed simp_all end lemma dense_interval: assumes "finite L" "l \ L" "l < x" "P(x::'a::dlo)" and dense: "\y l. \ \y\{l<.. L; lx \ \ P y" shows "\l\L. l (\y\{l<.. L) \ (\y. l y\x \ P y)" proof - let ?L = "{l\L. l < x}" let ?ll = "Max ?L" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) hence "\y. ?ll y y \ L" using \finite L\ by force moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp ultimately show ?thesis using \l < x\ \?L \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed lemma I_subst_peps: "nqfree \ \ DLO.I (subst\<^sub>+ \ l) xs \ (\leps>xs!l. \x. xs!l < x \ x \ leps \ DLO.I \ (x#xs))" proof(induct \) case TrueF thus ?case by simp (metis no_ub) next case (Atom a) show ?case proof (cases "(l,a)" rule: asubst_peps.cases) case 2 thus ?thesis using Atom apply(auto) apply(drule dense) apply(metis One_nat_def xt1(7)) done next case 3 thus ?thesis using Atom apply(auto) apply (metis no_ub) apply (metis no_ub less_trans) apply (metis no_ub) done next case 4 thus ?thesis using Atom by(auto)(metis no_ub) next case 5 thus ?thesis using Atom by(auto)(metis no_ub) next case 8 thus ?thesis using Atom by(auto)(metis no_ub) qed (insert Atom, auto) next case And thus ?case apply clarsimp apply(rule_tac x="min leps lepsa" in exI) apply simp done next case Or thus ?case by force qed simp_all definition "qe_eps\<^sub>1(\) = (let as = DLO.atoms\<^sub>0 \; lbs = lbounds as; ebs = ebounds as in list_disj (inf\<^sub>- \ # map (subst\<^sub>+ \) lbs @ map (subst \) ebs))" theorem I_qe_eps1: assumes "nqfree \" shows "DLO.I (qe_eps\<^sub>1 \) xs = (\x. DLO.I \ (x#xs))" (is "?QE = ?EX") proof let ?as = "DLO.atoms\<^sub>0 \" let ?ebs = "ebounds ?as" assume ?QE { assume "DLO.I (inf\<^sub>- \) xs" hence ?EX using \?QE\ min_inf[of \ xs] \nqfree \\ by(auto simp add:qe_eps\<^sub>1_def amap_fm_list_disj) } moreover { assume "\i \ set ?ebs. \DLO.I \ (xs!i # xs)" "\ DLO.I (inf\<^sub>- \) xs" with \?QE\ \nqfree \\ obtain l where "DLO.I (subst\<^sub>+ \ l) xs" by(fastforce simp: I_subst qe_eps\<^sub>1_def set_ebounds set_lbounds) then obtain leps where "DLO.I \ (leps#xs)" using I_subst_peps[OF \nqfree \\] by fastforce hence ?EX .. } ultimately show ?EX by blast next let ?as = "DLO.atoms\<^sub>0 \" let ?ebs = "ebounds ?as" assume ?EX then obtain x where x: "DLO.I \ (x#xs)" .. { assume "DLO.I (inf\<^sub>- \) xs" hence ?QE using \nqfree \\ by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\k \ set ?ebs. DLO.I (subst \ k) xs" hence ?QE by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\ DLO.I (inf\<^sub>- \) xs" and "\k \ set ?ebs. \ DLO.I (subst \ k) xs" hence noE: "\e \ EQ \ xs. \ DLO.I \ (e#xs)" using \nqfree \\ by (auto simp:set_ebounds EQ_def I_subst nth_Cons' split:if_split_asm) hence "x \ EQ \ xs" using x by fastforce obtain l where "l \ LB \ xs" "l < x" using LBex[OF \nqfree \\ x \\ DLO.I(inf\<^sub>- \) xs\ \x \ EQ \ xs\] .. have "\l\LB \ xs. l nolb \ xs l x \ (\y. l < y \ y \ x \ DLO.I \ (y#xs))" using dense_interval[where P = "\x. DLO.I \ (x#xs)", OF finite_LB \l\LB \ xs\ \l x] x innermost_intvl[OF \nqfree \\ _ _ \x \ EQ \ xs\] by (simp add:nolb_def) then obtain m where *: "Less (Suc m) 0 \ set ?as \ xs!m < x \ nolb \ xs (xs!m) x \ (\y. xs!m < y \ y \ x \ DLO.I \ (y#xs))" by blast then have "DLO.I (subst\<^sub>+ \ m) xs" using noE by(auto intro!: I_subst_peps2[OF \nqfree \\]) with * have ?QE by(simp add:qe_eps\<^sub>1_def bex_Un set_lbounds set_ebounds) metis } ultimately show ?QE by blast qed lemma qfree_asubst_peps: "qfree (asubst\<^sub>+ k a)" by(cases "(k,a)" rule:asubst_peps.cases) simp_all lemma qfree_subst_peps: "nqfree \ \ qfree (subst\<^sub>+ \ k)" by(induct \) (simp_all add:qfree_asubst_peps) lemma qfree_qe_eps\<^sub>1: "nqfree \ \ qfree(qe_eps\<^sub>1 \)" apply(simp add:qe_eps\<^sub>1_def) apply(rule qfree_list_disj) apply (auto simp:qfree_min_inf qfree_subst_peps qfree_map_fm) done definition "qe_eps = DLO.lift_nnf_qe qe_eps\<^sub>1" lemma qfree_qe_eps: "qfree(qe_eps \)" by(simp add: qe_eps_def DLO.qfree_lift_nnf_qe qfree_qe_eps\<^sub>1) lemma I_qe_eps: "DLO.I (qe_eps \) xs = DLO.I \ xs" by(simp add:qe_eps_def DLO.I_lift_nnf_qe qfree_qe_eps\<^sub>1 I_qe_eps1) end diff --git a/thys/LinearQuantifierElim/Thys/QElin_inf.thy b/thys/LinearQuantifierElim/Thys/QElin_inf.thy --- a/thys/LinearQuantifierElim/Thys/QElin_inf.thy +++ b/thys/LinearQuantifierElim/Thys/QElin_inf.thy @@ -1,302 +1,302 @@ (* Author: Tobias Nipkow, 2007 *) theory QElin_inf imports LinArith begin subsection \Quantifier elimination with infinitesimals \label{sec:lin-inf}\ text\This section formalizes Loos and Weispfenning's quantifier elimination procedure based on (the simulation of) infinitesimals~\cite{LoosW93}.\ fun asubst_peps :: "real * real list \ atom \ atom fm" ("asubst\<^sub>+") where "asubst_peps (r,cs) (Less s (d#ds)) = (if d=0 then Atom(Less s ds) else let u = s - d*r; v = d *\<^sub>s cs + ds; less = Atom(Less u v) in if d<0 then less else Or less (Atom(Eq u v)))" | "asubst_peps rcs (Eq r (d#ds)) = (if d=0 then Atom(Eq r ds) else FalseF)" | "asubst_peps rcs a = Atom a" abbreviation subst_peps :: "atom fm \ real * real list \ atom fm" ("subst\<^sub>+") where "subst\<^sub>+ \ rcs \ amap\<^sub>f\<^sub>m (asubst\<^sub>+ rcs) \" definition "nolb f xs l x = (\y\{l<.. LB f xs)" lemma nolb_And[simp]: "nolb (And f g) xs l x = (nolb f xs l x \ nolb g xs l x)" apply(clarsimp simp:nolb_def) apply blast done lemma nolb_Or[simp]: "nolb (Or f g) xs l x = (nolb f xs l x \ nolb g xs l x)" apply(clarsimp simp:nolb_def) apply blast done context notes [[simp_depth_limit=4]] begin lemma innermost_intvl: "\ nqfree f; nolb f xs l x; l < x; x \ EQ f xs; R.I f (x#xs); l < y; y \ x\ \ R.I f (y#xs)" proof(induct f) case (Atom a) show ?case proof (cases a) case [simp]: (Less r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom by (simp add:depends\<^sub>R_def) next case [simp]: (Cons c cs) hence "r < c*x + \cs,xs\" using Atom by simp { assume "c=0" hence ?thesis using Atom by simp } moreover { assume "c<0" hence "x < (r - \cs,xs\)/c" (is "_ < ?u") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?u \ y" using Atom \c<0\ by (auto simp add: field_simps) with \x show False using Atom \c<0\ by(auto simp:depends\<^sub>R_def) qed } moreover { assume "c>0" hence "x > (r - \cs,xs\)/c" (is "_ > ?l") using \r < c*x + \cs,xs\\ by (simp add: field_simps) then have "?l < y" using Atom \c>0\ by (auto simp:depends\<^sub>R_def Ball_def nolb_def) (metis linorder_not_le antisym order_less_trans) hence ?thesis using \c>0\ by (simp add: field_simps) } ultimately show ?thesis by force qed next case [simp]: (Eq r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom by (simp add:depends\<^sub>R_def) next case [simp]: (Cons c cs) hence "r = c*x + \cs,xs\" using Atom by simp { assume "c=0" hence ?thesis using Atom by simp } moreover { assume "c\0" hence ?thesis using \r = c*x + \cs,xs\\ Atom by(auto simp: ac_simps depends\<^sub>R_def split:if_splits) } ultimately show ?thesis by force qed qed next case (And f1 f2) thus ?case by (fastforce) next case (Or f1 f2) thus ?case by (fastforce) qed simp+ definition "EQ2 = EQ" lemma EQ2_Or[simp]: "EQ2 (Or f g) xs = (EQ2 f xs \ EQ2 g xs)" by(auto simp:EQ2_def) lemma EQ2_And[simp]: "EQ2 (And f g) xs = (EQ2 f xs \ EQ2 g xs)" by(auto simp:EQ2_def) lemma innermost_intvl2: "\ nqfree f; nolb f xs l x; l < x; x \ EQ2 f xs; R.I f (x#xs); l < y; y \ x\ \ R.I f (y#xs)" unfolding EQ2_def by(blast intro:innermost_intvl) lemma I_subst_peps2: "nqfree f \ r+\cs,xs\ < x \ nolb f xs (r+\cs,xs\) x \ \y \ {r+\cs,xs\ <.. x}. R.I f (y#xs) \ y \ EQ2 f xs \ R.I (subst\<^sub>+ f (r,cs)) xs" proof(induct f) case FalseF thus ?case - by simp (metis linorder_antisym_conv1 linorder_neq_iff) + by simp (metis antisym_conv1 linorder_neq_iff) next case (Atom a) show ?case proof(cases "((r,cs),a)" rule:asubst_peps.cases) case (1 r cs s d ds) { assume "d=0" hence ?thesis using Atom 1 by auto } moreover { assume "d<0" have "s < d*x + \ds,xs\" using Atom 1 by simp moreover have "d*x < d*(r + \cs,xs\)" using \d<0\ Atom 1 by (simp add: mult_strict_left_mono_neg) ultimately have "s < d * (r + \cs,xs\) + \ds,xs\" by(simp add:algebra_simps) hence ?thesis using 1 by (auto simp add: iprod_left_add_distrib algebra_simps) } moreover { let ?L = "(s - \ds,xs\) / d" let ?U = "r + \cs,xs\" assume "d>0" hence "?U < x" and "\y. ?U < y \ y < x \ y \ ?L" and "\y. ?U < y \ y \ x \ ?L < y" using Atom 1 by(simp_all add:nolb_def depends\<^sub>R_def Ball_def field_simps) hence "?L < ?U \ ?L = ?U" by (metis linorder_neqE_linordered_idom order_refl) hence ?thesis using Atom 1 \d>0\ by (simp add: iprod_left_add_distrib field_simps) } ultimately show ?thesis by force next case 2 thus ?thesis using Atom by (fastforce simp: nolb_def EQ2_def depends\<^sub>R_def field_simps split: if_split_asm) qed (insert Atom, auto) next case Or thus ?case by(simp add:Ball_def)(metis order_refl innermost_intvl2) qed simp_all end lemma I_subst_peps: "nqfree f \ R.I (subst\<^sub>+ f (r,cs)) xs \ (\leps>r+\cs,xs\. \x. r+\cs,xs\ < x \ x \ leps \ R.I f (x#xs))" proof(induct f) case TrueF thus ?case by simp (metis less_add_one) next case (Atom a) show ?case proof (cases "((r,cs),a)" rule: asubst_peps.cases) case (1 r cs s d ds) { assume "d=0" hence ?thesis using Atom 1 by auto (metis less_add_one) } moreover { assume "d<0" with Atom 1 have "r + \cs,xs\ < (s - \ds,xs\)/d" (is "?a < ?b") by(simp add:field_simps iprod_left_add_distrib) then obtain x where "?a < x" "x < ?b" by(metis dense) hence " \y. ?a < y \ y \ x \ s < d*y + \ds,xs\" using \d<0\ by (simp add:field_simps) (metis add_le_cancel_right mult_le_cancel_left order_antisym linear mult.commute xt1(8)) hence ?thesis using 1 \?a by auto } moreover { let ?a = "s - d * r" let ?b = "\d *\<^sub>s cs + ds,xs\" assume "d>0" with Atom 1 have "?a < ?b \ ?a = ?b" by auto hence ?thesis proof assume "?a = ?b" thus ?thesis using \d>0\ Atom 1 by(simp add:field_simps iprod_left_add_distrib) (metis add_0_left add_less_cancel_right distrib_left mult.commute mult_strict_left_mono) next assume "?a < ?b" { fix x assume "r+\cs,xs\ < x \ x \ r+\cs,xs\ + 1" hence "d*(r + \cs,xs\) < d*x" using \d>0\ by(metis mult_strict_left_mono) hence "s < d*x + \ds,xs\" using \d>0\ \?a < ?b\ by (simp add:algebra_simps iprod_left_add_distrib) } thus ?thesis using 1 \d>0\ by(force simp: iprod_left_add_distrib) qed } ultimately show ?thesis by (metis less_linear) qed (insert Atom, auto split:if_split_asm intro: less_add_one) next case And thus ?case apply clarsimp apply(rule_tac x="min leps lepsa" in exI) apply simp done next case Or thus ?case by force qed simp_all lemma dense_interval: assumes "finite L" "l \ L" "l < x" "P(x::real)" and dense: "\y l. \ \y\{l<.. L; lx \ \ P y" shows "\l\L. l (\y\{l<.. L) \ (\y. l y\x \ P y)" proof - let ?L = "{l\L. l < x}" let ?ll = "Max ?L" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) hence "\y. ?ll y y \ L" using \finite L\ by force moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp ultimately show ?thesis using \l < x\ \?L \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed definition "qe_eps\<^sub>1(f) = (let as = R.atoms\<^sub>0 f; lbs = lbounds as; ebs = ebounds as in list_disj (inf\<^sub>- f # map (subst\<^sub>+ f) lbs @ map (subst f) ebs))" theorem I_eps1: assumes "nqfree f" shows "R.I (qe_eps\<^sub>1 f) xs = (\x. R.I f (x#xs))" (is "?QE = ?EX") proof let ?as = "R.atoms\<^sub>0 f" let ?ebs = "ebounds ?as" assume ?QE { assume "R.I (inf\<^sub>- f) xs" hence ?EX using \?QE\ min_inf[of f xs] \nqfree f\ by(auto simp add:qe_eps\<^sub>1_def amap_fm_list_disj) } moreover { assume "\x \ EQ f xs. \R.I f (x#xs)" "\ R.I (inf\<^sub>- f) xs" with \?QE\ \nqfree f\ obtain r cs where "R.I (subst\<^sub>+ f (r,cs)) xs" by (fastforce simp: qe_eps\<^sub>1_def set_ebounds diff_divide_distrib eval_def I_subst \nqfree f\) then obtain leps where "R.I f (leps#xs)" using I_subst_peps[OF \nqfree f\] by fastforce hence ?EX .. } ultimately show ?EX by blast next let ?as = "R.atoms\<^sub>0 f" let ?ebs = "ebounds ?as" assume ?EX then obtain x where x: "R.I f (x#xs)" .. { assume "R.I (inf\<^sub>- f) xs" hence ?QE using \nqfree f\ by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\rcs \ set ?ebs. R.I (subst f rcs) xs" hence ?QE by(auto simp:qe_eps\<^sub>1_def) } moreover { assume "\ R.I (inf\<^sub>- f) xs" and "\rcs \ set ?ebs. \ R.I (subst f rcs) xs" hence noE: "\e \ EQ f xs. \ R.I f (e#xs)" using \nqfree f\ by (force simp:set_ebounds I_subst diff_divide_distrib eval_def split:if_split_asm) hence "x \ EQ f xs" using x by fastforce obtain l where "l \ LB f xs" "l < x" using LBex[OF \nqfree f\ x \\ R.I(inf\<^sub>- f) xs\ \x \ EQ f xs\] .. have "\l\LB f xs. l nolb f xs l x \ (\y. l < y \ y \ x \ R.I f (y#xs))" using dense_interval[where P = "\x. R.I f (x#xs)", OF finite_LB \l\LB f xs\ \l x] x innermost_intvl[OF \nqfree f\ _ _ \x \ EQ f xs\] by (simp add:nolb_def) then obtain r c cs where *: "Less r (c#cs) \ set(R.atoms\<^sub>0 f) \ c>0 \ (r - \cs,xs\)/c < x \ nolb f xs ((r - \cs,xs\)/c) x \ (\y. (r - \cs,xs\)/c < y \ y \ x \ R.I f (y#xs))" by blast then have "R.I (subst\<^sub>+ f (r/c, (-1/c) *\<^sub>s cs)) xs" using noE by(auto intro!: I_subst_peps2[OF \nqfree f\] simp:EQ2_def diff_divide_distrib algebra_simps) with * have ?QE by(simp add:qe_eps\<^sub>1_def bex_Un set_lbounds) metis } ultimately show ?QE by blast qed lemma qfree_asubst_peps: "qfree (asubst\<^sub>+ rcs a)" by(cases "(rcs,a)" rule:asubst_peps.cases) simp_all lemma qfree_subst_peps: "nqfree \ \ qfree (subst\<^sub>+ \ rcs)" by(induct \) (simp_all add:qfree_asubst_peps) lemma qfree_qe_eps\<^sub>1: "nqfree \ \ qfree(qe_eps\<^sub>1 \)" apply(simp add:qe_eps\<^sub>1_def) apply(rule qfree_list_disj) apply (auto simp:qfree_min_inf qfree_subst_peps qfree_map_fm) done definition "qe_eps = R.lift_nnf_qe qe_eps\<^sub>1" lemma qfree_qe_eps: "qfree(qe_eps \)" by(simp add: qe_eps_def R.qfree_lift_nnf_qe qfree_qe_eps\<^sub>1) lemma I_qe_eps: "R.I (qe_eps \) xs = R.I \ xs" by(simp add:qe_eps_def R.I_lift_nnf_qe qfree_qe_eps\<^sub>1 I_eps1) end diff --git a/thys/Projective_Geometry/Matroid_Rank_Properties.thy b/thys/Projective_Geometry/Matroid_Rank_Properties.thy --- a/thys/Projective_Geometry/Matroid_Rank_Properties.thy +++ b/thys/Projective_Geometry/Matroid_Rank_Properties.thy @@ -1,248 +1,248 @@ theory Matroid_Rank_Properties imports Main Higher_Projective_Space_Rank_Axioms begin (* Author: Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk .*) text \ Contents: \<^item> In this file we introduce the basic lemmas and properties derived from our based-rank axioms that will allow us to simplify our future proofs. \ section \Proof Techniques Using Ranks\ lemma matroid_ax_3_alt: assumes "I \ X \ Y" shows "rk (X \ Y) + rk I \ rk X + rk Y" by (metis add.commute add_le_cancel_right assms matroid_ax_2 matroid_ax_3 order_trans) lemma rk_uniqueness: assumes "rk {A,B} = 2" and "rk {C,D} = 2" and "rk {A,B,M} \ 2" and "rk {C,D,M} \ 2" and "rk {A,B,P} \ 2" and "rk {C,D,P} \ 2" and "rk {A,B,C,D} \ 3" shows "rk {M,P} = 1" proof- have "{A,B,M} \ {A,B,P} = {A,B,M,P}" by auto have "rk {A,B,M,P} + rk {A,B} \ rk {A,B,M} + rk {A,B,P}" by (metis (full_types) \{A, B, M} \ {A, B, P} = {A, B, M, P}\ insert_commute le_inf_iff matroid_ax_3_alt subset_insertI) then have "rk {A,B,M,P} = 2" by (smt Un_upper1 Un_upper2 \{A, B, M} \ {A, B, P} = {A, B, M, P}\ add_diff_cancel_left' add_diff_cancel_right' add_mono antisym assms(1) assms(3) assms(5) le_diff_conv matroid_ax_2) have "{C,D,M} \ {C,D,P} = {C,D,M,P}" by auto have "rk {C,D,M,P} + rk {C,D} \ rk {C,D,M} + rk {C,D,P}" by (metis Un_insert_left Un_upper1 \{C, D, M} \ {C, D, P} = {C, D, M, P}\ insert_is_Un le_inf_iff matroid_ax_3_alt) then have i1:"rk {C,D,M,P} + 2 \ 4" using assms(2) assms(4) assms(6) by linarith moreover have i2:"rk {C,D,M,P} \ 2" by (metis assms(2) insertI1 insert_subset matroid_ax_2 subset_insertI) from i1 and i2 have "rk {C,D,M,P} = 2" by linarith have "rk {A,B,C,D,M,P} \ 3" by (metis Un_insert_right Un_upper2 assms(7) matroid_ax_2 order_trans sup_bot.right_neutral) have "{A,B,M,P} \ {C,D,M,P} = {A,B,C,D,M,P}" by auto then have "rk {A,B,C,D,M,P} + rk {M,P} \ rk {A,B,M,P} + rk {C,D,M,P}" by (smt le_inf_iff matroid_ax_3_alt order_trans subset_insertI) then have i3:"rk {A,B,C,D,M,P} + rk {M,P} \ 4" using \rk {A, B, M, P} = 2\ \rk {C, D, M, P} = 2\ by linarith have i4:"rk {A,B,C,D,M,P} + rk {M,P} \ 3 + rk{M,P}" by (simp add: \3 \ rk {A, B, C, D, M, P}\) from i3 and i4 show "rk {M,P} = 1" by (metis (no_types, lifting) \rk {A, B, C, D, M, P} + rk {M, P} \ rk {A, B, M, P} + rk {C, D, M, P}\ \rk {A, B, M, P} = 2\ \rk {C, D, M, P} = 2\ add_le_cancel_left add_numeral_left antisym insert_absorb2 numeral_Bit1 numeral_One numeral_plus_one one_add_one one_le_numeral order_trans rk_ax_couple rk_ax_singleton) qed (* The following lemma allows to derive that there exists two lines that do not meet, i.e that belong to two different planes *) lemma rk_ax_dim_alt: "\A B C D. \M. rk {A,B,M} \ 2 \ rk {C,D,M} \ 2" proof- obtain A B C D where f1:"rk {A,B,C,D} \ 4" using rk_ax_dim by auto have "\M. rk {A,B,M} \ 2 \ rk {C,D,M} \ 2" proof fix M have "{A,B,C,D,M} = {A,B,M} \ {C,D,M}" by auto then have "rk {A,B,C,D,M} + rk {M} \ rk {A,B,M} + rk {C,D,M}" by (smt le_inf_iff matroid_ax_3_alt order_trans subset_insertI) then have "rk {A,B,C,D,M} \ 3" if "rk {A,B,M} = 2" and "rk {C,D,M} = 2" by (smt One_nat_def Suc_leI Suc_le_mono Suc_numeral add_Suc_right add_leD1 add_mono le_antisym not_le numeral_3_eq_3 numeral_One numeral_plus_one one_add_one rk_ax_singleton that(1) that(2)) then have "rk {A,B,C,D} \ 3" if "rk {A,B,M} = 2" and "rk {C,D,M} = 2" by (smt insert_commute matroid_ax_2 order_trans subset_insertI that(1) that(2)) thus "rk {A, B, M} \ 2 \ rk {C, D, M} \ 2 " using \4 \ rk {A, B, C, D}\ by linarith qed thus "\A B C D. \M. rk {A, B, M} \ 2 \ rk {C, D, M} \ 2" by blast qed lemma rk_empty: "rk {} = 0" proof- have "rk {} \ 0" by simp have "rk {} \ 0" by (metis card.empty matroid_ax_1b) thus "rk {} = 0" by blast qed lemma matroid_ax_2_alt: "rk X \ rk (X \ {x}) \ rk (X \ {x}) \ rk X + 1" proof have "X \ X \ {x}" by auto thus "rk X \ rk (X \ {x})" by (simp add: matroid_ax_2) have "rk {x} \ 1" by (metis One_nat_def card.empty card_Suc_eq insert_absorb insert_not_empty matroid_ax_1b) thus "rk (X \ {x}) \ rk X + 1" by (metis add_leD1 le_antisym matroid_ax_3 rk_ax_singleton) qed lemma matroid_ax_3_alt': "rk (X \ {y}) = rk (X \ {z}) \ rk (X \ {z}) = rk X \ rk X = rk (X \ {y,z})" proof- have i1:"rk X \ rk (X \ {y,z})" using matroid_ax_2 by blast have i2:"rk X \ rk (X \ {y,z})" if "rk (X \ {y}) = rk (X \ {z})" and "rk (X \ {z}) = rk X" proof- have "(X \ {y}) \ (X \ {z}) = X \ {y,z}" by blast then have "rk (X \ {y,z}) + rk X \ rk X + rk X" by (metis \rk (X \ {y}) = rk (X \ {z})\ \rk (X \ {z}) = rk X\ inf_sup_ord(3) le_inf_iff matroid_ax_3_alt) thus "rk (X \ {y,z}) \ rk X" by simp qed thus "rk (X \ {y}) = rk (X \ {z}) \ rk (X \ {z}) = rk X \ rk X = rk (X \ {y, z})" using antisym i1 by blast qed lemma rk_ext: assumes "rk X \ 3" shows "\P. rk(X \ {P}) = rk X + 1" proof- obtain A B C D where "rk {A,B,C,D} \ 4" using rk_ax_dim by auto have f1:"rk (X \ {A, B, C, D}) \ 4" by (metis Un_upper2 \4 \ rk {A, B, C, D}\ matroid_ax_2 sup.coboundedI2 sup.orderE) have "rk (X \ {A, B, C, D}) = rk X" if "rk(X \ {A}) = rk(X \ {B})" and "rk(X \ {B}) = rk(X \ {C})" and "rk(X \ {C}) = rk(X \ {D})" and "rk(X \ {D}) = rk X" using matroid_ax_3_alt' that(1) that(2) that(3) that(4) by auto then have f2:"rk (X \ {A, B, C, D}) \ 3" if "rk(X \ {A}) = rk(X \ {B})" and "rk(X \ {B}) = rk(X \ {C})" and "rk(X \ {C}) = rk(X \ {D})" and "rk(X \ {D}) = rk X" using assms that(1) that(2) that(3) that(4) by linarith from f1 and f2 have "False" if "rk(X \ {A}) = rk(X \ {B})" and "rk(X \ {B}) = rk(X \ {C})" and "rk(X \ {C}) = rk(X \ {D})" and "rk(X \ {D}) = rk X" using that(1) that(2) that(3) that(4) by linarith then have "rk (X \ {A}) = rk X + 1 \ rk (X \ {B}) = rk X + 1 \ rk (X \ {C}) = rk X + 1 \ rk (X \ {D}) = rk X + 1" by (smt One_nat_def Suc_le_eq Suc_numeral Un_upper2 \4 \ rk {A, B, C, D}\ \\rk (X \ {A}) = rk (X \ {B}); rk (X \ {B}) = rk (X \ {C}); rk (X \ {C}) = rk (X \ {D}); rk (X \ {D}) = rk X\ \ rk (X \ {A, B, C, D}) = rk X\ - add.right_neutral add_Suc_right assms linorder_antisym_conv1 matroid_ax_2 matroid_ax_2_alt + add.right_neutral add_Suc_right assms antisym_conv1 matroid_ax_2 matroid_ax_2_alt not_less semiring_norm(2) semiring_norm(8) sup.coboundedI2 sup.orderE) thus "\P . rk (X \ {P}) = rk X + 1" by blast qed lemma rk_singleton : "\P. rk {P} = 1" proof fix P have f1:"rk {P} \ 1" by (metis One_nat_def card.empty card_Suc_eq insert_absorb insert_not_empty matroid_ax_1b) have f2:"rk {P} \ 1" using rk_ax_singleton by auto from f1 and f2 show "rk {P} = 1" using antisym by blast qed lemma rk_singleton_bis : assumes "A = B" shows "rk {A, B} = 1" by (simp add: assms rk_singleton) lemma rk_couple : assumes "A \ B" shows "rk {A, B} = 2" proof- have f1:"rk {A, B} \ 2" by (metis insert_is_Un matroid_ax_2_alt one_add_one rk_singleton) have f2:"rk {A, B} \ 2" by (simp add: assms rk_ax_couple) from f1 and f2 show "?thesis" by (simp add: f1 le_antisym) qed lemma rk_triple_le : "rk {A, B, C} \ 3" by (metis Suc_numeral Un_commute insert_absorb2 insert_is_Un linear matroid_ax_2_alt numeral_2_eq_2 numeral_3_eq_3 numeral_le_one_iff numeral_plus_one rk_couple rk_singleton semiring_norm(70)) lemma rk_couple_to_singleton : assumes "rk {A, B} = 1" shows "A = B" proof- have "rk {A, B} = 2" if "A \ B" using rk_couple by (simp add: that) thus "A = B" using assms by auto qed lemma rk_triple_to_rk_couple : assumes "rk {A, B, C} = 3" shows "rk {A, B} = 2" proof- have "rk {A, B} \ 2" using matroid_ax_1b by (metis one_le_numeral rk_ax_couple rk_couple rk_singleton_bis) have "rk {A, B, C} \ 2" if "rk {A, B} = 1" using matroid_ax_2_alt[of "{A, B}" C] by (simp add: insert_commute that) then have "rk {A, B} \ 2" using assms rk_ax_couple rk_singleton_bis by force thus "rk {A, B} = 2" by (simp add: \rk {A, B} \ 2\ le_antisym) qed end