diff --git a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy --- a/thys/Extended_Finite_State_Machines/EFSM_LTL.thy +++ b/thys/Extended_Finite_State_Machines/EFSM_LTL.thy @@ -1,281 +1,281 @@ section\LTL for EFSMs\ text\This theory builds off the \texttt{Linear\_Temporal\_Logic\_on\_Streams} theory from the HOL library and defines functions to ease the expression of LTL properties over EFSMs. Since the LTL operators effectively act over traces of models we must find a way to express models as streams.\ theory EFSM_LTL imports "Extended_Finite_State_Machines.EFSM" "HOL-Library.Linear_Temporal_Logic_on_Streams" begin text_raw\\snip{statedef}{1}{2}{%\ record state = statename :: "nat option" datastate :: registers action :: action "output" :: outputs text_raw\}%endsnip\ text_raw\\snip{whitebox}{1}{2}{%\ type_synonym whitebox_trace = "state stream" text_raw\}%endsnip\ type_synonym property = "whitebox_trace \ bool" abbreviation label :: "state \ String.literal" where "label s \ fst (action s)" abbreviation inputs :: "state \ value list" where "inputs s \ snd (action s)" text_raw\\snip{ltlStep}{1}{2}{%\ fun ltl_step :: "transition_matrix \ cfstate option \ registers \ action \ (nat option \ outputs \ registers)" where "ltl_step _ None r _ = (None, [], r)" | "ltl_step e (Some s) r (l, i) = (let possibilities = possible_steps e s r l i in if possibilities = {||} then (None, [], r) else let (s', t) = Eps (\x. x |\| possibilities) in (Some s', (evaluate_outputs t i r), (evaluate_updates t i r)) )" text_raw\}%endsnip\ lemma ltl_step_singleton: "\t. possible_steps e n r (fst v) (snd v) = {|(aa, t)|} \ evaluate_outputs t (snd v) r = b \ evaluate_updates t (snd v) r = c\ ltl_step e (Some n) r v = (Some aa, b, c)" apply (cases v) by auto lemma ltl_step_none: "possible_steps e s r a b = {||} \ ltl_step e (Some s) r (a, b) = (None, [], r)" by simp lemma ltl_step_none_2: "possible_steps e s r (fst ie) (snd ie) = {||} \ ltl_step e (Some s) r ie = (None, [], r)" by (metis ltl_step_none prod.exhaust_sel) lemma ltl_step_alt: "ltl_step e (Some s) r t = ( let possibilities = possible_steps e s r (fst t) (snd t) in if possibilities = {||} then (None, [], r) else let (s', t') = Eps (\x. x |\| possibilities) in (Some s', (apply_outputs (Outputs t') (join_ir (snd t) r)), (apply_updates (Updates t') (join_ir (snd t) r) r)) )" by (case_tac t, simp add: Let_def) lemma ltl_step_some: assumes "possible_steps e s r l i = {|(s', t)|}" and "evaluate_outputs t i r = p" and "evaluate_updates t i r = r'" shows "ltl_step e (Some s) r (l, i) = (Some s', p, r')" by (simp add: assms) lemma ltl_step_cases: assumes invalid: "P (None, [], r)" and valid: "\(s', t) |\| (possible_steps e s r l i). P (Some s', (evaluate_outputs t i r), (evaluate_updates t i r))" shows "P (ltl_step e (Some s) r (l, i))" apply simp apply (case_tac "possible_steps e s r l i") apply (simp add: invalid) apply simp subgoal for x S' apply (case_tac "SOME xa. xa = x \ xa |\| S'") apply simp apply (insert assms(2)) apply (simp add: fBall_def Ball_def fmember_def) by (metis (mono_tags, lifting) fst_conv prod.case_eq_if snd_conv someI_ex) done text\The \texttt{make\_full\_observation} function behaves similarly to \texttt{observe\_execution} from the \texttt{EFSM} theory. The main difference in behaviour is what is recorded. While the observe execution function simply observes an execution of the EFSM to produce the corresponding output for each action, the intention here is to record every detail of execution, including the values of internal variables. Thinking of each action as a step forward in time, there are five components which characterise a given point in the execution of an EFSM. At each point, the model has a current control state and data state. Each action has a label and some input parameters, and its execution may produce some observableoutput. It is therefore sufficient to provide a stream of 5-tuples containing the current control state, data state, the label and inputs of the action, and computed output. The make full observation function can then be defined as in Figure 9.1, with an additional function watch defined on top of this which starts the make full observation off in the initial control state with the empty data state. Careful inspection of the definition reveals another way that \texttt{make\_full\_observation} differs from \texttt{observe\_execution}. Rather than taking a cfstate, it takes a cfstate option. The reason for this is that we need to make our EFSM models complete. That is, we need them to be able to respond to every action from every state like a DFA. If a model does not recognise a given action in a given state, we cannot simply stop processing because we are working with necessarily infinite traces. Since these traces are generated by observing action sequences, the make full observation function must keep processing whether there is a viable transition or not. To support this, the make full observation adds an implicit ``sink state'' to every EFSM it processes by lifting control flow state indices from \texttt{nat} to \texttt{nat option} such that state $n$ is seen as state \texttt{Some} $n$. The control flow state \texttt{None} represents a sink state. If a model is unable to recognise a particular action from its current state, it moves into the \texttt{None} state. From here, the behaviour is constant for the rest of the time --- the control flow state remains None; the data state does not change, and no output is produced.\ text_raw\\snip{makeFullObservation}{1}{2}{%\ primcorec make_full_observation :: "transition_matrix \ cfstate option \ registers \ outputs \ action stream \ whitebox_trace" where "make_full_observation e s d p i = ( let (s', o', d') = ltl_step e s d (shd i) in \statename = s, datastate = d, action=(shd i), output = p\##(make_full_observation e s' d' o' (stl i)) )" text_raw\}%endsnip\ text_raw\\snip{watch}{1}{2}{%\ abbreviation watch :: "transition_matrix \ action stream \ whitebox_trace" where "watch e i \ (make_full_observation e (Some 0) <> [] i)" text_raw\}%endsnip\ subsection\Expressing Properties\ text\In order to simplify the expression and understanding of properties, this theory defines a number of named functions which can be used to express certain properties of EFSMs.\ subsubsection\State Equality\ text\The \textsc{state\_eq} takes a cfstate option representing a control flow state index and returns true if this is the control flow state at the head of the full observation.\ abbreviation state_eq :: "cfstate option \ whitebox_trace \ bool" where "state_eq v s \ statename (shd s) = v" lemma state_eq_holds: "state_eq s = holds (\x. statename x = s)" apply (rule ext) by (simp add: holds_def) lemma state_eq_None_not_Some: "state_eq None s \ \ state_eq (Some n) s" by simp subsubsection\Label Equality\ text\The \textsc{label\_eq} function takes a string and returns true if this is equal to the label at the head of the full observation.\ abbreviation "label_eq v s \ fst (action (shd s)) = (String.implode v)" lemma watch_label: "label_eq l (watch e t) = (fst (shd t) = String.implode l)" by (simp add: ) subsubsection\Input Equality\ text\The \textsc{input\_eq} function takes a value list and returns true if this is equal to the input at the head of the full observation.\ abbreviation "input_eq v s \ inputs (shd s) = v" subsubsection\Action Equality\ text\The \textsc{action\_eq} function takes a (label, value list) pair and returns true if this is equal to the action at the head of the full observation. This effectively combines \texttt{label\_eq} and \texttt{input\_eq} into one function.\ abbreviation "action_eq e \ label_eq (fst e) aand input_eq (snd e)" subsubsection\Output Equality\ text\The \textsc{output\_eq} function takes a takes a value option list and returns true if this is equal to the output at the head of the full observation.\ abbreviation "output_eq v s \ output (shd s) = v" text_raw\\snip{ltlVName}{1}{2}{%\ datatype ltl_vname = Ip nat | Op nat | Rg nat text_raw\}%endsnip\ subsubsection\Checking Arbitrary Expressions\ text\The \textsc{check\_exp} function takes a guard expression and returns true if the guard expression evaluates to true in the given state.\ type_synonym ltl_gexp = "ltl_vname gexp" definition join_iro :: "value list \ registers \ outputs \ ltl_vname datastate" where "join_iro i r p = (\x. case x of Rg n \ r $ n | Ip n \ Some (i ! n) | Op n \ p ! n )" lemma join_iro_R [simp]: "join_iro i r p (Rg n) = r $ n" by (simp add: join_iro_def) abbreviation "check_exp g s \ (gval g (join_iro (snd (action (shd s))) (datastate (shd s)) (output (shd s))) = trilean.true)" lemma alw_ev: "alw f = not (ev (\s. \f s))" by simp lemma alw_state_eq_smap: "alw (state_eq s) ss = alw (\ss. shd ss = s) (smap statename ss)" apply standard apply (simp add: alw_iff_sdrop ) by (simp add: alw_mono alw_smap ) subsection\Sink State\ text\Once the sink state is entered, it cannot be left and there are no outputs or updates henceforth.\ lemma shd_state_is_none: "(state_eq None) (make_full_observation e None r p t)" by (simp add: ) lemma unfold_observe_none: "make_full_observation e None d p t = (\statename = None, datastate = d, action=(shd t), output = p\##(make_full_observation e None d [] (stl t)))" by (simp add: stream.expand) lemma once_none_always_none_aux: assumes "\ p r i. j = (make_full_observation e None r p) i" shows "alw (state_eq None) j" using assms apply coinduct apply simp by fastforce lemma once_none_always_none: "alw (state_eq None) (make_full_observation e None r p t)" using once_none_always_none_aux by blast lemma once_none_nxt_always_none: "alw (nxt (state_eq None)) (make_full_observation e None r p t)" using once_none_always_none by (simp add: alw_iff_sdrop del: sdrop.simps) lemma snth_sconst: "(\i. s !! i = h) = (s = sconst h)" - by (metis funpow_code_def id_funpow sdrop_simps(1) sdrop_siterate siterate.simps(1) smap_alt smap_sconst snth.simps(1) stream.map_id) + by (auto simp add: sconst_alt sset_range) lemma alw_sconst: "(alw (\xs. shd xs = h) t) = (t = sconst h)" by (simp add: snth_sconst[symmetric] alw_iff_sdrop) lemma smap_statename_None: "smap statename (make_full_observation e None r p i) = sconst None" by (meson EFSM_LTL.alw_sconst alw_state_eq_smap once_none_always_none) lemma alw_not_some: "alw (\xs. statename (shd xs) \ Some s) (make_full_observation e None r p t)" by (metis (mono_tags, lifting) alw_mono once_none_always_none option.distinct(1) ) lemma state_none: "((state_eq None) impl nxt (state_eq None)) (make_full_observation e s r p t)" by (simp add: ) lemma state_none_2: "(state_eq None) (make_full_observation e s r p t) \ (state_eq None) (make_full_observation e s r p (stl t))" by (simp add: ) lemma no_output_none_aux: assumes "\ p r i. j = (make_full_observation e None r []) i" shows "alw (output_eq []) j" using assms apply coinduct apply simp by fastforce lemma no_output_none: "nxt (alw (output_eq [])) (make_full_observation e None r p t)" using no_output_none_aux by auto lemma nxt_alw: "nxt (alw P) s \ alw (nxt P) s" by (simp add: alw_iff_sdrop) lemma no_output_none_nxt: "alw (nxt (output_eq [])) (make_full_observation e None r p t)" using nxt_alw no_output_none by blast lemma no_output_none_if_empty: "alw (output_eq []) (make_full_observation e None r [] t)" by (metis (mono_tags, lifting) alw_nxt make_full_observation.simps(1) no_output_none state.select_convs(4)) lemma no_updates_none_aux: assumes "\ p i. j = (make_full_observation e None r p) i" shows "alw (\x. datastate (shd x) = r) j" using assms apply coinduct by fastforce lemma no_updates_none: "alw (\x. datastate (shd x) = r) (make_full_observation e None r p t)" using no_updates_none_aux by blast lemma action_components: "(label_eq l aand input_eq i) s = (action (shd s) = (String.implode l, i))" by (metis fst_conv prod.collapse snd_conv) end diff --git a/thys/Extended_Finite_State_Machines/VName.thy b/thys/Extended_Finite_State_Machines/VName.thy --- a/thys/Extended_Finite_State_Machines/VName.thy +++ b/thys/Extended_Finite_State_Machines/VName.thy @@ -1,44 +1,38 @@ section\Variables\ text\Variables can either be inputs or registers. Here we define the \texttt{vname} datatype which allows us to write expressions in terms of variables and case match during evaluation. We also make the \texttt{vname} datatype a member of linorder such that we can establish a linear order on arithmetic expressions, guards, and subsequently transitions.\ theory VName imports Main begin text_raw\\snip{vnametype}{1}{2}{%\ datatype vname = I nat | R nat text_raw\}%endsnip\ instantiation vname :: linorder begin text_raw\\snip{vnameorder}{1}{2}{%\ fun less_vname :: "vname \ vname \ bool" where "(I n1) < (R n2) = True" | "(R n1) < (I n2) = False" | "(I n1) < (I n2) = (n1 < n2)" | "(R n1) < (R n2) = (n1 < n2)" text_raw\}%endsnip\ definition less_eq_vname :: "vname \ vname \ bool" where "less_eq_vname v1 v2 = (v1 < v2 \ v1 = v2)" declare less_eq_vname_def [simp] instance apply standard - apply (metis (full_types) dual_order.asym less_eq_vname_def less_vname.simps(2) less_vname.simps(3) less_vname.simps(4) vname.exhaust) - apply simp + apply (auto elim: less_vname.elims) subgoal for x y z - apply (induct x y rule: less_vname.induct) - apply (metis less_eq_vname_def less_vname.elims(2) less_vname.elims(3) vname.simps(4)) - apply simp - apply (metis less_eq_vname_def less_trans less_vname.elims(3) less_vname.simps(3) vname.simps(4)) - by (metis le_less_trans less_eq_vname_def less_imp_le_nat less_vname.elims(2) less_vname.simps(4) vname.simps(4)) - apply (metis dual_order.asym less_eq_vname_def less_vname.elims(2) less_vname.simps(3) less_vname.simps(4)) - subgoal for x y - by (induct x y rule: less_vname.induct, auto) - done + apply (cases x; cases y; cases z) + apply simp_all + done + done end end diff --git a/thys/Graph_Theory/Funpow.thy b/thys/Graph_Theory/Auxiliary.thy rename from thys/Graph_Theory/Funpow.thy rename to thys/Graph_Theory/Auxiliary.thy --- a/thys/Graph_Theory/Funpow.thy +++ b/thys/Graph_Theory/Auxiliary.thy @@ -1,1051 +1,294 @@ -theory Funpow +theory Auxiliary imports "HOL-Library.FuncSet" - "HOL-Combinatorics.Permutations" + "HOL-Combinatorics.Orbits" begin -section \Auxiliary Lemmas about @{term "(^^)"}\ - -lemma funpow_simp_l: "f ((f ^^ n) x) = (f ^^ Suc n) x" - by (metis comp_apply funpow.simps(2)) - -lemma funpow_add_app: "(f ^^ n) ((f ^^ m) x) = (f ^^ (n + m)) x" - by (metis comp_apply funpow_add) - -lemma funpow_mod_eq: - assumes "(f ^^ n) x = x" "0 < n" shows "(f ^^ (m mod n)) x = (f ^^ m) x" -proof (induct m rule: less_induct) - case (less m) - { assume "m < n" then have ?case by simp } - moreover - { assume "m = n" then have ?case by (simp add: \_ = x\)} - moreover - { assume "n < m" - then have "m - n < m" "0 < m - n" using \0 < n\ by arith+ - - have "(f ^^ (m mod n)) x = (f ^^ ((m - n) mod n)) x" - using \0 < m - n\ by (simp add: mod_geq) - also have "\ = (f ^^ (m - n)) x" - using \m - n < m\ by (rule less) - also have "\ = (f ^^ (m - n)) ((f ^^ n) x)" - by (simp add: assms) - also have "\ = (f ^^ m) x" - using \0 < m - n\ by (simp add: funpow_add_app) - finally have ?case . } - ultimately show ?case by (metis linorder_neqE_nat) -qed - -lemma id_funpow_id: - assumes "f x = x" shows "(f ^^ n) x = x" - using assms by (induct n) auto - -lemma inv_id_abs[simp]: "inv (\a. a) = id" unfolding id_def[symmetric] by simp - -lemma inj_funpow: - fixes f :: "'a \ 'a" - assumes "inj f" shows "inj (f ^^ n)" -proof (induct n) - case 0 then show ?case by (auto simp: id_def[symmetric]) -next - case (Suc n) with assms show ?case unfolding funpow.simps by (rule inj_compose) -qed - -lemma funpow_inj_finite: - assumes "inj p" "finite {(p ^^ n) x |n. True}" - shows "\n>0. (p ^^ n) x = x" -proof - - have "\finite {0::nat..}" by simp - moreover - have "{(p ^^ n) x |n. True} = (\n. (p ^^ n) x) ` {0..}" by auto - with assms have "finite \" by simp - ultimately have "\n\{0..}. \ finite {m \ {0..}. (p ^^ m) x = (p ^^ n) x}" - by (rule pigeonhole_infinite) - then obtain n where "\finite {m. (p ^^ m) x = (p ^^ n) x}" by auto - then have "\finite ({m. (p ^^ m) x = (p ^^ n) x} - {n})" by auto - then have "({m. (p ^^ m) x = (p ^^ n) x} - {n}) \ {}" - by (metis finite.emptyI) - then obtain m where m: "(p ^^ m) x = (p ^^ n) x" "m \ n" by auto - - { fix m n assume "(p ^^ n) x = (p ^^ m) x" "m < n" - have "(p ^^ (n - m)) x = inv (p ^^ m) ((p ^^ m) ((p ^^ (n - m)) x))" - using \inj p\ by (simp add: inv_f_f inj_funpow) - also have "((p ^^ m) ((p ^^ (n - m)) x)) = (p ^^ n) x" - using \m < n\ by (simp add: funpow_add_app) - also have "inv (p ^^ m) \ = x" - using \inj p\ by (simp add: \(p ^^ n) x = _\ inj_funpow) - finally have "(p ^^ (n - m)) x = x" "0 < n - m" - using \m < n\ by auto } - note general = this - - show ?thesis - proof (cases m n rule: linorder_cases) - case less - then show ?thesis using general m by metis - next - case equal - then show ?thesis using m by metis - next - case greater - then show ?thesis using general m by metis - qed -qed - -lemma permutes_in_funpow_image: - assumes "f permutes S" "x \ S" - shows "(f ^^ n) x \ S" - using assms by (induct n) (auto simp: permutes_in_image) - -(* XXX move*) -lemma permutation_self: - assumes "permutation p" shows "\n>0. (p ^^ n) x = x" -proof cases - assume "p x = x" then show ?thesis by auto -next - assume "p x \ x" - from assms have "inj p" by (intro permutation_bijective bij_is_inj) - { fix n - from \p x \ x\ have "(p ^^ Suc n) x \ (p ^^ n) x" - proof (induct n arbitrary: x) - case 0 then show ?case by simp - next - case (Suc n) - have "p (p x) \ p x" - proof (rule notI) - assume "p (p x) = p x" - then show False using \p x \ x\ \inj p\ by (simp add: inj_eq) - qed - have "(p ^^ Suc (Suc n)) x = (p ^^ Suc n) (p x)" - by (metis funpow_simp_l funpow_swap1) - also have "\ \ (p ^^ n) (p x)" - by (rule Suc) fact - also have "(p ^^ n) (p x) = (p ^^ Suc n) x" - by (metis funpow_simp_l funpow_swap1) - finally show ?case by simp - qed } - then have "{(p ^^ n) x | n. True} \ {x. p x \ x}" - by auto - then have "finite {(p ^^ n) x | n. True}" - using permutation_finite_support[OF assms] by (rule finite_subset) - with \inj p\ show ?thesis by (rule funpow_inj_finite) -qed - -(* XXX move *) -lemma (in -) funpow_invs: +lemma funpow_invs: assumes "m \ n" and inv: "\x. f (g x) = x" shows "(f ^^ m) ((g ^^ n) x) = (g ^^ (n - m)) x" using \m \ n\ proof (induction m) case (Suc m) moreover then have "n - m = Suc (n - Suc m)" by auto ultimately show ?case by (auto simp: inv) qed simp - - -section \Function-power distance between values\ - -(* xxx move *) -definition funpow_dist :: "('a \ 'a) \ 'a \ 'a \ nat" where - "funpow_dist f x y \ LEAST n. (f ^^ n) x = y" - -abbreviation funpow_dist1 :: "('a \ 'a) \ 'a \ 'a \ nat" where - "funpow_dist1 f x y \ Suc (funpow_dist f (f x) y)" - -lemma funpow_dist_0: - assumes "x = y" shows "funpow_dist f x y = 0" - using assms unfolding funpow_dist_def by (intro Least_eq_0) simp - -lemma funpow_dist_least: - assumes "n < funpow_dist f x y" shows "(f ^^ n) x \ y" -proof (rule notI) - assume "(f ^^ n) x = y" - then have "funpow_dist f x y \ n" unfolding funpow_dist_def by (rule Least_le) - with assms show False by linarith -qed - -lemma funpow_dist1_least: - assumes "0 < n" "n < funpow_dist1 f x y" shows "(f ^^ n) x \ y" -proof (rule notI) - assume "(f ^^ n) x = y" - then have "(f ^^ (n - 1)) (f x) = y" - using \0 < n\ by (cases n) (simp_all add: funpow_swap1) - then have "funpow_dist f (f x) y \ n - 1" unfolding funpow_dist_def by (rule Least_le) - with assms show False by simp -qed - - -section \Cyclic Permutations\ - -inductive_set orbit :: "('a \ 'a) \ 'a \ 'a set" for f x where - base: "f x \ orbit f x" | - step: "y \ orbit f x \ f y \ orbit f x" - -definition cyclic_on :: "('a \ 'a) \ 'a set \ bool" where - "cyclic_on f S \ (\s\S. S = orbit f s)" - -lemma orbit_altdef: "orbit f x = {(f ^^ n) x | n. 0 < n}" (is "?L = ?R") -proof (intro set_eqI iffI) - fix y assume "y \ ?L" then show "y \ ?R" - by (induct rule: orbit.induct) (auto simp: exI[where x=1] exI[where x="Suc n" for n]) -next - fix y assume "y \ ?R" - then obtain n where "y = (f ^^ n) x" "0 < n" by blast - then show "y \ ?L" - proof (induction n arbitrary: y) - case (Suc n) then show ?case by (cases "n = 0") (auto intro: orbit.intros) - qed simp -qed - -lemma orbit_trans: - assumes "s \ orbit f t" "t \ orbit f u" shows "s \ orbit f u" - using assms by induct (auto intro: orbit.intros) - -lemma orbit_subset: - assumes "s \ orbit f (f t)" shows "s \ orbit f t" - using assms by (induct) (auto intro: orbit.intros) - -lemma orbit_sim_step: - assumes "s \ orbit f t" shows "f s \ orbit f (f t)" - using assms by induct (auto intro: orbit.intros) - -lemma orbit_step: - assumes "y \ orbit f x" "f x \ y" shows "y \ orbit f (f x)" - using assms -proof induction - case (step y) then show ?case by (cases "x = y") (auto intro: orbit.intros) -qed simp - -lemma self_in_orbit_trans: - assumes "s \ orbit f s" "t \ orbit f s" shows "t \ orbit f t" - using assms(2,1) by induct (auto intro: orbit_sim_step) - -lemma orbit_swap: - assumes "s \ orbit f s" "t \ orbit f s" shows "s \ orbit f t" - using assms(2,1) -proof induction - case base then show ?case by (cases "f s = s") (auto intro: orbit_step) -next - case (step x) then show ?case by (cases "f x = s") (auto intro: orbit_step) -qed - -lemma permutation_self_in_orbit: - assumes "permutation f" shows "s \ orbit f s" - unfolding orbit_altdef using permutation_self[OF assms, of s] by simp metis - -lemma orbit_altdef_self_in: - assumes "s \ orbit f s" shows "orbit f s = {(f ^^ n) s | n. True}" -proof (intro set_eqI iffI) - fix x assume "x \ {(f ^^ n) s | n. True}" - then obtain n where "x = (f ^^ n) s" by auto - then show "x \ orbit f s" using assms by (cases "n = 0") (auto simp: orbit_altdef) -qed (auto simp: orbit_altdef) - -lemma orbit_altdef_permutation: - assumes "permutation f" shows "orbit f s = {(f ^^ n) s | n. True}" - using assms by (intro orbit_altdef_self_in permutation_self_in_orbit) - -lemma orbit_altdef_bounded: - assumes "(f ^^ n) s = s" "0 < n" shows "orbit f s = {(f ^^ m) s| m. m < n}" -proof - - from assms have "s \ orbit f s" unfolding orbit_altdef by auto metis - then have "orbit f s = {(f ^^ m) s|m. True}" by (rule orbit_altdef_self_in) - also have "\ = {(f ^^ m) s| m. m < n}" - using assms by (auto simp: funpow_mod_eq intro: exI[where x="m mod n" for m]) - finally show ?thesis . -qed - -lemma funpow_in_orbit: - assumes "s \ orbit f t" shows "(f ^^ n) s \ orbit f t" - using assms by (induct n) (auto intro: orbit.intros) - -lemma finite_orbit: - assumes "s \ orbit f s" shows "finite (orbit f s)" -proof - - from assms obtain n where n: "0 < n" "(f ^^n) s = s" by (auto simp: orbit_altdef) - then show ?thesis by (auto simp: orbit_altdef_bounded) -qed - -lemma self_in_orbit_step: - assumes "s \ orbit f s" shows "orbit f (f s) = orbit f s" -proof (intro set_eqI iffI) - fix t assume "t \ orbit f s" then show "t \ orbit f (f s)" - using assms by (auto intro: orbit_step orbit_sim_step) -qed (auto intro: orbit_subset) - -lemma permutation_orbit_step: - assumes "permutation f" shows "orbit f (f s) = orbit f s" - using assms by (intro self_in_orbit_step permutation_self_in_orbit) - -lemma orbit_nonempty: - "orbit f s \ {}" - using orbit.base by fastforce - -lemma orbit_inv_eq: - assumes "permutation f" - shows "orbit (inv f) x = orbit f x" (is "?L = ?R") -proof - - { fix g y assume A: "permutation g" "y \ orbit (inv g) x" - have "y \ orbit g x" - proof - - have inv_g: "\y. x = g y \ inv g x = y" "\y. inv g (g y) = y" - by (metis A(1) bij_inv_eq_iff permutation_bijective)+ - - { fix y assume "y \ orbit g x" - then have "inv g y \ orbit g x" - by (cases) (simp_all add: inv_g A(1) permutation_self_in_orbit) - } note inv_g_in_orb = this - - from A(2) show ?thesis - by induct (simp_all add: inv_g_in_orb A permutation_self_in_orbit) - qed - } note orb_inv_ss = this - - have "inv (inv f) = f" - by (simp add: assms inv_inv_eq permutation_bijective) - then show ?thesis - using orb_inv_ss[OF assms] orb_inv_ss[OF permutation_inverse[OF assms]] by auto -qed - -lemma cyclic_on_alldef: - "cyclic_on f S \ S \ {} \ (\s\S. S = orbit f s)" - unfolding cyclic_on_def by (auto intro: orbit.step orbit_swap orbit_trans) - - -lemma cyclic_on_funpow_in: - assumes "cyclic_on f S" "s \ S" shows "(f^^n) s \ S" - using assms unfolding cyclic_on_def by (auto intro: funpow_in_orbit) - -lemma finite_cyclic_on: - assumes "cyclic_on f S" shows "finite S" - using assms by (auto simp: cyclic_on_def finite_orbit) - -lemma cyclic_on_singleI: - assumes "s \ S" "S = orbit f s" shows "cyclic_on f S" - using assms unfolding cyclic_on_def by blast - -lemma inj_on_funpow_least: - assumes "(f ^^ n) s = s" "\m. \m < n; 0 < m\ \ (f ^^ m) s \ s" - shows "inj_on (\k. (f^^k) s) {0.. l" "(f ^^ k) s = (f ^^ l) s" - define k' l' where "k' = min k l" and "l' = max k l" - with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" - by (auto simp: min_def max_def) - - have "s = (f ^^ ((n - l') + l')) s" using assms \l' < n\ by simp - also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) - also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') - also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) - finally have "(f ^^ (n - l' + k')) s = s" by simp - moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ - ultimately have False using assms(2) by auto - } - then show ?thesis by (intro inj_onI) auto -qed - -lemma cyclic_on_inI: - assumes "cyclic_on f S" "s \ S" shows "f s \ S" - using assms by (auto simp: cyclic_on_def intro: orbit.intros) - -lemma bij_betw_funpow: - assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" -proof (induct n) - case 0 then show ?case by (auto simp: id_def[symmetric]) -next - case (Suc n) - then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) -qed - -(*XXX rename move*) -lemma orbit_FOO: - assumes self:"a \ orbit g a" - and eq: "\x. x \ orbit g a \ g' (f x) = f (g x)" - shows "f ` orbit g a = orbit g' (f a)" (is "?L = ?R") -proof (intro set_eqI iffI) - fix x assume "x \ ?L" - then obtain x0 where "x0 \ orbit g a" "x = f x0" by auto - then show "x \ ?R" - proof (induct arbitrary: x) - case base then show ?case by (auto simp: self orbit.base eq[symmetric]) - next - case step then show ?case by cases (auto simp: eq[symmetric] orbit.intros) - qed -next - fix x assume "x \ ?R" - then show "x \ ?L" - proof (induct arbitrary: ) - case base then show ?case by (auto simp: self orbit.base eq) - next - case step then show ?case by cases (auto simp: eq orbit.intros) - qed -qed - -lemma cyclic_on_FOO: - assumes "cyclic_on f S" - assumes "\x. x \ S \ g (h x) = h (f x)" - shows "cyclic_on g (h ` S)" - using assms by (auto simp: cyclic_on_def) (meson orbit_FOO) - -lemma cyclic_on_f_in: - assumes "f permutes S" "cyclic_on f A" "f x \ A" - shows "x \ A" -proof - - from assms have fx_in_orb: "f x \ orbit f (f x)" by (auto simp: cyclic_on_alldef) - from assms have "A = orbit f (f x)" by (auto simp: cyclic_on_alldef) - moreover - then have "\ = orbit f x" using \f x \ A\ by (auto intro: orbit_step orbit_subset) - ultimately - show ?thesis by (metis (no_types) orbit.simps permutes_inverses(2)[OF assms(1)]) -qed - -lemma permutes_not_in: - assumes "f permutes S" "x \ S" shows "f x = x" - using assms by (auto simp: permutes_def) - -lemma orbit_cong0: - assumes "x \ A" "f \ A \ A" "\y. y \ A \ f y = g y" shows "orbit f x = orbit g x" -proof - - { fix n have "(f ^^ n) x = (g ^^ n) x \ (f ^^ n) x \ A" - by (induct n rule: nat.induct) (insert assms, auto) - } then show ?thesis by (auto simp: orbit_altdef) -qed - -lemma orbit_cong: - assumes self_in: "t \ orbit f t" and eq: "\s. s \ orbit f t \ g s = f s" - shows "orbit g t = orbit f t" - using assms(1) _ assms(2) by (rule orbit_cong0) (auto simp: orbit.step eq) - -lemma cyclic_cong: - assumes "\s. s \ S \ f s = g s" shows "cyclic_on f S = cyclic_on g S" -proof - - have "(\s\S. orbit f s = orbit g s) \ cyclic_on f S = cyclic_on g S" - by (metis cyclic_on_alldef cyclic_on_def) - then show ?thesis by (metis assms orbit_cong cyclic_on_def) -qed - -lemma permutes_comp_preserves_cyclic1: - assumes "g permutes B" "cyclic_on f C" - assumes "A \ B = {}" "C \ A" - shows "cyclic_on (f o g) C" -proof - - have *: "\c. c \ C \ f (g c) = f c" - using assms by (subst permutes_not_in[where f=g]) auto - with assms(2) show ?thesis by (simp cong: cyclic_cong) -qed - -lemma permutes_comp_preserves_cyclic2: - assumes "f permutes A" "cyclic_on g C" - assumes "A \ B = {}" "C \ B" - shows "cyclic_on (f o g) C" -proof - - obtain c where c: "c \ C" "C = orbit g c" "c \ orbit g c" - using \cyclic_on g C\ by (auto simp: cyclic_on_def) - then have "\c. c \ C \ f (g c) = g c" - using assms c by (subst permutes_not_in[where f=f]) (auto intro: orbit.intros) - with assms(2) show ?thesis by (simp cong: cyclic_cong) -qed - - -(*XXX merge with previous section?*) -subsection \Orbits\ - -lemma permutes_orbit_subset: - assumes "f permutes S" "x \ S" shows "orbit f x \ S" -proof - fix y assume "y \ orbit f x" - then show "y \ S" by induct (auto simp: permutes_in_image assms) -qed - -lemma cyclic_on_orbit': - assumes "permutation f" shows "cyclic_on f (orbit f x)" - unfolding cyclic_on_alldef using orbit_nonempty[of f x] - by (auto intro: assms orbit_swap orbit_trans permutation_self_in_orbit) - -(* XXX remove? *) -lemma cyclic_on_orbit: - assumes "f permutes S" "finite S" shows "cyclic_on f (orbit f x)" - using assms by (intro cyclic_on_orbit') (auto simp: permutation_permutes) - -lemma orbit_cyclic_eq3: - assumes "cyclic_on f S" "y \ S" shows "orbit f y = S" - using assms unfolding cyclic_on_alldef by simp - -(*XXX move*) -lemma orbit_eq_singleton_iff: "orbit f x = {x} \ f x = x" (is "?L \ ?R") -proof - assume A: ?R - { fix y assume "y \ orbit f x" then have "y = x" - by induct (auto simp: A) - } then show ?L by (metis orbit_nonempty singletonI subsetI subset_singletonD) -next - assume A: ?L - then have "\y. y \ orbit f x \ f x = y" - by - (erule orbit.cases, simp_all) - then show ?R using A by blast -qed - -(* XXX move *) -lemma eq_on_cyclic_on_iff1: - assumes "cyclic_on f S" "x \ S" - obtains "f x \ S" "f x = x \ card S = 1" -proof - from assms show "f x \ S" by (auto simp: cyclic_on_def intro: orbit.intros) - from assms have "S = orbit f x" by (auto simp: cyclic_on_alldef) - then have "f x = x \ S = {x}" by (metis orbit_eq_singleton_iff) - then show "f x = x \ card S = 1" using \x \ S\ by (auto simp: card_Suc_eq) -qed - - - - - - - -subsection \Decomposition of Arbitrary Permutations\ - -definition perm_restrict :: "('a \ 'a) \ 'a set \ ('a \ 'a)" where - "perm_restrict f S x \ if x \ S then f x else x" - -lemma perm_restrict_comp: - assumes "A \ B = {}" "cyclic_on f B" - shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" -proof - - have "\x. x \ B \ f x \ B" using \cyclic_on f B\ by (rule cyclic_on_inI) - with assms show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) -qed - -lemma perm_restrict_simps: - "x \ S \ perm_restrict f S x = f x" - "x \ S \ perm_restrict f S x = x" - by (auto simp: perm_restrict_def) - -lemma perm_restrict_perm_restrict: - "perm_restrict (perm_restrict f A) B = perm_restrict f (A \ B)" - by (auto simp: perm_restrict_def) - -lemma perm_restrict_union: - assumes "perm_restrict f A permutes A" "perm_restrict f B permutes B" "A \ B = {}" - shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" - using assms by (auto simp: fun_eq_iff perm_restrict_def permutes_def) (metis Diff_iff Diff_triv) - -lemma perm_restrict_id[simp]: - assumes "f permutes S" shows "perm_restrict f S = f" - using assms by (auto simp: permutes_def perm_restrict_def) - -lemma cyclic_on_perm_restrict: - "cyclic_on (perm_restrict f S) S \ cyclic_on f S" - by (simp add: perm_restrict_def cong: cyclic_cong) - -lemma perm_restrict_diff_cyclic: - assumes "f permutes S" "cyclic_on f A" - shows "perm_restrict f (S - A) permutes (S - A)" -proof - - { fix y - have "\x. perm_restrict f (S - A) x = y" - proof cases - assume A: "y \ S - A" - with \f permutes S\ obtain x where "f x = y" "x \ S" - unfolding permutes_def by auto metis - moreover - with A have "x \ A" by (metis Diff_iff assms(2) cyclic_on_inI) - ultimately - have "perm_restrict f (S - A) x = y" by (simp add: perm_restrict_simps) - then show ?thesis .. - next - assume "y \ S - A" - then have "perm_restrict f (S - A) y = y" by (simp add: perm_restrict_simps) - then show ?thesis .. - qed - } note X = this - - { fix x y assume "perm_restrict f (S - A) x = perm_restrict f (S - A) y" - with assms have "x = y" - by (auto simp: perm_restrict_def permutes_def split: if_splits intro: cyclic_on_f_in) - } note Y = this - - show ?thesis by (auto simp: permutes_def perm_restrict_simps X intro: Y) -qed - -lemma orbit_eqI: - "y = f x \ y \ orbit f x" - "z = f y \y \ orbit f x \z \ orbit f x" - by (metis orbit.base) (metis orbit.step) - -lemma permutes_decompose: - assumes "f permutes S" "finite S" - shows "\C. (\c \ C. cyclic_on f c) \ \C = S \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {})" - using assms(2,1) -proof (induction arbitrary: f rule: finite_psubset_induct) - case (psubset S) - - show ?case - proof (cases "S = {}") - case True then show ?thesis by (intro exI[where x="{}"]) auto - next - case False - then obtain s where "s \ S" by auto - with \f permutes S\ have "orbit f s \ S" - by (rule permutes_orbit_subset) - have cyclic_orbit: "cyclic_on f (orbit f s)" - using \f permutes S\ \finite S\ by (rule cyclic_on_orbit) - - let ?f' = "perm_restrict f (S - orbit f s)" - - have "f s \ S" using \f permutes S\ \s \ S\ by (auto simp: permutes_in_image) - then have "S - orbit f s \ S" using orbit.base[of f s] \s \ S\ by blast - moreover - have "?f' permutes (S - orbit f s)" - using \f permutes S\ cyclic_orbit by (rule perm_restrict_diff_cyclic) - ultimately - obtain C where C: "\c. c \ C \ cyclic_on ?f' c" "\C = S - orbit f s" - "\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}" - using psubset.IH by metis - - { fix c assume "c \ C" - then have *: "\x. x \ c \ perm_restrict f (S - orbit f s) x = f x" - using C(2) \f permutes S\ by (auto simp add: perm_restrict_def) - then have "cyclic_on f c" using C(1)[OF \c \ C\] by (simp cong: cyclic_cong add: *) - } note in_C_cyclic = this - - have Un_ins: "\(insert (orbit f s) C) = S" - using \\C = _\ \orbit f s \ S\ by blast - - have Disj_ins: "(\c1 \ insert (orbit f s) C. \c2 \ insert (orbit f s) C. c1 \ c2 \ c1 \ c2 = {})" - using C by auto - - show ?thesis - by (intro conjI Un_ins Disj_ins exI[where x="insert (orbit f s) C"]) - (auto simp: cyclic_orbit in_C_cyclic) - qed -qed - - -subsection \Funpow + Orbit\ - -lemma funpow_dist_prop: - "y \ orbit f x \ (f ^^ funpow_dist f x y) x = y" - unfolding funpow_dist_def by (rule LeastI_ex) (auto simp: orbit_altdef) - -lemma funpow_dist_0_eq: - assumes "y \ orbit f x" shows "funpow_dist f x y = 0 \ x = y" - using assms by (auto simp: funpow_dist_0 dest: funpow_dist_prop) - -lemma funpow_dist_step: - assumes "x \ y" "y \ orbit f x" shows "funpow_dist f x y = Suc (funpow_dist f (f x) y)" -proof - - from \y \ _\ obtain n where "(f ^^ n) x = y" by (auto simp: orbit_altdef) - with \x \ y\ obtain n' where [simp]: "n = Suc n'" by (cases n) auto - - show ?thesis - unfolding funpow_dist_def - proof (rule Least_Suc2) - show "(f ^^ n) x = y" by fact - then show "(f ^^ n') (f x) = y" by (simp add: funpow_swap1) - show "(f ^^ 0) x \ y" using \x \ y\ by simp - show "\k. ((f ^^ Suc k) x = y) = ((f ^^ k) (f x) = y)" - by (simp add: funpow_swap1) - qed -qed - -lemma funpow_dist1_prop: - assumes "y \ orbit f x" shows "(f ^^ funpow_dist1 f x y) x = y" - by (metis assms funpow.simps(1) funpow_dist_0 funpow_dist_prop funpow_simp_l funpow_swap1 id_apply orbit_step) - -(*XXX simplify? *) -lemma funpow_neq_less_funpow_dist: - assumes "y \ orbit f x" "m \ funpow_dist f x y" "n \ funpow_dist f x y" "m \ n" - shows "(f ^^ m) x \ (f ^^ n) x" -proof (rule notI) - assume A: "(f ^^ m) x = (f ^^ n) x" - - define m' n' where "m' = min m n" and "n' = max m n" - with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' \ funpow_dist f x y" - by (auto simp: min_def max_def) - - have "y = (f ^^ funpow_dist f x y) x" - using \y \ _\ by (simp only: funpow_dist_prop) - also have "\ = (f ^^ ((funpow_dist f x y - n') + n')) x" - using \n' \ _\ by simp - also have "\ = (f ^^ ((funpow_dist f x y - n') + m')) x" - by (simp add: funpow_add \(f ^^ m') x = _\) - also have "(f ^^ ((funpow_dist f x y - n') + m')) x \ y" - using A' by (intro funpow_dist_least) linarith - finally show "False" by simp -qed - -(* XXX reduce to funpow_neq_less_funpow_dist? *) -lemma funpow_neq_less_funpow_dist1: - assumes "y \ orbit f x" "m < funpow_dist1 f x y" "n < funpow_dist1 f x y" "m \ n" - shows "(f ^^ m) x \ (f ^^ n) x" -proof (rule notI) - assume A: "(f ^^ m) x = (f ^^ n) x" - - define m' n' where "m' = min m n" and "n' = max m n" - with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' < funpow_dist1 f x y" - by (auto simp: min_def max_def) - - have "y = (f ^^ funpow_dist1 f x y) x" - using \y \ _\ by (simp only: funpow_dist1_prop) - also have "\ = (f ^^ ((funpow_dist1 f x y - n') + n')) x" - using \n' < _\ by simp - also have "\ = (f ^^ ((funpow_dist1 f x y - n') + m')) x" - by (simp add: funpow_add \(f ^^ m') x = _\) - also have "(f ^^ ((funpow_dist1 f x y - n') + m')) x \ y" - using A' by (intro funpow_dist1_least) linarith+ - finally show "False" by simp -qed - -lemma inj_on_funpow_dist: - assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0..funpow_dist f x y}" - using funpow_neq_less_funpow_dist[OF assms] by (intro inj_onI) auto - -lemma inj_on_funpow_dist1: - assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0.. orbit f x" - shows "orbit f x = (\n. (f ^^ n) x) ` {0.. orbit f x" by (auto simp: orbit_altdef) - then show ?thesis by (rule funpow_dist1_prop) -qed - -lemma funpow_dist1_dist: - assumes "funpow_dist1 f x y < funpow_dist1 f x z" - assumes "{y,z} \ orbit f x" - shows "funpow_dist1 f x z = funpow_dist1 f x y + funpow_dist1 f y z" (is "?L = ?R") -proof - - have x_z: "(f ^^ funpow_dist1 f x z) x = z" using assms by (blast intro: funpow_dist1_prop) - have x_y: "(f ^^ funpow_dist1 f x y) x = y" using assms by (blast intro: funpow_dist1_prop) - - have "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y - = (f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) ((f ^^ funpow_dist1 f x y) x)" - using x_y by simp - also have "\ = z" - using assms x_z by (simp del: funpow.simps add: funpow_add_app) - finally have y_z_diff: "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y = z" . - then have "(f ^^ funpow_dist1 f y z) y = z" - using assms by (intro funpow_dist1_prop1) auto - then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" - using x_y by simp - then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" - by (simp del: funpow.simps add: funpow_add_app) - - show ?thesis - proof (rule antisym) - from y_z_diff have "(f ^^ funpow_dist1 f y z) y = z" - using assms by (intro funpow_dist1_prop1) auto - then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" - using x_y by simp - then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" - by (simp del: funpow.simps add: funpow_add_app) - then have "funpow_dist1 f x z \ funpow_dist1 f y z + funpow_dist1 f x y" - using funpow_dist1_least not_less by fastforce - then show "?L \ ?R" by presburger - next - have "funpow_dist1 f y z \ funpow_dist1 f x z - funpow_dist1 f x y" - using y_z_diff assms(1) by (metis not_less zero_less_diff funpow_dist1_least) - then show "?R \ ?L" by linarith - qed -qed - -lemma funpow_dist1_le_self: - assumes "(f ^^ m) x = x" "0 < m" "y \ orbit f x" - shows "funpow_dist1 f x y \ m" -proof (cases "x = y") - case True with assms show ?thesis by (auto dest!: funpow_dist1_least) -next - case False - have "(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x" - using assms by (simp add: funpow_mod_eq) - with False \y \ orbit f x\ have "funpow_dist1 f x y \ funpow_dist1 f x y mod m" - by auto (metis funpow_dist_least funpow_dist_prop funpow_dist_step funpow_simp_l not_less) - with \m > 0\ show ?thesis - by (auto intro: order_trans) -qed - - -subsection \Permutation Domains\ +section \Permutation Domains\ definition has_dom :: "('a \ 'a) \ 'a set \ bool" where "has_dom f S \ \s. s \ S \ f s = s" lemma permutes_conv_has_dom: "f permutes S \ bij f \ has_dom f S" by (auto simp: permutes_def has_dom_def bij_iff) - section \Segments\ inductive_set segment :: "('a \ 'a) \ 'a \ 'a \ 'a set" for f a b where base: "f a \ b \ f a \ segment f a b" | step: "x \ segment f a b \ f x \ b \ f x \ segment f a b" lemma segment_step_2D: assumes "x \ segment f a (f b)" shows "x \ segment f a b \ x = b" using assms by induct (auto intro: segment.intros) lemma not_in_segment2D: assumes "x \ segment f a b" shows "x \ b" using assms by induct auto lemma segment_altdef: assumes "b \ orbit f a" shows "segment f a b = (\n. (f ^^ n) a) ` {1.. ?L" have "f a \b \ b \ orbit f (f a)" using assms by (simp add: orbit_step) then have *: "f a \ b \ 0 < funpow_dist f (f a) b" using assms using gr0I funpow_dist_0_eq[OF \_ \ b \ orbit f (f a)\] by (simp add: orbit.intros) from \x \ ?L\ show "x \ ?R" proof induct case base then show ?case by (intro image_eqI[where x=1]) (auto simp: *) next case step then show ?case using assms funpow_dist1_prop less_antisym by (fastforce intro!: image_eqI[where x="Suc n" for n]) qed next fix x assume "x \ ?R" then obtain n where "(f ^^ n ) a = x" "0 < n" "n < funpow_dist1 f a b" by auto then show "x \ ?L" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have "(f ^^ Suc n) a \ b" using Suc by (meson funpow_dist1_least) with Suc show ?case by (cases "n = 0") (auto intro: segment.intros) qed qed (*XXX move up*) lemma segmentD_orbit: assumes "x \ segment f y z" shows "x \ orbit f y" using assms by induct (auto intro: orbit.intros) lemma segment1_empty: "segment f x (f x) = {}" by (auto simp: segment_altdef orbit.base funpow_dist_0) lemma segment_subset: assumes "y \ segment f x z" assumes "w \ segment f x y" shows "w \ segment f x z" using assms by (induct arbitrary: w) (auto simp: segment1_empty intro: segment.intros dest: segment_step_2D elim: segment.cases) (* XXX move up*) lemma not_in_segment1: assumes "y \ orbit f x" shows "x \ segment f x y" proof assume "x \ segment f x y" then obtain n where n: "0 < n" "n < funpow_dist1 f x y" "(f ^^ n) x = x" using assms by (auto simp: segment_altdef Suc_le_eq) then have neq_y: "(f ^^ (funpow_dist1 f x y - n)) x \ y" by (simp add: funpow_dist1_least) have "(f ^^ (funpow_dist1 f x y - n)) x = (f ^^ (funpow_dist1 f x y - n)) ((f ^^ n) x)" - using n by (simp add: funpow_add_app) + using n by (simp add: funpow_add) also have "\ = (f ^^ funpow_dist1 f x y) x" - using \n < _\ by (simp add: funpow_add_app) + using \n < _\ by (simp add: funpow_add) + (metis assms funpow_0 funpow_neq_less_funpow_dist1 n(1) n(3) nat_neq_iff zero_less_Suc) also have "\ = y" using assms by (rule funpow_dist1_prop) finally show False using neq_y by contradiction qed lemma not_in_segment2: "y \ segment f x y" using not_in_segment2D by metis (*XXX move*) lemma in_segmentE: assumes "y \ segment f x z" "z \ orbit f x" obtains "(f ^^ funpow_dist1 f x y) x = y" "funpow_dist1 f x y < funpow_dist1 f x z" proof from assms show "(f ^^ funpow_dist1 f x y) x = y" by (intro segmentD_orbit funpow_dist1_prop) moreover obtain n where "(f ^^ n) x = y" "0 < n" "n < funpow_dist1 f x z" using assms by (auto simp: segment_altdef) moreover then have "funpow_dist1 f x y \ n" by (meson funpow_dist1_least not_less) ultimately show "funpow_dist1 f x y < funpow_dist1 f x z" by linarith qed (*XXX move*) lemma cyclic_split_segment: assumes S: "cyclic_on f S" "a \ S" "b \ S" and "a \ b" shows "S = {a,b} \ segment f a b \ segment f b a" (is "?L = ?R") proof (intro set_eqI iffI) fix c assume "c \ ?L" with S have "c \ orbit f a" unfolding cyclic_on_alldef by auto then show "c \ ?R" by induct (auto intro: segment.intros) next fix c assume "c \ ?R" moreover have "segment f a b \ orbit f a" "segment f b a \ orbit f b" by (auto dest: segmentD_orbit) ultimately show "c \ ?L" using S by (auto simp: cyclic_on_alldef) qed (*XXX move*) lemma segment_split: assumes y_in_seg: "y \ segment f x z" shows "segment f x z = segment f x y \ {y} \ segment f y z" (is "?L = ?R") proof (intro set_eqI iffI) fix w assume "w \ ?L" then show "w \ ?R" by induct (auto intro: segment.intros) next fix w assume "w \ ?R" moreover { assume "w \ segment f x y" then have "w \ segment f x z" using segment_subset[OF y_in_seg] by auto } moreover { assume "w \ segment f y z" then have "w \ segment f x z" using y_in_seg by induct (auto intro: segment.intros) } ultimately show "w \ ?L" using y_in_seg by (auto intro: segment.intros) qed lemma in_segmentD_inv: assumes "x \ segment f a b" "x \ f a" assumes "inj f" shows "inv f x \ segment f a b" using assms by (auto elim: segment.cases) lemma in_orbit_invI: assumes "b \ orbit f a" assumes "inj f" shows "a \ orbit (inv f) b" using assms(1) apply induct apply (simp add: assms(2) orbit_eqI(1)) by (metis assms(2) inv_f_f orbit.base orbit_trans) lemma segment_step_2: assumes A: "x \ segment f a b" "b \ a" and "inj f" shows "x \ segment f a (f b)" using A by induct (auto intro: segment.intros dest: not_in_segment2D injD[OF \inj f\]) lemma inv_end_in_segment: assumes "b \ orbit f a" "f a \ b" "bij f" shows "inv f b \ segment f a b" using assms(1,2) proof induct case base then show ?case by simp next case (step x) moreover from \bij f\ have "inj f" by (rule bij_is_inj) moreover then have "x \ f x \ f a = x \ x \ segment f a (f x)" by (meson segment.simps) moreover have "x \ f x" using step \inj f\ by (metis in_orbit_invI inv_f_eq not_in_segment1 segment.base) then have "inv f x \ segment f a (f x) \ x \ segment f a (f x)" using \bij f\ \inj f\ by (auto dest: segment.step simp: surj_f_inv_f bij_is_surj) then have "inv f x \ segment f a x \ x \ segment f a (f x)" using \f a \ f x\ \inj f\ by (auto dest: segment_step_2 injD) ultimately show ?case by (cases "f a = x") simp_all qed lemma segment_overlapping: assumes "x \ orbit f a" "x \ orbit f b" "bij f" shows "segment f a x \ segment f b x \ segment f b x \ segment f a x" using assms(1,2) proof induction case base then show ?case by (simp add: segment1_empty) next case (step x) from \bij f\ have "inj f" by (simp add: bij_is_inj) have *: "\f x y. y \ segment f x (f x) \ False" by (simp add: segment1_empty) { fix y z assume A: "y \ segment f b (f x)" "y \ segment f a (f x)" "z \ segment f a (f x)" from \x \ orbit f a\ \f x \ orbit f b\ \y \ segment f b (f x)\ have "x \ orbit f b" by (metis * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq[OF \inj f\] segmentD_orbit) moreover with \x \ orbit f a\ step.IH have "segment f a (f x) \ segment f b (f x) \ segment f b (f x) \ segment f a (f x)" apply auto apply (metis * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq[OF \inj f\] segment_step_2D segment_subset step.prems subsetCE) by (metis (no_types, lifting) \inj f\ * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq orbit_eqI(2) segment_step_2D segment_subset subsetCE) ultimately have "segment f a (f x) \ segment f b (f x)" using A by auto } note C = this then show ?case by auto qed lemma segment_disj: assumes "a \ b" "bij f" shows "segment f a b \ segment f b a = {}" proof (rule ccontr) assume "\?thesis" then obtain x where x: "x \ segment f a b" "x \ segment f b a" by blast then have "segment f a b = segment f a x \ {x} \ segment f x b" "segment f b a = segment f b x \ {x} \ segment f x a" by (auto dest: segment_split) then have o: "x \ orbit f a" "x \ orbit f b" by (auto dest: segmentD_orbit) note * = segment_overlapping[OF o \bij f\] have "inj f" using \bij f\ by (simp add: bij_is_inj) have "segment f a x = segment f b x" proof (intro set_eqI iffI) fix y assume A: "y \ segment f b x" then have "y \ segment f a x \ f a \ segment f b a" using * x(2) by (auto intro: segment.base segment_subset) then show "y \ segment f a x" using \inj f\ A by (metis (no_types) not_in_segment2 segment_step_2) next fix y assume A: "y \ segment f a x " then have "y \ segment f b x \ f b \ segment f a b" using * x(1) by (auto intro: segment.base segment_subset) then show "y \ segment f b x" using \inj f\ A by (metis (no_types) not_in_segment2 segment_step_2) qed moreover have "segment f a x \ segment f b x" by (metis assms bij_is_inj not_in_segment2 segment.base segment_step_2 segment_subset x(1)) ultimately show False by contradiction qed lemma segment_x_x_eq: assumes "permutation f" shows "segment f x x = orbit f x - {x}" (is "?L = ?R") proof (intro set_eqI iffI) fix y assume "y \ ?L" then show "y \ ?R" by (auto dest: segmentD_orbit simp: not_in_segment2) next fix y assume "y \ ?R" then have "y \ orbit f x" "y \ x" by auto then show "y \ ?L" by induct (auto intro: segment.intros) qed section \Lists of Powers\ definition iterate :: "nat \ nat \ ('a \ 'a ) \ 'a \ 'a list" where "iterate m n f x = map (\n. (f^^n) x) [m..k. (f ^^ k) x) ` {m.. m \ n" by (auto simp: iterate_def) lemma iterate_length[simp]: "length (iterate m n f x) = n - m" by (auto simp: iterate_def) lemma iterate_nth[simp]: assumes "k < n - m" shows "iterate m n f x ! k = (f^^(m+k)) x" using assms by (induct k arbitrary: m) (auto simp: iterate_def) lemma iterate_applied: "iterate n m f (f x) = iterate (Suc n) (Suc m) f x" by (induct m arbitrary: n) (auto simp: iterate_def funpow_swap1) end diff --git a/thys/Graph_Theory/Graph_Theory.thy b/thys/Graph_Theory/Graph_Theory.thy --- a/thys/Graph_Theory/Graph_Theory.thy +++ b/thys/Graph_Theory/Graph_Theory.thy @@ -1,24 +1,24 @@ - (* Title: Graph_Theory.thy +(* Title: Graph_Theory.thy Author: Lars Noschinski, TU München *) theory Graph_Theory imports Digraph Bidirected_Digraph Arc_Walk Digraph_Component Digraph_Component_Vwalk Digraph_Isomorphism Pair_Digraph Vertex_Walk Subdivision Euler Kuratowski Shortest_Path begin end diff --git a/thys/Graph_Theory/Subdivision.thy b/thys/Graph_Theory/Subdivision.thy --- a/thys/Graph_Theory/Subdivision.thy +++ b/thys/Graph_Theory/Subdivision.thy @@ -1,430 +1,430 @@ theory Subdivision imports Arc_Walk Digraph_Component Pair_Digraph Bidirected_Digraph - Funpow + Auxiliary begin section \Subdivision on Digraphs\ definition subdivision_step :: "('a, 'b) pre_digraph \ ('b \ 'b) \ ('a, 'b) pre_digraph \ ('b \ 'b) \ 'a \ 'a \ 'a \ 'b \ 'b \ 'b \ bool" where "subdivision_step G rev_G H rev_H \ \(u, v, w) (uv, uw, vw). bidirected_digraph G rev_G \ bidirected_digraph H rev_H \ perm_restrict rev_H (arcs G) = perm_restrict rev_G (arcs H) \ compatible G H \ verts H = verts G \ {w} \ w \ verts G \ arcs H = {uw, rev_H uw, vw, rev_H vw} \ arcs G - {uv, rev_G uv} \ uv \ arcs G \ distinct [uw, rev_H uw, vw, rev_H vw] \ arc_to_ends G uv = (u,v) \ arc_to_ends H uw = (u,w) \ arc_to_ends H vw = (v,w) " inductive subdivision :: "('a,'b) pre_digraph \ ('b \ 'b) \ ('a,'b) pre_digraph \ ('b \ 'b) \ bool" for biG where base: "bidirected_digraph (fst biG) (snd biG) \ subdivision biG biG" | divide: "\subdivision biG biI; subdivision_step (fst biI) (snd biI) (fst biH) (snd biH) (u,v,w) (uv,uw,vw)\ \ subdivision biG biH" lemma subdivision_induct[case_names base divide, induct pred: subdivision]: assumes "subdivision (G, rev_G) (H, rev_H)" and "bidirected_digraph G rev_G \ P G rev_G" and "\I rev_I H rev_H u v w uv uw vw. subdivision (G, rev_G) (I, rev_I) \ P I rev_I \ subdivision_step I rev_I H rev_H (u, v, w) (uv, uw, vw) \ P H rev_H" shows "P H rev_H" using assms(1) by (induct biH\"(H, rev_H)" arbitrary: H rev_H) (auto intro: assms(2,3)) lemma subdivision_base: "bidirected_digraph G rev_G \ subdivision (G, rev_G) (G, rev_G)" by (rule subdivision.base) simp lemma subdivision_step_rev: assumes "subdivision_step G rev_G H rev_H (u, v, w) (uv, uw, vw)" "subdivision (H, rev_H) (I, rev_I)" shows "subdivision (G, rev_G) (I, rev_I)" proof - have "bidirected_digraph (fst (G, rev_G)) (snd (G, rev_G))" using assms by (auto simp: subdivision_step_def) with assms(2,1) show ?thesis using assms(2,1) by induct (auto intro: subdivision.intros dest: subdivision_base) qed lemma subdivision_trans: assumes "subdivision (G, rev_G) (H, rev_H)" "subdivision (H, rev_H) (I, rev_I)" shows "subdivision (G, rev_G) (I, rev_I)" using assms by induction (auto intro: subdivision_step_rev) locale subdiv_step = fixes G rev_G H rev_H u v w uv uw vw assumes subdiv_step: "subdivision_step G rev_G H rev_H (u, v, w) (uv, uw, vw)" sublocale subdiv_step \ G: bidirected_digraph G rev_G using subdiv_step unfolding subdivision_step_def by simp sublocale subdiv_step \ H: bidirected_digraph H rev_H using subdiv_step unfolding subdivision_step_def by simp context subdiv_step begin (*XXX*) abbreviation (input) "vu \ rev_G uv" abbreviation (input) "wu \ rev_H uw" abbreviation (input) "wv \ rev_H vw" lemma subdiv_compat: "compatible G H" using subdiv_step by (simp add: subdivision_step_def) lemma arc_to_ends_eq: "arc_to_ends H = arc_to_ends G" using subdiv_compat by (simp add: compatible_def arc_to_ends_def fun_eq_iff) lemma head_eq: "head H = head G" using subdiv_compat by (simp add: compatible_def fun_eq_iff) lemma tail_eq: "tail H = tail G" using subdiv_compat by (simp add: compatible_def fun_eq_iff) lemma verts_H: "verts H = verts G \ {w}" using subdiv_step by (simp add: subdivision_step_def) lemma verts_G: "verts G = verts H - {w}" using subdiv_step by (auto simp: subdivision_step_def) lemma arcs_H: "arcs H = {uw, wu, vw, wv} \ arcs G - {uv, vu}" using subdiv_step by (simp add: subdivision_step_def) lemma not_in_verts_G: "w \ verts G" using subdiv_step by (simp add: subdivision_step_def) lemma in_arcs_G: "{uv, vu} \ arcs G" using subdiv_step by (simp add: subdivision_step_def) lemma not_in_arcs_H: "{uv,vu} \ arcs H = {}" using arcs_H by auto lemma subdiv_ate: "arc_to_ends G uv = (u,v)" "arc_to_ends H uv = (u,v)" "arc_to_ends H uw = (u,w)" "arc_to_ends H vw = (v,w)" using subdiv_step subdiv_compat by (auto simp: subdivision_step_def arc_to_ends_def compatible_def) lemma subdiv_ends[simp]: "tail G uv = u" "head G uv = v" "tail H uv = u" "head H uv = v" "tail H uw = u" "head H uw = w" "tail H vw = v" "head H vw = w" using subdiv_ate by (auto simp: arc_to_ends_def) lemma subdiv_ends_G_rev[simp]: "tail G (vu) = v" "head G (vu) = u" "tail H (vu) = v" "head H (vu) = u" using in_arcs_G by (auto simp: tail_eq head_eq) lemma subdiv_distinct_verts0: "u \ w" "v \ w" using in_arcs_G not_in_verts_G using subdiv_ate by (auto simp: arc_to_ends_def dest: G.wellformed) lemma in_arcs_H: "{uw, wu, vw, wv} \ arcs H" proof - { assume "uv = uw" then have "arc_to_ends H uv = arc_to_ends H uw" by simp then have "v = w" by (simp add: arc_to_ends_def) } moreover { assume "uv = vw" then have "arc_to_ends H uv = arc_to_ends H vw" by simp then have "v = w" by (simp add: arc_to_ends_def) } moreover { assume "vu = uw" then have "arc_to_ends H (vu) = arc_to_ends H uw" by simp then have "u = w" by (simp add: arc_to_ends_def) } moreover { assume "vu = vw" then have "arc_to_ends H (vu) = arc_to_ends H vw" by simp then have "u = w" by (simp add: arc_to_ends_def) } ultimately have "{uw,vw} \ arcs H" unfolding arcs_H using subdiv_distinct_verts0 by auto then show ?thesis by auto qed lemma subdiv_ends_H_rev[simp]: "tail H (wu) = w" "tail H (wv) = w" "head H (wu) = u" "head H (wv) = v" using in_arcs_H subdiv_ate by simp_all lemma in_verts_G: "{u,v} \ verts G" using in_arcs_G by (auto dest: G.wellformed) lemma not_in_arcs_G: "{uw, wu, vw, wv} \ arcs G = {}" proof - note X = G.wellformed[simplified tail_eq[symmetric] head_eq[symmetric]] show ?thesis using not_in_verts_G in_arcs_H by (auto dest: X ) qed lemma subdiv_distinct_arcs: "distinct [uv, vu, uw, wu, vw, wv]" proof - have "distinct [uw, wu, vw, wv]" using subdiv_step by (simp add: subdivision_step_def) moreover have "distinct [uv, vu]" using in_arcs_G G.arev_dom by auto moreover have "{uv, vu} \ {uw, wu, vw, wv} = {}" using arcs_H in_arcs_H by auto ultimately show ?thesis by auto qed lemma arcs_G: "arcs G = arcs H \ {uv, vu} - {uw, wu, vw, wv}" using in_arcs_G not_in_arcs_G unfolding arcs_H by auto lemma subdiv_ate_H_rev: "arc_to_ends H (wu) = (w,u)" "arc_to_ends H (wv) = (w,v)" using in_arcs_H subdiv_ate by simp_all lemma adj_with_w: "u \\<^bsub>H\<^esub> w" "w \\<^bsub>H\<^esub> u" "v \\<^bsub>H\<^esub> w" "w \\<^bsub>H\<^esub> v" using in_arcs_H subdiv_ate by (auto intro: H.dominatesI[rotated]) lemma w_reach: "u \\<^sup>*\<^bsub>H\<^esub> w" "w \\<^sup>*\<^bsub>H\<^esub> u" "v \\<^sup>*\<^bsub>H\<^esub> w" "w \\<^sup>*\<^bsub>H\<^esub> v" using adj_with_w by auto lemma G_reach: "v \\<^sup>*\<^bsub>G\<^esub> u" "u \\<^sup>*\<^bsub>G\<^esub> v" using subdiv_ate in_arcs_G by (simp add: G.dominatesI G.symmetric_reachable')+ lemma out_arcs_w: "out_arcs H w = {wu, wv}" using subdiv_distinct_verts0 in_arcs_H by (auto simp: arcs_H) (auto simp: tail_eq verts_G dest: G.tail_in_verts) lemma out_degree_w: "out_degree H w = 2" using subdiv_distinct_arcs by (auto simp: out_degree_def out_arcs_w card_insert_if) end lemma subdivision_compatible: assumes "subdivision (G, rev_G) (H, rev_H)" shows "compatible G H" using assms by induct (auto simp: compatible_def subdivision_step_def) lemma subdivision_bidir: assumes "subdivision (G, rev_G) (H, rev_H)" shows "bidirected_digraph H rev_H" using assms by induct (auto simp: subdivision_step_def) lemma subdivision_choose_rev: assumes "subdivision (G, rev_G) (H, rev_H)" "bidirected_digraph H rev_H'" shows "\rev_G'. subdivision (G, rev_G') (H, rev_H')" using assms proof (induction arbitrary: rev_H') case base then show ?case by (auto dest: subdivision_base) next case (divide I rev_I H rev_H u v w uv uw vw) interpret subdiv_step I rev_I H rev_H u v w uv uw vw using divide by unfold_locales interpret H': bidirected_digraph H rev_H' by fact define rev_I' where "rev_I' x = (if x = uv then rev_I uv else if x = rev_I uv then uv else if x \ arcs I then rev_H' x else x)" for x have rev_H_injD: "\x y z. rev_H' x = z \ rev_H' y = z \ x \ y \ False" by (metis H'.arev_eq_iff) have rev_H'_simps: "rev_H' uw = rev_H uw \ rev_H' vw = rev_H vw \ rev_H' uw = rev_H vw \ rev_H' vw = rev_H uw" proof - have "arc_to_ends H (rev_H' uw) = (w,u)" "arc_to_ends H (rev_H' vw) = (w,v)" using in_arcs_H by (auto simp: subdiv_ate) moreover have "\x. x \ arcs H \ tail H x = w \ x \ {rev_H uw, rev_H vw}" using subdiv_distinct_verts0 not_in_verts_G by (auto simp: arcs_H) (simp add: tail_eq) ultimately have "rev_H' uw \ {rev_H uw, rev_H vw}" "rev_H' vw \ {rev_H uw, rev_H vw}" using in_arcs_H by auto then show ?thesis using in_arcs_H by (auto dest: rev_H_injD) qed have rev_H_uv: "rev_H' uv = uv" "rev_H' (rev_I uv) = rev_I uv" using not_in_arcs_H by (auto simp: H'.arev_eq) have bd_I': "bidirected_digraph I rev_I'" proof fix a have "\a. a \ uv \ a \ rev_I uv \ a \ arcs I \ a \ arcs H" by (auto simp: arcs_H) then show "(a \ arcs I) = (rev_I' a \ a)" using in_arcs_G by (auto simp: rev_I'_def dest: G.arev_neq H'.arev_neq) next fix a have *: "\a. rev_H' a = rev_I uv \ a = rev_I uv" by (metis H'.arev_arev H'.arev_dom insert_disjoint(1) not_in_arcs_H) have **: "\a. uv = rev_H' a \ a = uv" using H'.arev_eq not_in_arcs_H by force have ***: "\a. a \ arcs I \ rev_H' a \ arcs I" using rev_H'_simps by (case_tac "a \ {uv,vu}") (fastforce simp: rev_H_uv, auto simp: arcs_G dest: rev_H_injD) show "rev_I' (rev_I' a) = a" by (auto simp: rev_I'_def H'.arev_eq rev_H_uv * ** ***) next fix a assume "a \ arcs I" then show "tail I (rev_I' a) = head I a" using in_arcs_G by (auto simp: rev_I'_def tail_eq[symmetric] head_eq[symmetric] arcs_H) qed moreover have "\x. rev_H' x = uv \ x = uv" "\x. rev_H' x = rev_I uv \ x = rev_I uv" using not_in_arcs_H by (auto dest: H'.arev_eq) (metis H'.arev_arev H'.arev_eq) then have "perm_restrict rev_H' (arcs I) = perm_restrict rev_I' (arcs H)" using not_in_arcs_H by (auto simp: rev_I'_def perm_restrict_def H'.arev_eq) ultimately have sds_I'H': "subdivision_step I rev_I' H rev_H' (u, v, w) (uv, uw, vw)" using divide(2,4) rev_H'_simps unfolding subdivision_step_def by (fastforce simp: rev_I'_def) then have "subdivision (I, rev_I') (H, rev_H')" "\rev_G'. subdivision (G, rev_G') (I, rev_I')" using bd_I' divide by (auto intro: subdivision.intros dest: subdivision_base) then show ?case by (blast intro: subdivision_trans) qed lemma subdivision_verts_subset: assumes "subdivision (G,rev_G) (H,rev_H)" "x \ verts G" shows "x \ verts H" using assms by induct (auto simp: subdiv_step.verts_H subdiv_step_def) subsection \Subdivision on Pair Digraphs\ text \ In this section, we introduce specialized rules for pair digraphs. \ abbreviation "subdivision_pair G H \ subdivision (with_proj G, swap_in (parcs G)) (with_proj H, swap_in (parcs H))" lemma arc_to_ends_with_proj[simp]: "arc_to_ends (with_proj G) = id" by (auto simp: arc_to_ends_def) context begin text \ We use the @{verbatim inductive} command to define an inductive definition pair graphs. This is proven to be equivalent to @{const subdivision}. This allows us to transfer the rules proven by @{verbatim inductive} to @{const subdivision}. To spare the user confusion, we hide this new constant. \ private inductive pair_sd :: "'a pair_pre_digraph \ 'a pair_pre_digraph \ bool" for G where base: "pair_bidirected_digraph G \ pair_sd G G" | divide: "\e w H. \e \ parcs H; w \ pverts H; pair_sd G H\ \ pair_sd G (subdivide H e w)" private lemma bidirected_digraphI_pair_sd: assumes "pair_sd G H" shows "pair_bidirected_digraph H" using assms proof induct case base then show ?case by auto next case (divide e w H) interpret H: pair_bidirected_digraph H by fact from divide show ?case by (intro H.pair_bidirected_digraph_subdivide) qed private lemma subdivision_with_projI: assumes "pair_sd G H" shows "subdivision_pair G H" using assms proof induct case base then show ?case by (blast intro: pair_bidirected_digraph.bidirected_digraph subdivision_base) next case (divide e w H) obtain u v where "e = (u,v)" by (cases e) interpret H: pair_bidirected_digraph H using divide(3) by (rule bidirected_digraphI_pair_sd) interpret I: pair_bidirected_digraph "subdivide H e w" using divide(1,2) by (rule H.pair_bidirected_digraph_subdivide) have uvw: "u \ v" "u \ w" "v \ w" using divide by (auto simp: \e = _\ dest: H.adj_not_same H.wellformed) have "subdivision (with_proj G, swap_in (parcs G)) (H, swap_in (parcs H))" using divide by auto moreover have *: "perm_restrict (swap_in (parcs (subdivide H e w))) (parcs H) = perm_restrict (swap_in (parcs H)) (parcs (subdivide H e w))" by (auto simp: perm_restrict_def fun_eq_iff swap_in_def) have "subdivision_step (with_proj H) (swap_in (arcs H)) (with_proj (subdivide H e w)) (swap_in (arcs (subdivide H e w))) (u, v, w) (e, (u,w), (v,w))" unfolding subdivision_step_def unfolding prod.simps with_proj_simps using divide uvw by (intro conjI H.bidirected_digraph I.bidirected_digraph *) (auto simp add: swap_in_def \e = _\ compatibleI_with_proj) ultimately show ?case by (auto intro: subdivision.divide) qed private lemma subdivision_with_projD: assumes "subdivision_pair G H" shows "pair_sd G H" using assms proof (induct "with_proj H" "swap_in (parcs H)" arbitrary: H rule: subdivision_induct) case base interpret bidirected_digraph "with_proj G" "swap_in (parcs G)" by fact from base have "G = H" by (simp add: with_proj_def) with base show ?case by (auto intro: pair_sd.base pair_bidirected_digraphI_bidirected_digraph) next case (divide I rev_I u v w uv uw vw) define I' where "I' = \ pverts = verts I, parcs = arcs I \" have "compatible G I " using \subdivision (with_proj G, _) (I, _)\ by (rule subdivision_compatible) then have "tail I = fst" "head I = snd" by (auto simp: compatible_def) then have I: "I = I'" by (auto simp: I'_def) moreover from I have "rev_I = swap_in (parcs I')" using \subdivision_step _ _ _ _ _ _\ by (simp add: subdivision_step_def bidirected_digraph_rev_conv_pair) ultimately have pd_sd: "pair_sd G I'" by (auto intro: divide.hyps) interpret sd: subdiv_step I' "swap_in (parcs I')" H "swap_in (parcs H)" u v w uv uw vw using \subdivision_step _ _ _ _ _ _\ unfolding \I = _\ \rev_I = _\ by unfold_locales have ends: "uv = (u,v)" "uw = (u,w)" "vw = (v,w)" using sd.subdiv_ate by simp_all then have si_ends: "swap_in (parcs H) (u,w) = (w,u)" "swap_in (parcs H) (v,w) = (w,v)" "swap_in (parcs I') (u,v) = (v,u)" using sd.subdiv_ends_H_rev sd.subdiv_ends_G_rev by (auto simp: swap_in_def) have "parcs H = parcs I' - {(u, v), (v, u)} \ {(u, w), (w, u), (w, v), (v, w)}" using sd.in_arcs_G sd.not_in_arcs_G sd.arcs_H by (auto simp: si_ends ends) then have "H = subdivide I' uv w" using sd.verts_H by (simp add: ends subdivide.simps) then show ?case using sd.in_arcs_G sd.not_in_verts_G by (auto intro: pd_sd pair_sd.divide) qed private lemma subdivision_pair_conv: "pair_sd G H = subdivision_pair G H " by (metis subdivision_with_projD subdivision_with_projI) lemmas subdivision_pair_induct = pair_sd.induct[ unfolded subdivision_pair_conv, case_names base divide, induct pred: pair_sd] lemmas subdivision_pair_base = pair_sd.base[unfolded subdivision_pair_conv] lemmas subdivision_pair_divide = pair_sd.divide[unfolded subdivision_pair_conv] lemmas subdivision_pair_intros = pair_sd.intros[unfolded subdivision_pair_conv] lemmas subdivision_pair_cases = pair_sd.cases[unfolded subdivision_pair_conv] lemmas subdivision_pair_simps = pair_sd.simps[unfolded subdivision_pair_conv] lemmas bidirected_digraphI_subdivision = bidirected_digraphI_pair_sd[unfolded subdivision_pair_conv] end lemma (in pair_graph) pair_graph_subdivision: assumes "subdivision_pair G H" shows "pair_graph H" using assms by (induct rule: subdivision_pair_induct) (blast intro: pair_graph.pair_graph_subdivide divide)+ end diff --git a/thys/Gromov_Hyperbolicity/Isometries_Classification.thy b/thys/Gromov_Hyperbolicity/Isometries_Classification.thy --- a/thys/Gromov_Hyperbolicity/Isometries_Classification.thy +++ b/thys/Gromov_Hyperbolicity/Isometries_Classification.thy @@ -1,1119 +1,1119 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Classification of isometries on a Gromov hyperbolic space\ theory Isometries_Classification imports Gromov_Boundary Busemann_Function begin text \Isometries of Gromov hyperbolic spaces are of three types: \begin{itemize} \item Elliptic ones, for which orbits are bounded. \item Parabolic ones, which are not elliptic and have exactly one fixed point at infinity. \item Loxodromic ones, which are not elliptic and have exactly two fixed points at infinity. \end{itemize} In this file, we show that all isometries are indeed of this form, and give further properties for each type. For the definition, we use another characterization in terms of stable translation length: for isometries which are not elliptic, then they are parabolic if the stable translation length is $0$, loxodromic if it is positive. This gives a very efficient definition, and it is clear from this definition that the three categories of isometries are disjoint. All the work is then to go from this general definition to the dynamical properties in terms of fixed points on the boundary. \ subsection \The translation length\ text \The translation length is the minimal translation distance of an isometry. The stable translation length is the limit of the translation length of $f^n$ divided by $n$.\ definition translation_length::"(('a::metric_space) \ 'a) \ real" where "translation_length f = Inf {dist x (f x)|x. True}" lemma translation_length_nonneg [simp, mono_intros]: "translation_length f \ 0" unfolding translation_length_def by (rule cInf_greatest, auto) lemma translation_length_le [mono_intros]: "translation_length f \ dist x (f x)" unfolding translation_length_def apply (rule cInf_lower) by (auto intro: bdd_belowI[of _ 0]) definition stable_translation_length::"(('a::metric_space) \ 'a) \ real" where "stable_translation_length f = Inf {translation_length (f^^n)/n |n. n > 0}" lemma stable_translation_length_nonneg [simp]: "stable_translation_length f \ 0" unfolding stable_translation_length_def by (rule cInf_greatest, auto) lemma stable_translation_length_le_translation_length [mono_intros]: "n * stable_translation_length f \ translation_length (f^^n)" proof - have *: "stable_translation_length f \ translation_length (f^^n)/n" if "n > 0" for n unfolding stable_translation_length_def apply (rule cInf_lower) using that by (auto intro: bdd_belowI[of _ 0]) show ?thesis apply (cases "n = 0") using * by (auto simp add: divide_simps algebra_simps) qed lemma semicontraction_iterates: fixes f::"('a::metric_space) \ 'a" assumes "1-lipschitz_on UNIV f" shows "1-lipschitz_on UNIV (f^^n)" by (induction n, auto intro!: lipschitz_onI lipschitz_on_compose2[of 1 UNIV _ 1 f, simplified] lipschitz_on_subset[OF assms]) text \If $f$ is a semicontraction, then its stable translation length is the limit of $d(x, f^n x)/n$ for any $n$. While it is obvious that the liminf of this quantity is at least the stable translation length (which is defined as an inf over all points and all times), the opposite inequality is more interesting. One may find a point $y$ and a time $k$ for which $d(y, f^k y)/k$ is very close to the stable translation length. By subadditivity of the sequence $n \mapsto f(y, f^n y)$ and Fekete's Lemma, it follows that, for any large $n$, then $d(y, f^n y)/n$ is also very close to the stable translation length. Since this is equal to $d(x, f^n x)/n$ up to $\pm 2 d(x,y)/n$, the result follows.\ proposition stable_translation_length_as_pointwise_limit: assumes "1-lipschitz_on UNIV f" shows "(\n. dist x ((f^^n) x)/n) \ stable_translation_length f" proof - have *: "subadditive (\n. dist y ((f^^n) y))" for y proof (rule subadditiveI) fix m n::nat have "dist y ((f ^^ (m + n)) y) \ dist y ((f^^m) y) + dist ((f^^m) y) ((f^^(m+n)) y)" by (rule dist_triangle) also have "... = dist y ((f^^m) y) + dist ((f^^m) y) ((f^^m) ((f^^n) y))" by (auto simp add: funpow_add) also have "... \ dist y ((f^^m) y) + dist y ((f^^n) y)" using semicontraction_iterates[OF assms, of m] unfolding lipschitz_on_def by auto finally show "dist y ((f ^^ (m + n)) y) \ dist y ((f ^^ m) y) + dist y ((f ^^ n) y)" by simp qed have Ly: "(\n. dist y ((f^^n) y) / n) \ Inf {dist y ((f^^n) y) / n |n. n > 0}" for y by (auto intro!: bdd_belowI[of _ 0] subadditive_converges_bounded'[OF subadditive_imp_eventually_subadditive[OF *]]) have "eventually (\n. dist x ((f^^n) x)/n < l) sequentially" if "stable_translation_length f < l" for l proof - obtain m where m: "stable_translation_length f < m" "m < l" using \stable_translation_length f < l\ dense by auto have "\t \ {translation_length (f^^n)/n |n. n > 0}. t < m" apply (subst cInf_less_iff[symmetric]) using m unfolding stable_translation_length_def by (auto intro!: bdd_belowI[of _ 0]) then obtain k where k: "k > 0" "translation_length (f^^k)/k < m" by auto have "translation_length (f^^k) < k * m" using k by (simp add: divide_simps algebra_simps) then have "\t \ {dist y ((f^^k) y) |y. True}. t < k * m" apply (subst cInf_less_iff[symmetric]) unfolding translation_length_def by (auto intro!: bdd_belowI[of _ 0]) then obtain y where y: "dist y ((f^^k) y) < k * m" by auto have A: "eventually (\n. dist y ((f^^n) y)/n < m) sequentially" apply (auto intro!: order_tendstoD[OF Ly] iffD2[OF cInf_less_iff] bdd_belowI[of _ 0] exI[of _ "dist y ((f^^k) y)/k"]) using y k by (auto simp add: algebra_simps divide_simps) have B: "eventually (\n. dist x y * (1/n) < (l-m)/2) sequentially" apply (intro order_tendstoD[of _ "dist x y * 0"] tendsto_intros) using \m < l\ by simp have C: "dist x ((f^^n) x)/n < l" if "n > 0" "dist y ((f^^n) y)/n < m" "dist x y * (1/n) < (l-m)/2" for n proof - have "dist x ((f^^n) x) \ dist x y + dist y ((f^^n) y) + dist ((f^^n) y) ((f^^n) x)" by (intro mono_intros) also have "... \ dist x y + dist y ((f^^n) y) + dist y x" using semicontraction_iterates[OF assms, of n] unfolding lipschitz_on_def by auto also have "... = 2 * dist x y + dist y ((f^^n) y)" by (simp add: dist_commute) also have "... < 2 * real n * (l-m)/2 + n * m" apply (intro mono_intros) using that by (auto simp add: algebra_simps divide_simps) also have "... = n * l" by (simp add: algebra_simps divide_simps) finally show ?thesis using that by (simp add: algebra_simps divide_simps) qed show "eventually (\n. dist x ((f^^n) x)/n < l) sequentially" by (rule eventually_mono[OF eventually_conj[OF eventually_conj[OF A B] eventually_gt_at_top[of 0]] C], auto) qed moreover have "eventually (\n. dist x ((f^^n) x)/n > l) sequentially" if "stable_translation_length f > l" for l proof - have *: "dist x ((f^^n) x)/n > l" if "n > 0" for n proof - have "n * l < n * stable_translation_length f" using \stable_translation_length f > l\ \n > 0\ by auto also have "... \ translation_length (f^^n)" by (intro mono_intros) also have "... \ dist x ((f^^n) x)" by (intro mono_intros) finally show ?thesis using \n > 0\ by (auto simp add: algebra_simps divide_simps) qed then show ?thesis by (rule eventually_mono[rotated], auto) qed ultimately show ?thesis by (rule order_tendstoI[rotated]) qed text \It follows from the previous proposition that the stable translation length is also the limit of the renormalized translation length of $f^n$.\ proposition stable_translation_length_as_limit: assumes "1-lipschitz_on UNIV f" shows "(\n. translation_length (f^^n) / n) \ stable_translation_length f" proof - obtain x::'a where True by auto show ?thesis proof (rule tendsto_sandwich[of "\n. stable_translation_length f" _ _ "\n. dist x ((f^^n) x)/n"]) have "stable_translation_length f \ translation_length (f ^^ n) / real n" if "n > 0" for n using stable_translation_length_le_translation_length[of n f] that by (simp add: divide_simps algebra_simps) then show "eventually (\n. stable_translation_length f \ translation_length (f ^^ n) / real n) sequentially" by (rule eventually_mono[rotated], auto) have "translation_length (f ^^ n) / real n \ dist x ((f ^^ n) x) / real n" for n using translation_length_le[of "f^^n" x] by (auto simp add: divide_simps) then show "eventually (\n. translation_length (f ^^ n) / real n \ dist x ((f ^^ n) x) / real n) sequentially" by auto qed (auto simp add: stable_translation_length_as_pointwise_limit[OF assms]) qed lemma stable_translation_length_inv: assumes "isometry f" shows "stable_translation_length (inv f) = stable_translation_length f" proof - have *: "dist basepoint (((inv f)^^n) basepoint) = dist basepoint ((f^^n) basepoint)" for n proof - have "basepoint = (f^^n) (((inv f)^^n) basepoint)" by (metis assms comp_apply fn_o_inv_fn_is_id isometry_inverse(2)) then have "dist basepoint ((f^^n) basepoint) = dist ((f^^n) (((inv f)^^n) basepoint)) ((f^^n) basepoint)" by auto also have "... = dist (((inv f)^^n) basepoint) basepoint" unfolding isometryD(2)[OF isometry_iterates[OF assms]] by simp finally show ?thesis by (simp add: dist_commute) qed have "(\n. dist basepoint ((f^^n) basepoint)/n) \ stable_translation_length f" using stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF assms]] by simp moreover have "(\n. dist basepoint ((f^^n) basepoint)/n) \ stable_translation_length (inv f)" unfolding *[symmetric] using stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF isometry_inverse(1)[OF assms]]] by simp ultimately show ?thesis using LIMSEQ_unique by auto qed subsection \The strength of an isometry at a fixed point at infinity\ text \The additive strength of an isometry at a fixed point at infinity is the asymptotic average every point is moved towards the fixed point at each step. It is measured using the Busemann function.\ definition additive_strength::"('a::Gromov_hyperbolic_space \ 'a) \ ('a Gromov_completion) \ real" where "additive_strength f xi = lim (\n. (Busemann_function_at xi ((f^^n) basepoint) basepoint)/n)" text \For additivity reasons, as the Busemann function is a quasi-morphism, the additive strength measures the deplacement even at finite times. It is also uniform in terms of the basepoint. This shows that an isometry sends horoballs centered at a fixed point to horoballs, up to a uniformly bounded error depending only on $\delta$.\ lemma Busemann_function_eq_additive_strength: assumes "isometry f" "Gromov_extension f xi = xi" shows "\Busemann_function_at xi ((f^^n) x) (x::'a::Gromov_hyperbolic_space) - real n * additive_strength f xi\ \ 2 * deltaG(TYPE('a))" proof - define u where "u = (\y n. Busemann_function_at xi ((f^^n) y) y)" have *: "abs(u y (m+n) - u y m - u y n) \ 2 * deltaG(TYPE('a))" for n m y proof - have P: "Gromov_extension (f^^m) xi = xi" unfolding Gromov_extension_isometry_iterates[OF assms(1)] apply (induction m) using assms by auto have *: "u y n = Busemann_function_at xi ((f^^m) ((f^^n) y)) ((f^^m) y)" apply (subst P[symmetric]) unfolding Busemann_function_isometry[OF isometry_iterates[OF \isometry f\]] u_def by auto show ?thesis unfolding * unfolding u_def using Busemann_function_quasi_morphism[of xi "(f^^(m+n)) y" "(f^^m) y" y] unfolding funpow_add comp_def by auto qed define l where "l = (\y. lim (\n. u y n/n))" have A: "abs(u y k - k * l y) \ 2 * deltaG(TYPE('a))" for y k unfolding l_def using almost_additive_converges(2) * by auto then have *: "abs(Busemann_function_at xi ((f^^k) y) y - k * l y) \ 2 * deltaG(TYPE('a))" for y k unfolding u_def by auto have "l basepoint = additive_strength f xi" unfolding l_def u_def additive_strength_def by auto have "abs(k * l basepoint - k * l x) \ 4 * deltaG(TYPE('a)) + 2 * dist basepoint x" for k::nat proof - have "abs(k * l basepoint - k * l x) = abs((Busemann_function_at xi ((f^^k) x) x - k * l x) - (Busemann_function_at xi ((f^^k) basepoint) basepoint - k * l basepoint) + (Busemann_function_at xi ((f^^k) basepoint) basepoint - Busemann_function_at xi ((f^^k) x) x))" by auto also have "... \ abs (Busemann_function_at xi ((f^^k) x) x - k * l x) + abs (Busemann_function_at xi ((f^^k) basepoint) basepoint - k * l basepoint) + abs (Busemann_function_at xi ((f^^k) basepoint) basepoint - Busemann_function_at xi ((f^^k) x) x)" by auto also have "... \ 2 * deltaG(TYPE('a)) + 2 * deltaG(TYPE('a)) + (dist ((f^^k) basepoint) ((f^^k) x) + dist basepoint x)" by (intro mono_intros *) also have "... = 4 * deltaG(TYPE('a)) + 2 * dist basepoint x" unfolding isometryD[OF isometry_iterates[OF assms(1)]] by auto finally show ?thesis by auto qed moreover have "u = v" if H: "\k::nat. abs(k * u - k * v) \ C" for u v C::real proof - have "(\n. abs(u - v)) \ 0" proof (rule tendsto_sandwich[of "\n. 0" _ _ "\n::nat. C/n"], auto) have "(\n. C*(1/n)) \ C * 0" by (intro tendsto_intros) then show "(\n. C/n) \ 0" by auto have "\u - v\ \ C / real n" if "n \ 1" for n using H[of n] that apply (simp add: divide_simps algebra_simps) by (metis H abs_mult abs_of_nat right_diff_distrib') then show "\\<^sub>F n in sequentially. \u - v\ \ C / real n" unfolding eventually_sequentially by auto qed then show ?thesis by (metis LIMSEQ_const_iff abs_0_eq eq_iff_diff_eq_0) qed ultimately have "l basepoint = l x" by auto show ?thesis using A[of x n] unfolding u_def \l basepoint = l x\[symmetric] \l basepoint = additive_strength f xi\ by auto qed lemma additive_strength_as_limit [tendsto_intros]: assumes "isometry f" "Gromov_extension f xi = xi" shows "(\n. Busemann_function_at xi ((f^^n) x) x/n) \ additive_strength f xi" proof - have "(\n. abs(Busemann_function_at xi ((f^^n) x) x/n - additive_strength f xi)) \ 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. (2 * deltaG(TYPE('a))) * (1/real n)"], auto) unfolding eventually_sequentially apply (rule exI[of _ 1]) using Busemann_function_eq_additive_strength[OF assms] apply (simp add: divide_simps algebra_simps) using tendsto_mult[OF _ lim_1_over_n] by auto then show ?thesis using LIM_zero_iff tendsto_rabs_zero_cancel by blast qed text \The additive strength measures the amount of displacement towards a fixed point at infinity. Therefore, the distance from $x$ to $f^n x$ is at least $n$ times the additive strength, but one might think that it might be larger, if there is displacement along the horospheres. It turns out that this is not the case: the displacement along the horospheres is at most logarithmic (this is a classical property of parabolic isometries in hyperbolic spaces), and in fact it is bounded for loxodromic elements. We prove here that the growth is at most logarithmic in all cases, using a small computation based on the hyperbolicity inequality, expressed in Lemma \verb+dist_minus_Busemann_max_ineq+ above. This lemma will be used below to show that the translation length is the absolute value of the additive strength.\ lemma dist_le_additive_strength: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" "n \ 1" shows "dist x ((f^^n) x) \ dist x (f x) + real n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" proof - have A: "\n. n \ {1..2^k} \ dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + k * 16 * deltaG(TYPE('a))" for k proof (induction k) case 0 fix n::nat assume "n \ {1..2^0}" then have "n = 1" by auto then show "dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + real 0 * 16 * deltaG(TYPE('a))" using assms(3) by auto next case (Suc k) fix N::nat assume "N \ {1..2^(Suc k)}" then consider "N \ {1..2^k}" | "N \ {2^k<..2^(Suc k)}" using not_le by auto then show "dist x ((f ^^ N) x) - real N * additive_strength f xi \ dist x (f x) + real (Suc k) * 16 * deltaG TYPE('a)" proof (cases) case 1 show ?thesis by (rule order_trans[OF Suc.IH[OF 1]], auto simp add: algebra_simps) next case 2 define m::nat where "m = N - 2^k" define n::nat where "n = 2^k" have nm: "N = n+m" "m \ {1..2^k}" "n \ {1..2^k}"unfolding m_def n_def using 2 by auto have *: "(f^^(n+m)) x = (f^^n) ((f^^m) x)" unfolding funpow_add comp_def by auto have **: "(f^^(n+m)) x = (f^^m) ((f^^n) x)" apply (subst add.commute) unfolding funpow_add comp_def by auto have "dist x ((f^^N) x) - N * additive_strength f xi - 2 * deltaG(TYPE('a)) \ dist x ((f^^(n+m)) x) - Busemann_function_at xi ((f^^(n+m)) x) x" unfolding nm(1) using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of "n+m" x] by auto also have "... \ max (dist x ((f^^n) x) - Busemann_function_at xi ((f^^n) x) x) (dist ((f^^n) x) ((f^^(n+m)) x) - Busemann_function_at xi ((f^^(n+m)) x) ((f^^n) x) - 2 * Busemann_function_at xi ((f^^n) x) x) + 8 * deltaG(TYPE('a))" using dist_minus_Busemann_max_ineq by auto also have "... \ max (dist x ((f^^n) x) - (n * additive_strength f xi - 2 * deltaG(TYPE('a)))) (dist ((f^^n) x) ((f^^(n+m)) x) - (m * additive_strength f xi - 2 * deltaG(TYPE('a))) - 2 * (n * additive_strength f xi - 2 * deltaG(TYPE('a)))) + 8 * deltaG(TYPE('a))" unfolding ** apply (intro mono_intros) using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of n x] Busemann_function_eq_additive_strength[OF assms(1) assms(2), of m "(f^^n) x"] by auto also have "... \ max (dist x ((f^^n) x) - n * additive_strength f xi + 6 * deltaG(TYPE('a))) (dist x ((f^^m) x) - m * additive_strength f xi + 6 * deltaG(TYPE('a))) + 8 * deltaG(TYPE('a))" unfolding * isometryD(2)[OF isometry_iterates[OF assms(1)], of n] using assms(3) by (intro mono_intros, auto) also have "... = max (dist x ((f^^n) x) - n * additive_strength f xi) (dist x ((f^^m) x) - m * additive_strength f xi) + 14 * deltaG(TYPE('a))" unfolding max_add_distrib_left[symmetric] by auto also have "... \ dist x (f x) + k * 16 * deltaG(TYPE('a)) + 14 * deltaG(TYPE('a))" using nm by (auto intro!: Suc.IH) finally show ?thesis by (auto simp add: algebra_simps) qed qed define k::nat where "k = nat(ceiling (log 2 n))" have "n \ 2^k" unfolding k_def by (meson less_log2_of_power not_le real_nat_ceiling_ge) then have "dist x ((f^^n) x) - real n * additive_strength f xi \ dist x (f x) + k * 16 * deltaG(TYPE('a))" using A[of n k] \n \ 1\ by auto moreover have "real (nat \log 2 (real n)\) = real_of_int \log 2 (real n)\" by (metis Transcendental.log_one \n \ 2 ^ k\ assms(4) ceiling_zero int_ops(2) k_def le_antisym nat_eq_iff2 of_int_1 of_int_of_nat_eq order_refl power_0) ultimately show ?thesis unfolding k_def by (auto simp add: algebra_simps) qed text \The strength of the inverse of a map is the opposite of the strength of the map.\ lemma additive_strength_inv: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" shows "additive_strength (inv f) xi = - additive_strength f xi" proof - have *: "(inv f ^^ n) ((f ^^ n) x) = x" for n x - by (metis assms(1) comp_apply funpow_code_def inv_fn_o_fn_is_id isometry_inverse(2)) + by (metis assms(1) comp_apply inv_fn_o_fn_is_id isometry_inverse(2)) have A: "abs(real n * additive_strength f xi + real n * additive_strength (inv f) xi) \ 6 * deltaG (TYPE('a))" for n::nat and x::'a using Busemann_function_quasi_morphism[of xi x "(f^^n) x" x] Busemann_function_eq_additive_strength[OF assms, of n x] Busemann_function_eq_additive_strength[OF isometry_inverse(1)[OF assms(1)] Gromov_extension_inv_fixed_point[OF assms], of n "(f^^n) x"] unfolding * by auto have B: "abs(additive_strength f xi + additive_strength (inv f) xi) \ 6 * deltaG(TYPE('a)) * (1/n)" if "n \ 1" for n::nat using that A[of n] apply (simp add: divide_simps algebra_simps) unfolding distrib_left[symmetric] by auto have "(\n. abs(additive_strength f xi + additive_strength (inv f) xi)) \ 6 * deltaG(TYPE('a)) * 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. 6 * deltaG(TYPE('a)) * (1/real n)"], simp) unfolding eventually_sequentially apply (rule exI[of _ 1]) using B apply simp by (simp, intro tendsto_intros) then show ?thesis using LIMSEQ_unique mult_zero_right tendsto_const by fastforce qed text \We will now prove that the stable translation length of an isometry is given by the absolute value of its strength at any fixed point. We start with the case where the strength is nonnegative, and then reduce to this case by considering the map or its inverse.\ lemma stable_translation_length_eq_additive_strength_aux: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" shows "stable_translation_length f = additive_strength f xi" proof - have "(\n. dist x ((f^^n) x)/n) \ additive_strength f xi" for x proof (rule tendsto_sandwich[of "\n. (n * additive_strength f xi - 2 * deltaG(TYPE('a)))/real n" _ _ "\n. (dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a)))/ n"]) have "n * additive_strength f xi - 2 * deltaG TYPE('a) \ dist x ((f ^^ n) x)" for n using Busemann_function_eq_additive_strength[OF assms(1) assms(2), of n x] Busemann_function_le_dist[of xi "(f^^n) x" x] by (simp add: dist_commute) then have "(n * additive_strength f xi - 2 * deltaG TYPE('a)) / n \ dist x ((f ^^ n) x) / n" if "n \ 1" for n using that by (simp add: divide_simps) then show "\\<^sub>F n in sequentially. (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n \ dist x ((f ^^ n) x) / real n" unfolding eventually_sequentially by auto have B: "(\n. additive_strength f xi - (2 * deltaG(TYPE('a))) * (1/n)) \ additive_strength f xi - (2 * deltaG(TYPE('a))) * 0" by (intro tendsto_intros) show "(\n. (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n) \ additive_strength f xi" proof (rule Lim_transform_eventually) show "eventually (\n. additive_strength f xi - (2 * deltaG(TYPE('a))) * (1/n) = (real n * additive_strength f xi - 2 * deltaG TYPE('a)) / real n) sequentially" unfolding eventually_sequentially apply (rule exI[of _ 1]) by (simp add: divide_simps) qed (use B in auto) have "dist x ((f^^n) x) \ dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" if "n \ 1" for n using dist_le_additive_strength[OF assms that] by simp then have "(dist x ((f^^n) x))/n \ (dist x (f x) + n * additive_strength f xi + ceiling (log 2 n) * 16 * deltaG(TYPE('a)))/n" if "n \ 1" for n using that by (simp add: divide_simps) then show "\\<^sub>F n in sequentially. dist x ((f ^^ n) x) / real n \ (dist x (f x) + real n * additive_strength f xi + real_of_int (\log 2 (real n)\ * 16) * deltaG TYPE('a)) / real n" unfolding eventually_sequentially by auto have B: "(\n. dist x (f x) * (1/n) + additive_strength f xi + 16 * deltaG TYPE('a) * (\log 2 n\ / n)) \ dist x (f x) * 0 + additive_strength f xi + 16 * deltaG TYPE('a) * 0" by (intro tendsto_intros) show "(\n. (dist x (f x) + n * additive_strength f xi + real_of_int (\log 2 n\ * 16) * deltaG TYPE('a)) / real n) \ additive_strength f xi" proof (rule Lim_transform_eventually) show "eventually (\n. dist x (f x) * (1/n) + additive_strength f xi + 16 * deltaG TYPE('a) * (\log 2 n\ / n) = (dist x (f x) + real n * additive_strength f xi + real_of_int (\log 2 (real n)\ * 16) * deltaG TYPE('a)) / real n) sequentially" unfolding eventually_sequentially apply (rule exI[of _ 1]) by (simp add: algebra_simps divide_simps) qed (use B in auto) qed then show ?thesis using LIMSEQ_unique stable_translation_length_as_pointwise_limit[OF isometryD(4)[OF assms(1)]] by blast qed lemma stable_translation_length_eq_additive_strength: assumes "isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "Gromov_extension f xi = xi" shows "stable_translation_length f = abs(additive_strength f xi)" proof (cases "additive_strength f xi \ 0") case True then show ?thesis using stable_translation_length_eq_additive_strength_aux[OF assms] by auto next case False then have *: "abs(additive_strength f xi) = additive_strength (inv f) xi" unfolding additive_strength_inv[OF assms] by auto show ?thesis unfolding * stable_translation_length_inv[OF assms(1), symmetric] using stable_translation_length_eq_additive_strength_aux[OF isometry_inverse(1)[OF assms(1)] Gromov_extension_inv_fixed_point[OF assms]] * by auto qed subsection \Elliptic isometries\ text \Elliptic isometries are the simplest ones: they have bounded orbits.\ definition elliptic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "elliptic_isometry f = (isometry f \ (\x. bounded {(f^^n) x|n. True}))" lemma elliptic_isometryD: assumes "elliptic_isometry f" shows "bounded {(f^^n) x |n. True}" "isometry f" using assms unfolding elliptic_isometry_def by auto lemma elliptic_isometryI [intro]: assumes "bounded {(f^^n) x |n. True}" "isometry f" shows "elliptic_isometry f" proof - have "bounded {(f^^n) y |n. True}" for y proof - obtain c e where c: "\n. dist c ((f^^n) x) \ e" using assms(1) unfolding bounded_def by auto have "dist c ((f^^n) y) \ e + dist x y" for n proof - have "dist c ((f^^n) y) \ dist c ((f^^n) x) + dist ((f^^n) x) ((f^^n) y)" by (intro mono_intros) also have "... \ e + dist x y" using c[of n] isometry_iterates[OF assms(2), of n] by (intro mono_intros, auto simp add: isometryD) finally show ?thesis by simp qed then show ?thesis unfolding bounded_def by auto qed then show ?thesis unfolding elliptic_isometry_def using assms by auto qed text \The inverse of an elliptic isometry is an elliptic isometry.\ lemma elliptic_isometry_inv: assumes "elliptic_isometry f" shows "elliptic_isometry (inv f)" proof - obtain c e where A: "\n. dist c ((f^^n) basepoint) \ e" using elliptic_isometryD(1)[OF assms, of basepoint] unfolding bounded_def by auto have "c = (f^^n) (((inv f)^^n) c)" for n using fn_o_inv_fn_is_id[OF isometry_inverse(2)[OF elliptic_isometryD(2)[OF assms]], of n] unfolding comp_def by metis then have "dist ((f^^n) (((inv f)^^n) c)) ((f^^n) basepoint) \ e" for n using A by auto then have B: "dist basepoint (((inv f)^^n) c) \ e" for n unfolding isometryD(2)[OF isometry_iterates[OF elliptic_isometryD(2)[OF assms]]] by (auto simp add: dist_commute) show ?thesis apply (rule elliptic_isometryI[of _ c]) using isometry_inverse(1)[OF elliptic_isometryD(2)[OF assms]] B unfolding bounded_def by auto qed text \The inverse of a bijective map is an elliptic isometry if and only if the original map is.\ lemma elliptic_isometry_inv_iff: assumes "bij f" shows "elliptic_isometry (inv f) \ elliptic_isometry f" using elliptic_isometry_inv[of f] elliptic_isometry_inv[of "inv f"] inv_inv_eq[OF assms] by auto text \The identity is an elliptic isometry.\ lemma elliptic_isometry_id: "elliptic_isometry id" by (intro elliptic_isometryI isometryI, auto) text \The translation length of an elliptic isometry is $0$.\ lemma elliptic_isometry_stable_translation_length: assumes "elliptic_isometry f" shows "stable_translation_length f = 0" proof - obtain x::'a where True by auto have "bounded {(f^^n) x|n. True}" using elliptic_isometryD[OF assms] by auto then obtain c C where cC: "\n. dist c ((f^^n) x) \ C" unfolding bounded_def by auto have "(\n. dist x ((f^^n) x)/n) \ 0" proof (rule tendsto_sandwich[of "\_. 0" _ sequentially "\n. 2 * C / n"]) have "(\n. 2 * C * (1 / real n)) \ 2 * C * 0" by (intro tendsto_intros) then show "(\n. 2 * C / real n) \ 0" by auto have "dist x ((f ^^ n) x) / real n \ 2 * C / real n" for n using cC[of 0] cC[of n] dist_triangle[of x "(f^^n) x" c] by (auto simp add: algebra_simps divide_simps dist_commute) then show "eventually (\n. dist x ((f ^^ n) x) / real n \ 2 * C / real n) sequentially" by auto qed (auto) moreover have "(\n. dist x ((f^^n) x)/n) \ stable_translation_length f" by (rule stable_translation_length_as_pointwise_limit[OF isometry_on_lipschitz[OF isometryD(1)[OF elliptic_isometryD(2)[OF assms]]]]) ultimately show ?thesis using LIMSEQ_unique by auto qed text \If an isometry has a fixed point, then it is elliptic.\ lemma isometry_with_fixed_point_is_elliptic: assumes "isometry f" "f x = x" shows "elliptic_isometry f" proof - have *: "(f^^n) x = x" for n apply (induction n) using assms(2) by auto show ?thesis apply (rule elliptic_isometryI[of _ x, OF _ assms(1)]) unfolding * by auto qed subsection \Parabolic and loxodromic isometries\ text \An isometry is parabolic if it is not elliptic and if its translation length vanishes.\ definition parabolic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "parabolic_isometry f = (isometry f \ \elliptic_isometry f \ stable_translation_length f = 0)" text \An isometry is loxodromic if it is not elliptic and if its translation length is nonzero.\ definition loxodromic_isometry::"('a::Gromov_hyperbolic_space \ 'a) \ bool" where "loxodromic_isometry f = (isometry f \ \elliptic_isometry f \ stable_translation_length f \ 0)" text \The main features of such isometries are expressed in terms of their fixed points at infinity. We define them now, but proving that the definitions make sense will take some work.\ definition neutral_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "neutral_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi = 0)" definition attracting_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "attracting_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi < 0)" definition repelling_fixed_point::"('a::Gromov_hyperbolic_space \ 'a) \ 'a Gromov_completion" where "repelling_fixed_point f = (SOME xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi > 0)" lemma parabolic_isometryD: assumes "parabolic_isometry f" shows "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f = 0" "\elliptic_isometry f" using assms unfolding parabolic_isometry_def by auto lemma parabolic_isometryI: assumes "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f = 0" shows "parabolic_isometry f" using assms unfolding parabolic_isometry_def elliptic_isometry_def by auto lemma loxodromic_isometryD: assumes "loxodromic_isometry f" shows "isometry f" "\bounded {(f^^n) x|n. True}" "stable_translation_length f > 0" "\elliptic_isometry f" using assms unfolding loxodromic_isometry_def by (auto, meson dual_order.antisym not_le stable_translation_length_nonneg) text \To have a loxodromic isometry, it suffices to know that the stable translation length is nonzero, as elliptic isometries have zero translation length.\ lemma loxodromic_isometryI: assumes "isometry f" "stable_translation_length f \ 0" shows "loxodromic_isometry f" using assms elliptic_isometry_stable_translation_length unfolding loxodromic_isometry_def by auto text \Any isometry is elliptic, or parabolic, or loxodromic, and these possibilities are mutually exclusive.\ lemma elliptic_or_parabolic_or_loxodromic: assumes "isometry f" shows "elliptic_isometry f \ parabolic_isometry f \ loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma elliptic_imp_not_parabolic_loxodromic: assumes "elliptic_isometry f" shows "\parabolic_isometry f" "\loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma parabolic_imp_not_elliptic_loxodromic: assumes "parabolic_isometry f" shows "\elliptic_isometry f" "\loxodromic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto lemma loxodromic_imp_not_elliptic_parabolic: assumes "loxodromic_isometry f" shows "\elliptic_isometry f" "\parabolic_isometry f" using assms unfolding parabolic_isometry_def loxodromic_isometry_def by auto text \The inverse of a parabolic isometry is parabolic.\ lemma parabolic_isometry_inv: assumes "parabolic_isometry f" shows "parabolic_isometry (inv f)" unfolding parabolic_isometry_def using isometry_inverse[of f] elliptic_isometry_inv_iff[of f] parabolic_isometryD[OF assms] stable_translation_length_inv[of f] by auto text \The inverse of a loxodromic isometry is loxodromic.\ lemma loxodromic_isometry_inv: assumes "loxodromic_isometry f" shows "loxodromic_isometry (inv f)" unfolding loxodromic_isometry_def using isometry_inverse[of f] elliptic_isometry_inv_iff[of f] loxodromic_isometryD[OF assms] stable_translation_length_inv[of f] by auto text \We will now prove that an isometry which is not elliptic has a fixed point at infinity. This is very easy if the space is proper (ensuring that the Gromov completion is compact), but in fact this holds in general. One constructs it by considering a sequence $r_n$ such that $f^{r_n} 0$ tends to infinity, and additionally $d(f^l 0, 0) < d(f^{r_n} 0, 0)$ for $l < r_n$: this implies the convergence at infinity of $f^{r_n} 0$, by an argument based on a Gromov product computation -- and the limit is a fixed point. Moreover, it has nonpositive additive strength, essentially by construction.\ lemma high_scores: fixes u::"nat \ real" and i::nat and C::real assumes "\(bdd_above (range u))" shows "\n. (\l \ n. u l \ u n) \ u n \ C \ n \ i" proof - define M where "M = max C (Max {u l|l. l < i})" define n where "n = Inf {m. u m > M}" have "\(range u \ {..M})" using assms by (meson bdd_above_Iic bdd_above_mono) then have "{m. u m > M} \ {}" using assms by (simp add: image_subset_iff not_less) then have "n \ {m. u m > M}" unfolding n_def using Inf_nat_def1 by metis then have "u n > M" by simp have "n \ i" proof (rule ccontr) assume "\ i \ n" then have *: "n < i" by simp have "u n \ Max {u l|l. l < i}" apply (rule Max_ge) using * by auto then show False using \u n > M\ M_def by auto qed moreover have "u l \ u n" if "l \ n" for l proof (cases "l = n") case True then show ?thesis by simp next case False then have "l < n" using \l \ n\ by auto then have "l \ {m. u m > M}" unfolding n_def by (meson bdd_below_def cInf_lower not_le zero_le) then show ?thesis using \u n > M\ by auto qed ultimately show ?thesis using \u n > M\ M_def less_eq_real_def by auto qed lemma isometry_not_elliptic_has_attracting_fixed_point: assumes "isometry f" "\(elliptic_isometry f)" shows "\xi \ Gromov_boundary. Gromov_extension f xi = xi \ additive_strength f xi \ 0" proof - define u where "u = (\n. dist basepoint ((f^^n) basepoint))" have NB: "\(bdd_above (range u))" proof assume "bdd_above (range u)" then obtain C where *: "\n. u n \ C" unfolding bdd_above_def by auto have "bounded {(f^^n) basepoint|n. True}" unfolding bounded_def apply (rule exI[of _ basepoint], rule exI[of _ C]) using * unfolding u_def by auto then show False using elliptic_isometryI assms by auto qed have "\r. \n. ((\l \ r n. u l \ u (r n)) \ u (r n) \ 2 * n) \ r (Suc n) \ r n + 1" apply (rule dependent_nat_choice) using high_scores[OF NB] by (auto) blast then obtain r::"nat \ nat" where r: "\n l. l \ r n \ u l \ u (r n)" "\n. u (r n) \ 2 * n" "\n. r (Suc n) \ r n + 1" by auto then have "strict_mono r" by (metis Suc_eq_plus1 Suc_le_lessD strict_monoI_Suc) then have "r n \ n" for n by (simp add: seq_suble) have A: "dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ dist basepoint ((f^^(r n)) basepoint)" if "n \ p" for n p proof - have "r n \ r p" using \n \ p\ \strict_mono r\ by (simp add: strict_mono_less_eq) then have *: "f^^((r n)) = (f^^(r p)) o (f^^(r n - r p))" unfolding funpow_add[symmetric] by auto have "dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) = dist ((f^^(r p)) basepoint) ((f^^(r p)) ((f^^(r n - r p)) basepoint))" unfolding * comp_def by auto also have "... = dist basepoint ((f^^(r n - r p)) basepoint)" using isometry_iterates[OF assms(1), of "r p"] isometryD by auto also have "... \ dist basepoint ((f^^(r n)) basepoint)" using r(1)[of "r n - r p" n] unfolding u_def by auto finally show ?thesis by simp qed have *: "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ p" if "n \ p" for n p proof - have "2 * Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) = dist basepoint ((f^^(r p)) basepoint) + dist basepoint ((f^^(r n)) basepoint) - dist ((f^^(r p)) basepoint) ((f^^(r n)) basepoint)" unfolding Gromov_product_at_def by auto also have "... \ dist basepoint ((f^^(r p)) basepoint)" using A[OF that] by auto finally show "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ p" using r(2)[of p] unfolding u_def by auto qed have *: "Gromov_product_at basepoint ((f^^(r p)) basepoint) ((f^^(r n)) basepoint) \ N" if "n \ N" "p \ N" for n p N using *[of n p] *[of p n] that by (auto simp add: Gromov_product_commute intro: le_cases[of n p]) have "Gromov_converging_at_boundary (\n. (f^^(r n)) basepoint)" apply (rule Gromov_converging_at_boundaryI[of basepoint]) using * by (meson dual_order.trans real_arch_simple) with Gromov_converging_at_boundary_converges[OF this] obtain xi where xi: "(\n. to_Gromov_completion ((f^^(r n)) basepoint)) \ xi" "xi \ Gromov_boundary" by auto then have *: "(\n. Gromov_extension f (to_Gromov_completion ((f^^(r n)) basepoint))) \ xi" apply (simp, rule Gromov_converging_at_boundary_bounded_perturbation[of _ _ _ "dist basepoint (f basepoint)"]) by (simp add: assms(1) funpow_swap1 isometryD(2) isometry_iterates) moreover have "(\n. Gromov_extension f (to_Gromov_completion ((f^^(r n)) basepoint))) \ Gromov_extension f xi" using continuous_on_tendsto_compose[OF Gromov_extension_isometry(2)[OF assms(1)] xi(1)] by auto ultimately have fxi: "Gromov_extension f xi = xi" using LIMSEQ_unique by auto have "Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ 0" if "n \ p" for n p unfolding Busemann_function_inner using A[OF that] by auto then have A: "eventually (\n. Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ 0) sequentially" for p unfolding eventually_sequentially by auto have B: "eventually (\n. Busemann_function_at (to_Gromov_completion ((f^^(r n)) basepoint)) ((f^^(r p)) basepoint) basepoint \ Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1) sequentially" for p by (rule eventually_mono[OF Busemann_function_inside_approx[OF _ xi(1), of 1 "((f^^(r p)) basepoint)" basepoint, simplified]], simp) have "eventually (\n. Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1 \ 0) sequentially" for p by (rule eventually_mono[OF eventually_conj[OF A[of p] B[of p]]], simp) then have *: "Busemann_function_at xi ((f^^(r p)) basepoint) basepoint - 2 * deltaG(TYPE('a)) - 1 \ 0" for p by auto then have A: "Busemann_function_at xi ((f^^(r p)) basepoint) basepoint / (r p) - (2 * deltaG(TYPE('a)) + 1) * (1/r p) \ 0" if "p \ 1" for p using order_trans[OF that \p \ r p\] by (auto simp add: algebra_simps divide_simps) have B: "(\p. Busemann_function_at xi ((f^^(p)) basepoint) basepoint / p - (2 * deltaG(TYPE('a)) + 1) * (1/p)) \ additive_strength f xi - (2 * deltaG(TYPE('a)) + 1) * 0" by (intro tendsto_intros assms fxi) have "additive_strength f xi - (2 * deltaG TYPE('a) + 1) * 0 \ 0" apply (rule LIMSEQ_le_const2[OF LIMSEQ_subseq_LIMSEQ[OF B \strict_mono r\]]) using A unfolding comp_def by auto then show ?thesis using xi fxi by auto qed text \Applying the previous result to the inverse map, we deduce that there is also a fixed point with nonnegative strength.\ lemma isometry_not_elliptic_has_repelling_fixed_point: assumes "isometry f" "\(elliptic_isometry f)" shows "\xi \ Gromov_boundary. Gromov_extension f xi = xi \ additive_strength f xi \ 0" proof - have *: "\elliptic_isometry (inv f)" using elliptic_isometry_inv_iff isometry_inverse(2)[OF assms(1)] assms(2) by auto obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension (inv f) xi = xi" "additive_strength (inv f) xi \ 0" using isometry_not_elliptic_has_attracting_fixed_point[OF isometry_inverse(1)[OF assms(1)] *] by auto have *: "Gromov_extension f xi = xi" using Gromov_extension_inv_fixed_point[OF isometry_inverse(1)[OF assms(1)] xi(2)] inv_inv_eq[OF isometry_inverse(2)[OF assms(1)]] by auto moreover have "additive_strength f xi \ 0" using additive_strength_inv[OF assms(1) *] xi(3) by auto ultimately show ?thesis using xi(1) by auto qed subsubsection \Parabolic isometries\ text \We show that a parabolic isometry has (at least) one neutral fixed point at infinity.\ lemma parabolic_fixed_point: assumes "parabolic_isometry f" shows "neutral_fixed_point f \ Gromov_boundary" "Gromov_extension f (neutral_fixed_point f) = neutral_fixed_point f" "additive_strength f (neutral_fixed_point f) = 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" using isometry_not_elliptic_has_attracting_fixed_point parabolic_isometryD[OF assms] by blast moreover have "additive_strength f xi = 0" using stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] xi(2)] parabolic_isometryD(3)[OF assms] by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi = 0" by auto show "neutral_fixed_point f \ Gromov_boundary" "Gromov_extension f (neutral_fixed_point f) = neutral_fixed_point f" "additive_strength f (neutral_fixed_point f) = 0" unfolding neutral_fixed_point_def using someI_ex[OF A] by auto qed text \Parabolic isometries have exactly one fixed point, the neutral fixed point at infinity. The proof goes as follows: if it has another fixed point, then the orbit of a basepoint would stay on the horospheres centered at both fixed points. But the intersection of two horospheres based at different points is a a bounded set. Hence, the map has a bounded orbit, and is therefore elliptic.\ theorem parabolic_unique_fixed_point: assumes "parabolic_isometry f" shows "Gromov_extension f xi = xi \ xi = neutral_fixed_point f" proof (auto simp add: parabolic_fixed_point[OF assms]) assume H: "Gromov_extension f xi = xi" have *: "additive_strength f xi = 0" using stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] H] parabolic_isometryD(3)[OF assms] by auto show "xi = neutral_fixed_point f" proof (rule ccontr) assume N: "xi \ neutral_fixed_point f" define C where "C = 2 * real_of_ereal (extended_Gromov_product_at basepoint xi (neutral_fixed_point f)) + 4 * deltaG(TYPE('a))" have A: "dist basepoint ((f^^n) basepoint) \ C" for n proof - have "dist ((f^^n) basepoint) basepoint - 2 * real_of_ereal (extended_Gromov_product_at basepoint xi (neutral_fixed_point f)) - 2 * deltaG(TYPE('a)) \ max (Busemann_function_at xi ((f^^n) basepoint) basepoint) (Busemann_function_at (neutral_fixed_point f) ((f^^n) basepoint) basepoint)" using dist_le_max_Busemann_functions[OF N] by (simp add: algebra_simps) also have "... \ max (n * additive_strength f xi + 2 * deltaG(TYPE('a))) (n * additive_strength f (neutral_fixed_point f) + 2 * deltaG(TYPE('a)))" apply (intro mono_intros) using Busemann_function_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] H, of n basepoint] Busemann_function_eq_additive_strength[OF parabolic_isometryD(1)[OF assms] parabolic_fixed_point(2)[OF assms], of n basepoint] by auto also have "... = 2 * deltaG(TYPE('a))" unfolding * parabolic_fixed_point[OF assms] by auto finally show ?thesis unfolding C_def by (simp add: algebra_simps dist_commute) qed have "elliptic_isometry f" apply (rule elliptic_isometryI[of _ basepoint]) using parabolic_isometryD(1)[OF assms] A unfolding bounded_def by auto then show False using elliptic_imp_not_parabolic_loxodromic assms by auto qed qed text \When one iterates a parabolic isometry, the distance to the starting point can grow at most logarithmically.\ lemma parabolic_logarithmic_growth: assumes "parabolic_isometry (f::'a::Gromov_hyperbolic_space \ 'a)" "n \ 1" shows "dist x ((f^^n) x) \ dist x (f x) + ceiling (log 2 n) * 16 * deltaG(TYPE('a))" using dist_le_additive_strength[OF parabolic_isometryD(1)[OF assms(1)] parabolic_fixed_point(2)[OF assms(1)] _ assms(2)] parabolic_isometryD(3)[OF assms(1)] stable_translation_length_eq_additive_strength[OF parabolic_isometryD(1)[OF assms(1)] parabolic_fixed_point(2)[OF assms(1)]] by auto text \It follows that there is no parabolic isometry in trees, since the formula in the previous lemma shows that there is no orbit growth as $\delta = 0$, and therefore orbits are bounded, contradicting the parabolicity of the isometry.\ lemma tree_no_parabolic_isometry: assumes "isometry (f::'a::Gromov_hyperbolic_space_0 \ 'a)" shows "elliptic_isometry f \ loxodromic_isometry f" proof - have "\parabolic_isometry f" proof assume P: "parabolic_isometry f" have "dist basepoint ((f^^n) basepoint) \ dist basepoint (f basepoint)" if "n \ 1" for n using parabolic_logarithmic_growth[OF P that, of basepoint] by auto then have *: "dist basepoint ((f^^n) basepoint) \ dist basepoint (f basepoint)" for n by (cases "n \ 1", auto simp add: not_less_eq_eq) have "elliptic_isometry f" apply (rule elliptic_isometryI[OF _ assms, of basepoint]) using * unfolding bounded_def by auto then show False using elliptic_imp_not_parabolic_loxodromic P by auto qed then show ?thesis using elliptic_or_parabolic_or_loxodromic[OF assms] by auto qed subsubsection \Loxodromic isometries\ text \A loxodromic isometry has (at least) two fixed points at infinity, one attracting and one repelling. We have already constructed fixed points with nonnegative and nonpositive strengths. Since the strength is nonzero (its absolute value is the stable translation length), then these fixed points correspond to what we want.\ lemma loxodromic_attracting_fixed_point: assumes "loxodromic_isometry f" shows "attracting_fixed_point f \ Gromov_boundary" "Gromov_extension f (attracting_fixed_point f) = attracting_fixed_point f" "additive_strength f (attracting_fixed_point f) < 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" using isometry_not_elliptic_has_attracting_fixed_point loxodromic_isometryD[OF assms] by blast moreover have "additive_strength f xi < 0" using stable_translation_length_eq_additive_strength[OF loxodromic_isometryD(1)[OF assms] xi(2)] loxodromic_isometryD(3)[OF assms] xi(3) by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi < 0" by auto show "attracting_fixed_point f \ Gromov_boundary" "Gromov_extension f (attracting_fixed_point f) = attracting_fixed_point f" "additive_strength f (attracting_fixed_point f) < 0" unfolding attracting_fixed_point_def using someI_ex[OF A] by auto qed lemma loxodromic_repelling_fixed_point: assumes "loxodromic_isometry f" shows "repelling_fixed_point f \ Gromov_boundary" "Gromov_extension f (repelling_fixed_point f) = repelling_fixed_point f" "additive_strength f (repelling_fixed_point f) > 0" proof - obtain xi where xi: "xi \ Gromov_boundary" "Gromov_extension f xi = xi" "additive_strength f xi \ 0" using isometry_not_elliptic_has_repelling_fixed_point loxodromic_isometryD[OF assms] by blast moreover have "additive_strength f xi > 0" using stable_translation_length_eq_additive_strength[OF loxodromic_isometryD(1)[OF assms] xi(2)] loxodromic_isometryD(3)[OF assms] xi(3) by auto ultimately have A: "\xi. xi \ Gromov_boundary \ Gromov_extension f xi = xi \ additive_strength f xi > 0" by auto show "repelling_fixed_point f \ Gromov_boundary" "Gromov_extension f (repelling_fixed_point f) = repelling_fixed_point f" "additive_strength f (repelling_fixed_point f) > 0" unfolding repelling_fixed_point_def using someI_ex[OF A] by auto qed text \The attracting and repelling fixed points of a loxodromic isometry are distinct -- precisely since one is attracting and the other is repelling.\ lemma attracting_fixed_point_neq_repelling_fixed_point: assumes "loxodromic_isometry f" shows "attracting_fixed_point f \ repelling_fixed_point f" using loxodromic_repelling_fixed_point[OF assms] loxodromic_attracting_fixed_point[OF assms] by auto text \The attracting fixed point of a loxodromic isometry is indeed attracting. Moreover, the convergence is uniform away from the repelling fixed point. This is expressed in the following proposition, where neighborhoods of the repelling and attracting fixed points are given by the property that the Gromov product with the fixed point is large. The proof goes as follows. First, the Busemann function with respect to the fixed points at infinity evolves like the strength. Therefore, $f^n e$ tends to the repulsive fixed point in negative time, and to the attracting one in positive time. Consider now a general point $x$ with $(\xi^-, x)_e \leq K$. This means that the geodesics from $e$ to $x$ and $\xi^-$ diverge before time $K$. For large $n$, since $f^{-n} e$ is close to $\xi^-$, we also get the inequality $(f^{-n} e, x)_e \leq K$. Applying $f^n$ and using the invariance of the Gromov product under isometries yields $(e, f^n x)_{f^n e} \leq K$. But this Gromov product is equal to $d(e, f^n e) - (f^n e, f^n x)_e$ (this is a general property of Gromov products). In particular, $(f^n e, f^n x) \geq d(e, f^n e) - K$, and moreover $d(e, f^n e)$ is large. Since $f^n e$ is close to $\xi^+$, it follows that $f^n x$ is also close to $\xi^+$, as desired. The real proof requires some more care as everything should be done in ereal, and moreover every inequality is only true up to some multiple of $\delta$. But everything works in the way just described above. \ proposition loxodromic_attracting_fixed_point_attracts_uniformly: assumes "loxodromic_isometry f" shows "\N. \n \ N. \x. extended_Gromov_product_at basepoint x (repelling_fixed_point f) \ ereal K \ extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) x) (attracting_fixed_point f) \ ereal M" proof - have I: "isometry f" using loxodromic_isometryD(1)[OF assms] by simp have J: "\ereal r\ \ \" for r by auto text \We show that $f^n e$ tends to the repelling fixed point in negative time.\ have "(\n. ereal (Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint)) \ - \" proof (rule tendsto_sandwich[of "\n. -\" _ _ "\n. ereal(- real n * additive_strength f (repelling_fixed_point f) + 2 * deltaG(TYPE('a)))", OF _ always_eventually], auto) fix n have "Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint \ real n * additive_strength (inv f) (repelling_fixed_point f) + 2 * deltaG(TYPE('a))" using Busemann_function_eq_additive_strength[OF isometry_inverse(1)[OF I] Gromov_extension_inv_fixed_point[OF I loxodromic_repelling_fixed_point(2)[OF assms]], of n basepoint] by auto then show "Busemann_function_at (repelling_fixed_point f) ((inv f ^^ n) basepoint) basepoint \ 2 * deltaG(TYPE('a)) - real n * additive_strength f (repelling_fixed_point f)" by (simp add: I additive_strength_inv assms loxodromic_repelling_fixed_point(2)) next have "(\n. ereal (2 * deltaG TYPE('a)) + ereal (- real n) * additive_strength f (repelling_fixed_point f)) \ ereal (2 * deltaG (TYPE('a))) + (- \) * additive_strength f (repelling_fixed_point f)" apply (intro tendsto_intros) using loxodromic_repelling_fixed_point(3)[OF assms] by auto then show "(\n. ereal (2 * deltaG TYPE('a) - real n * additive_strength f (repelling_fixed_point f))) \ - \" using loxodromic_repelling_fixed_point(3)[OF assms] by auto qed then have "(\n. to_Gromov_completion (((inv f)^^n) basepoint)) \ repelling_fixed_point f" by (rule Busemann_function_minus_infinity_imp_convergent[of _ _ basepoint]) then have "(\n. extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ \" unfolding Gromov_completion_boundary_limit[OF loxodromic_repelling_fixed_point(1)[OF assms]] by simp then obtain Nr where Nr: "\n. n \ Nr \ extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f) \ ereal (K + deltaG(TYPE('a)) + 1)" unfolding Lim_PInfty by auto text \We show that $f^n e$ tends to the attracting fixed point in positive time.\ have "(\n. ereal (Busemann_function_at (attracting_fixed_point f) ((f ^^ n) basepoint) basepoint)) \ - \" proof (rule tendsto_sandwich[of "\n. -\" _ _ "\n. ereal(real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG(TYPE('a)))", OF _ always_eventually], auto) fix n show "Busemann_function_at (attracting_fixed_point f) ((f ^^ n) basepoint) basepoint \ real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG(TYPE('a))" using Busemann_function_eq_additive_strength[OF I loxodromic_attracting_fixed_point(2)[OF assms], of n basepoint] by auto next have "(\n. ereal (2 * deltaG TYPE('a)) + ereal (real n) * additive_strength f (attracting_fixed_point f)) \ ereal (2 * deltaG (TYPE('a))) + (\) * additive_strength f (attracting_fixed_point f)" apply (intro tendsto_intros) using loxodromic_attracting_fixed_point(3)[OF assms] by auto then show "(\n. ereal (real n * additive_strength f (attracting_fixed_point f) + 2 * deltaG TYPE('a))) \ - \" using loxodromic_attracting_fixed_point(3)[OF assms] by (auto simp add: algebra_simps) qed then have *: "(\n. to_Gromov_completion (((f)^^n) basepoint)) \ attracting_fixed_point f" by (rule Busemann_function_minus_infinity_imp_convergent[of _ _ basepoint]) then have "(\n. extended_Gromov_product_at basepoint (to_Gromov_completion (((f)^^n) basepoint)) (attracting_fixed_point f)) \ \" unfolding Gromov_completion_boundary_limit[OF loxodromic_attracting_fixed_point(1)[OF assms]] by simp then obtain Na where Na: "\n. n \ Na \ extended_Gromov_product_at basepoint (to_Gromov_completion (((f)^^n) basepoint)) (attracting_fixed_point f) \ ereal (M + deltaG(TYPE('a)))" unfolding Lim_PInfty by auto text \We show that the distance between $e$ and $f^n e$ tends to infinity.\ have "(\n. extended_Gromov_distance (to_Gromov_completion basepoint) (to_Gromov_completion ((f^^n) basepoint))) \ extended_Gromov_distance (to_Gromov_completion basepoint) (attracting_fixed_point f)" using * extended_Gromov_distance_continuous[of "to_Gromov_completion basepoint"] by (metis UNIV_I continuous_on filterlim_compose tendsto_at_iff_tendsto_nhds) then have "(\n. extended_Gromov_distance (to_Gromov_completion basepoint) (to_Gromov_completion ((f^^n) basepoint))) \ \" using loxodromic_attracting_fixed_point(1)[OF assms] by simp then obtain Nd where Nd: "\n. n \ Nd \ dist basepoint ((f^^n) basepoint) \ M + K + 3 * deltaG(TYPE('a))" unfolding Lim_PInfty by auto text \Now, if $n$ is large enough so that all the convergences to infinity above are almost realized, we show that a point $x$ which is not too close to $\xi^-$ is sent by $f^n$ to a point close to $\xi^+$, uniformly.\ define N where "N = Nr + Na + Nd" have "extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) x) (attracting_fixed_point f) \ ereal M" if H: "extended_Gromov_product_at basepoint x (repelling_fixed_point f) \ K" "n \ N" for n x proof - have n: "n \ Nr" "n \ Na" "n \ Nd" using that unfolding N_def by auto have "min (extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ extended_Gromov_product_at basepoint x (repelling_fixed_point f) + deltaG(TYPE('a))" by (intro mono_intros) also have "... \ ereal K + deltaG(TYPE('a))" apply (intro mono_intros) using H by auto finally have "min (extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f)) \ ereal K + deltaG(TYPE('a))" by simp moreover have "extended_Gromov_product_at basepoint (to_Gromov_completion (((inv f)^^n) basepoint)) (repelling_fixed_point f) > ereal K + deltaG(TYPE('a))" using Nr[OF n(1)] ereal_le_less by auto ultimately have A: "extended_Gromov_product_at basepoint x (to_Gromov_completion (((inv f)^^n) basepoint)) \ ereal K + deltaG(TYPE('a))" using not_le by fastforce have *: "((inv f)^^n) ((f^^n) z) = z" for z by (metis I bij_is_inj inj_fn inv_f_f inv_fn isometry_inverse(2)) have **: "x = Gromov_extension ((inv f)^^n) (Gromov_extension (f^^n) x)" using Gromov_extension_isometry_composition[OF isometry_iterates[OF I] isometry_iterates[OF isometry_inverse(1)[OF I]], of n n] unfolding comp_def * apply auto by meson have "extended_Gromov_product_at (((inv f)^^n) ((f^^n) basepoint)) (Gromov_extension ((inv f)^^n) (Gromov_extension (f^^n) x)) (Gromov_extension (((inv f)^^n)) (to_Gromov_completion basepoint)) \ ereal K + deltaG(TYPE('a))" using A by (simp add: * **[symmetric]) then have B: "extended_Gromov_product_at ((f^^n) basepoint) (Gromov_extension (f^^n) x) (to_Gromov_completion basepoint) \ ereal K + deltaG(TYPE('a))" unfolding Gromov_extension_preserves_extended_Gromov_product[OF isometry_iterates[OF isometry_inverse(1)[OF I]]] by simp have "dist basepoint ((f^^n) basepoint) - deltaG(TYPE('a)) \ extended_Gromov_product_at ((f^^n) basepoint) (Gromov_extension (f^^n) x) (to_Gromov_completion basepoint) + extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))" using extended_Gromov_product_add_ge[of basepoint "(f^^n) basepoint" "Gromov_extension (f^^n) x"] by (auto simp add: algebra_simps) also have "... \ (ereal K + deltaG(TYPE('a))) + extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))" by (intro mono_intros B) finally have "extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint)) \ dist basepoint ((f^^n) basepoint) - (2 * deltaG(TYPE('a)) + K)" apply (simp only: ereal_minus_le [OF J] ereal_minus(1) [symmetric]) apply (auto simp add: algebra_simps) apply (metis (no_types, hide_lams) add.assoc add.left_commute mult_2_right plus_ereal.simps(1)) done moreover have "dist basepoint ((f ^^ n) basepoint) - (2 * deltaG TYPE('a) + K) \ M + deltaG(TYPE('a))" using Nd[OF n(3)] by auto ultimately have "extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint)) \ ereal (M + deltaG(TYPE('a)))" using order_trans ereal_le_le by auto then have "ereal (M + deltaG(TYPE('a))) \ min (extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (to_Gromov_completion ((f^^n) basepoint))) (extended_Gromov_product_at basepoint (to_Gromov_completion ((f^^n) basepoint)) (attracting_fixed_point f))" using Na[OF n(2)] by (simp add: extended_Gromov_product_at_commute) also have "... \ extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (attracting_fixed_point f) + deltaG(TYPE('a))" by (intro mono_intros) finally have "ereal M \ extended_Gromov_product_at basepoint (Gromov_extension (f^^n) x) (attracting_fixed_point f)" unfolding plus_ereal.simps(1)[symmetric] ereal_add_le_add_iff2 by auto then show ?thesis by (simp add: Gromov_extension_isometry_iterates I) qed then show ?thesis by auto qed text \We deduce pointwise convergence from the previous result.\ lemma loxodromic_attracting_fixed_point_attracts: assumes "loxodromic_isometry f" "xi \ repelling_fixed_point f" shows "(\n. ((Gromov_extension f)^^n) xi) \ attracting_fixed_point f" proof - have "(\n. extended_Gromov_product_at basepoint (((Gromov_extension f)^^n) xi) (attracting_fixed_point f)) \ \" unfolding Lim_PInfty using loxodromic_attracting_fixed_point_attracts_uniformly[OF assms(1)] - by (metis Gromov_boundary_extended_product_PInf assms(2) ereal_top funpow_code_def infinity_ereal_def linear) + by auto (metis Gromov_boundary_extended_product_PInf assms(2) dual_order.refl real_le_ereal_iff real_of_ereal_le_0 zero_ereal_def) then show ?thesis unfolding Gromov_completion_boundary_limit[OF loxodromic_attracting_fixed_point(1)[OF assms(1)]] by simp qed text \Finally, we show that a loxodromic isometry has exactly two fixed points, its attracting and repelling fixed points defined above. Indeed, we already know that these points are fixed. It remains to see that there is no other fixed point. But a fixed point which is not the repelling one is both stationary and attracted to the attracting fixed point by the previous lemma, hence it has to coincide with the attracting fixed point.\ theorem loxodromic_unique_fixed_points: assumes "loxodromic_isometry f" shows "Gromov_extension f xi = xi \ xi = attracting_fixed_point f \ xi = repelling_fixed_point f" proof - have "xi = attracting_fixed_point f" if H: "Gromov_extension f xi = xi" "xi \ repelling_fixed_point f" for xi proof - have "((Gromov_extension f)^^n) xi = xi" for n apply (induction n) using H(1) by auto then have "(\n. ((Gromov_extension f)^^n) xi) \ xi" by auto then show ?thesis using loxodromic_attracting_fixed_point_attracts[OF assms H(2)] LIMSEQ_unique by auto qed then show ?thesis using loxodromic_attracting_fixed_point(2)[OF assms] loxodromic_repelling_fixed_point(2)[OF assms] by auto qed end (*of theory Isometries_Classification*) diff --git a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy --- a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy +++ b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy @@ -1,1034 +1,1034 @@ section \Permutations as Products of Disjoint Cycles\ theory Executable_Permutations imports - Graph_Theory.Funpow - List_Aux "HOL-Library.Rewrite" "HOL-Combinatorics.Permutations" + Graph_Theory.Auxiliary + List_Aux begin subsection \Cyclic Permutations\ definition list_succ :: "'a list \ 'a \ 'a" where "list_succ xs x = (if x \ set xs then xs ! ((index xs x + 1) mod length xs) else x)" text \ We demonstrate the functions on the following simple lemmas @{lemma "list_succ [1 :: int, 2, 3] 1 = 2" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 2 = 3" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 3 = 1" by code_simp} \ lemma list_succ_altdef: "list_succ xs x = (let n = index xs x in if n + 1 = length xs then xs ! 0 else if n + 1 < length xs then xs ! (n + 1) else x)" using index_le_size[of xs x] unfolding list_succ_def index_less_size_conv[symmetric] by (auto simp: Let_def) lemma list_succ_Nil: "list_succ [] = id" by (simp add: list_succ_def fun_eq_iff) lemma list_succ_singleton: "list_succ [x] = list_succ []" by (simp add: fun_eq_iff list_succ_def) lemma list_succ_short: assumes "length xs < 2" shows "list_succ xs = id" using assms by (cases xs) (rename_tac [2] y ys, case_tac [2] ys, auto simp: list_succ_Nil list_succ_singleton) lemma list_succ_simps: "index xs x + 1 = length xs \ list_succ xs x = xs ! 0" "index xs x + 1 < length xs \ list_succ xs x = xs ! (index xs x + 1)" "length xs \ index xs x \ list_succ xs x = x" by (auto simp: list_succ_altdef) lemma list_succ_not_in: assumes "x \ set xs" shows "list_succ xs x = x" using assms by (auto simp: list_succ_def) lemma list_succ_list_succ_rev: assumes "distinct xs" shows "list_succ (rev xs) (list_succ xs x) = x" proof - { assume "index xs x + 1 < length xs" moreover then have "length xs - Suc (Suc (length xs - Suc (Suc (index xs x)))) = index xs x" by linarith ultimately have ?thesis using assms by (simp add: list_succ_def index_rev index_nth_id rev_nth) } moreover { assume A: "index xs x + 1 = length xs" moreover from A have "xs \ []" by auto moreover with A have "last xs = xs ! index xs x" by (cases "length xs") (auto simp: last_conv_nth) ultimately have ?thesis using assms by (auto simp add: list_succ_def rev_nth index_rev index_nth_id last_conv_nth) } moreover { assume A: "index xs x \ length xs" then have "x \ set xs" by (metis index_less less_irrefl) then have ?thesis by (auto simp: list_succ_def) } ultimately show ?thesis by (metis discrete le_less not_less) qed lemma inj_list_succ: "distinct xs \ inj (list_succ xs)" by (metis injI list_succ_list_succ_rev) lemma inv_list_succ_eq: "distinct xs \ inv (list_succ xs) = list_succ (rev xs)" by (metis distinct_rev inj_imp_inv_eq inj_list_succ list_succ_list_succ_rev) lemma bij_list_succ: "distinct xs \ bij (list_succ xs)" by (metis bij_def inj_list_succ distinct_rev list_succ_list_succ_rev surj_def) lemma list_succ_permutes: assumes "distinct xs" shows "list_succ xs permutes set xs" using assms by (auto simp: permutes_conv_has_dom bij_list_succ has_dom_def list_succ_def) lemma permutation_list_succ: assumes "distinct xs" shows "permutation (list_succ xs)" using list_succ_permutes[OF assms] by (auto simp: permutation_permutes) lemma list_succ_nth: assumes "distinct xs" "n < length xs" shows "list_succ xs (xs ! n) = xs ! (Suc n mod length xs)" using assms by (auto simp: list_succ_def index_nth_id) lemma list_succ_last[simp]: assumes "distinct xs" "xs \ []" shows "list_succ xs (last xs) = hd xs" using assms by (auto simp: list_succ_def hd_conv_nth) lemma list_succ_rotate1[simp]: assumes "distinct xs" shows "list_succ (rotate1 xs) = list_succ xs" proof (rule ext) fix y show "list_succ (rotate1 xs) y = list_succ xs y" using assms proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases "x = y") case True then have "index (xs @ [y]) y = length xs" using \distinct (x # xs)\ by (simp add: index_append) with True show ?thesis by (cases "xs=[]") (auto simp: list_succ_def nth_append) next case False then show ?thesis apply (cases "index xs y + 1 < length xs") apply (auto simp:list_succ_def index_append nth_append) by (metis Suc_lessI index_less_size_conv mod_self nth_Cons_0 nth_append nth_append_length) qed qed qed lemma list_succ_rotate[simp]: assumes "distinct xs" shows "list_succ (rotate n xs) = list_succ xs" using assms by (induct n) auto lemma list_succ_in_conv: "list_succ xs x \ set xs \ x \ set xs" by (auto simp: list_succ_def not_nil_if_in_set ) lemma list_succ_in_conv1: assumes "A \ set xs = {}" shows "list_succ xs x \ A \ x \ A" by (metis assms disjoint_iff_not_equal list_succ_in_conv list_succ_not_in) lemma list_succ_commute: assumes "set xs \ set ys = {}" shows "list_succ xs (list_succ ys x) = list_succ ys (list_succ xs x)" proof - have "\x. x \ set xs \ list_succ ys x = x" "\x. x \ set ys \ list_succ xs x = x" using assms by (blast intro: list_succ_not_in)+ then show ?thesis by (cases "x \ set xs \ set ys") (auto simp: list_succ_in_conv list_succ_not_in) qed subsection \Arbitrary Permutations\ fun lists_succ :: "'a list list \ 'a \ 'a" where "lists_succ [] x = x" | "lists_succ (xs # xss) x = list_succ xs (lists_succ xss x)" definition distincts :: "'a list list \ bool" where "distincts xss \ distinct xss \ (\xs \ set xss. distinct xs \ xs \ []) \ (\xs \ set xss. \ys \ set xss. xs \ ys \ set xs \ set ys = {})" lemma distincts_distinct: "distincts xss \ distinct xss" by (auto simp: distincts_def) lemma distincts_Nil[simp]: "distincts []" by (simp add: distincts_def) lemma distincts_single: "distincts [xs] \ distinct xs \ xs \ []" by (auto simp add: distincts_def) lemma distincts_Cons: "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (set xs \ (\ys \ set xss. set ys)) = {}" (is "?L \ ?R") proof assume ?L then show ?R by (auto simp: distincts_def) next assume ?R then have "distinct (xs # xss)" apply (auto simp: disjoint_iff_not_equal distincts_distinct) apply (metis length_greater_0_conv nth_mem) done moreover from \?R\ have "\xs \ set (xs # xss). distinct xs \ xs \ []" by (auto simp: distincts_def) moreover from \?R\ have "\xs' \ set (xs # xss). \ys \ set (xs # xss). xs' \ ys \ set xs' \ set ys = {}" by (simp add: distincts_def) blast ultimately show ?L unfolding distincts_def by (intro conjI) qed lemma distincts_Cons': "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (\ys \ set xss. set xs \ set ys = {})" (is "?L \ ?R") unfolding distincts_Cons by blast lemma distincts_rev: "distincts (map rev xss) \ distincts xss" by (simp add: distincts_def distinct_map) lemma length_distincts: assumes "distincts xss" shows "length xss = card (set ` set xss)" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons xs xss) then have "set xs \ set ` set xss" using equals0I[of "set xs"] by (auto simp: distincts_Cons disjoint_iff_not_equal ) with Cons show ?case by (auto simp add: distincts_Cons) qed lemma distincts_remove1: "distincts xss \ distincts (remove1 xs xss)" by (auto simp: distincts_def) lemma distinct_Cons_remove1: "x \ set xs \ distinct (x # remove1 x xs) = distinct xs" by (induct xs) auto lemma set_Cons_remove1: "x \ set xs \ set (x # remove1 x xs) = set xs" by (induct xs) auto lemma distincts_Cons_remove1: "xs \ set xss \ distincts (xs # remove1 xs xss) = distincts xss" by (simp only: distinct_Cons_remove1 set_Cons_remove1 distincts_def) lemma distincts_inj_on_set: assumes "distincts xss" shows "inj_on set (set xss)" by (rule inj_onI) (metis assms distincts_def inf.idem set_empty) lemma distincts_distinct_set: assumes "distincts xss" shows "distinct (map set xss)" using assms by (auto simp: distinct_map distincts_distinct distincts_inj_on_set) lemma distincts_distinct_nth: assumes "distincts xss" "n < length xss" shows "distinct (xss ! n)" using assms by (auto simp: distincts_def) lemma lists_succ_not_in: assumes "x \ (\xs\set xss. set xs)" shows "lists_succ xss x = x" using assms by (induct xss) (auto simp: list_succ_not_in) lemma lists_succ_in_conv: "lists_succ xss x \ (\xs\set xss. set xs) \ x \ (\xs\set xss. set xs)" by (induct xss) (auto simp: list_succ_in_conv lists_succ_not_in list_succ_not_in) lemma lists_succ_in_conv1: assumes "A \ (\xs\set xss. set xs) = {}" shows "lists_succ xss x \ A \ x \ A" by (metis Int_iff assms emptyE lists_succ_in_conv lists_succ_not_in) lemma lists_succ_Cons_pf: "lists_succ (xs # xss) = list_succ xs o lists_succ xss" by auto lemma lists_succ_Nil_pf: "lists_succ [] = id" by (simp add: fun_eq_iff) lemmas lists_succ_simps_pf = lists_succ_Cons_pf lists_succ_Nil_pf lemma lists_succ_permutes: assumes "distincts xss" shows "lists_succ xss permutes (\xs \ set xss. set xs)" using assms proof (induction xss) case Nil then show ?case by auto next case (Cons xs xss) have "list_succ xs permutes (set xs)" using Cons by (intro list_succ_permutes) (simp add: distincts_def in_set_member) moreover have "lists_succ xss permutes (\ys \ set xss. set ys)" using Cons by (auto simp: Cons distincts_def) ultimately show "lists_succ (xs # xss) permutes (\ys \ set (xs # xss). set ys)" using Cons by (auto simp: lists_succ_Cons_pf intro: permutes_compose permutes_subset) qed lemma bij_lists_succ: "distincts xss \ bij (lists_succ xss)" by (induct xss) (auto simp: lists_succ_simps_pf bij_comp bij_list_succ distincts_Cons) lemma lists_succ_snoc: "lists_succ (xss @ [xs]) = lists_succ xss o list_succ xs" by (induct xss) auto lemma inv_lists_succ_eq: assumes "distincts xss" shows "inv (lists_succ xss) = lists_succ (rev (map rev xss))" proof - have *: "\f g. inv (\b. f (g b)) = inv (f o g)" by (simp add: o_def) have **: "lists_succ [] = id" by auto show ?thesis using assms by (induct xss) (auto simp: * ** lists_succ_snoc lists_succ_Cons_pf o_inv_distrib inv_list_succ_eq distincts_Cons bij_list_succ bij_lists_succ) qed lemma lists_succ_remove1: assumes "distincts xss" "xs \ set xss" shows "lists_succ (xs # remove1 xs xss) = lists_succ xss" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons ys xss) show ?case proof cases assume "xs = ys" then show ?case by simp next assume "xs \ ys" with Cons.prems have inter: "set xs \ set ys = {}" and "xs \ set xss" by (auto simp: distincts_Cons) have dists: "distincts (xs # remove1 xs xss)" "distincts (xs # ys # remove1 xs xss)" using \distincts (ys # xss)\ \xs \ set xss\ by (auto simp: distincts_def) have "list_succ xs \ (list_succ ys \ lists_succ (remove1 xs xss)) = list_succ ys \ (list_succ xs \ lists_succ (remove1 xs xss))" using inter unfolding fun_eq_iff comp_def by (subst list_succ_commute) auto also have "\ = list_succ ys o (lists_succ (xs # remove1 xs xss))" using dists by (simp add: lists_succ_Cons_pf distincts_Cons) also have "\ = list_succ ys o lists_succ xss" using \xs \ set xss\ \distincts (ys # xss)\ by (simp add: distincts_Cons Cons.hyps) finally show "lists_succ (xs # remove1 xs (ys # xss)) = lists_succ (ys # xss)" using Cons dists by (auto simp: lists_succ_Cons_pf distincts_Cons) qed qed lemma lists_succ_no_order: assumes "distincts xss" "distincts yss" "set xss = set yss" shows "lists_succ xss = lists_succ yss" using assms proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) have "xs \ set xss" "xs \ set yss" using Cons.prems by (auto dest: distincts_distinct) have "lists_succ xss = lists_succ (remove1 xs yss)" using Cons.prems \xs \ _\ by (intro Cons.hyps) (auto simp add: distincts_Cons distincts_remove1 distincts_distinct) then have "lists_succ (xs # xss) = lists_succ (xs # remove1 xs yss)" using Cons.prems \xs \ _\ by (simp add: lists_succ_Cons_pf distincts_Cons_remove1) then show ?case using Cons.prems \xs \ _\ by (simp add: lists_succ_remove1) qed section \List Orbits\ text \Computes the orbit of @{term x} under @{term f}\ definition orbit_list :: "('a \ 'a) \ 'a \ 'a list" where "orbit_list f x \ iterate 0 (funpow_dist1 f x x) f x" partial_function (tailrec) orbit_list_impl :: "('a \ 'a) \ 'a \ 'a list \ 'a \ 'a list" where "orbit_list_impl f s acc x = (let x' = f x in if x' = s then rev (x # acc) else orbit_list_impl f s (x # acc) x')" context notes [simp] = length_fold_remove1_le begin text \Computes the list of orbits\ fun orbits_list :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list f [] = []" | "orbits_list f (x # xs) = orbit_list f x # orbits_list f (fold remove1 (orbit_list f x) xs)" fun orbits_list_impl :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list_impl f [] = []" | "orbits_list_impl f (x # xs) = (let fc = orbit_list_impl f x [] x in fc # orbits_list_impl f (fold remove1 fc xs))" declare orbit_list_impl.simps[code] end abbreviation sset :: "'a list list \ 'a set set" where "sset xss \ set ` set xss" lemma iterate_funpow_step: assumes "f x \ y" "y \ orbit f x" shows "iterate 0 (funpow_dist1 f x y) f x = x # iterate 0 (funpow_dist1 f (f x) y) f (f x)" proof - from assms have A: "y \ orbit f (f x)" by (simp add: orbit_step) have "iterate 0 (funpow_dist1 f x y) f x = x # iterate 1 (funpow_dist1 f x y) f x" (is "_ = _ # ?it") unfolding iterate_def by (rewrite in "\ = _" upt_conv_Cons) auto also have "?it = map (\n. (f ^^ n) x) (map Suc [0.. = map (\n. (f ^^ n) (f x)) [0.. = iterate 0 (funpow_dist1 f (f x) y) f (f x)" unfolding iterate_def unfolding iterate_def by (simp add: funpow_dist_step[OF assms(1) A]) finally show ?thesis . qed lemma orbit_list_impl_conv: assumes "y \ orbit f x" shows "orbit_list_impl f y acc x = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using assms proof (induct n\"funpow_dist1 f x y" arbitrary: x acc) case (Suc x) show ?case proof cases assume "f x = y" then show ?thesis by (subst orbit_list_impl.simps) (simp add: Let_def iterate_def funpow_dist_0) next assume not_y :"f x \ y" have y_in_succ: "y \ orbit f (f x)" by (intro orbit_step Suc.prems not_y) have "orbit_list_impl f y acc x = orbit_list_impl f y (x # acc) (f x)" using not_y by (subst orbit_list_impl.simps) simp also have "\ = rev (x # acc) @ iterate 0 (funpow_dist1 f (f x) y) f (f x)" (is "_ = ?rev @ ?it") by (intro Suc funpow_dist_step not_y y_in_succ) also have "\ = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using not_y Suc.prems by (simp add: iterate_funpow_step) finally show ?thesis . qed qed lemma orbit_list_conv_impl: assumes "x \ orbit f x" shows "orbit_list f x = orbit_list_impl f x [] x" unfolding orbit_list_impl_conv[OF assms] orbit_list_def by simp lemma set_orbit_list: assumes "x \ orbit f x" shows "set (orbit_list f x) = orbit f x" by (simp add: orbit_list_def orbit_conv_funpow_dist1[OF assms] set_iterate) lemma set_orbit_list': assumes "permutation f" shows "set (orbit_list f x) = orbit f x" using assms by (simp add: permutation_self_in_orbit set_orbit_list) lemma distinct_orbit_list: assumes "x \ orbit f x" shows "distinct (orbit_list f x)" by (simp del: upt_Suc add: orbit_list_def iterate_def distinct_map inj_on_funpow_dist1[OF assms]) lemma distinct_orbit_list': assumes "permutation f" shows "distinct (orbit_list f x)" using assms by (simp add: permutation_self_in_orbit distinct_orbit_list) lemma orbits_list_conv_impl: assumes "permutation f" shows "orbits_list f xs = orbits_list_impl f xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case using assms by (cases xs) (auto simp: assms less less_Suc_eq_le length_fold_remove1_le orbit_list_conv_impl permutation_self_in_orbit Let_def) qed lemma orbit_list_not_nil[simp]: "orbit_list f x \ []" by (simp add: orbit_list_def) lemma sset_orbits_list: assumes "permutation f" shows "sset (orbits_list f xs) = (orbit f) ` set xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case proof (cases xs) case Nil then show ?thesis by simp next case (Cons x' xs') let ?xs'' = "fold remove1 (orbit_list f x') xs'" have A: "sset (orbits_list f ?xs'') = orbit f ` set ?xs''" using Cons by (simp add: less_Suc_eq_le length_fold_remove1_le less.hyps) have B: "set (orbit_list f x') = orbit f x'" by (rule set_orbit_list) (simp add: permutation_self_in_orbit assms) have "orbit f ` set (fold remove1 (orbit_list f x') xs') \ orbit f ` set xs'" using set_fold_remove1[of _ xs'] by auto moreover have "orbit f ` set xs' - {orbit f x'} \ (orbit f ` set (fold remove1 (orbit_list f x') xs'))" (is "?L \ ?R") proof fix A assume "A \ ?L" then obtain y where "A = orbit f y" "y \ set xs'" by auto have "A \ orbit f x'" using \A \ ?L\ by auto from \A = _\ \A \ _\ have "y \ orbit f x'" by (meson assms cyclic_on_orbit orbit_cyclic_eq3 permutation_permutes) with \y \ _\ have "y \ set (fold remove1 (orbit_list f x') xs')" by (auto simp: set_fold_remove1' set_orbit_list permutation_self_in_orbit assms) then show "A \ ?R" using \A = _\ by auto qed ultimately show ?thesis by (auto simp: A B Cons) qed qed subsection \Relation to @{term cyclic_on}\ lemma list_succ_orbit_list: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "list_succ (orbit_list f s) = f" proof - have "distinct (orbit_list f s)" "\x. x \ set (orbit_list f s) \ x = f x" using assms by (simp_all add: distinct_orbit_list set_orbit_list) moreover have "\i. i < length (orbit_list f s) \ orbit_list f s ! (Suc i mod length (orbit_list f s)) = f (orbit_list f s ! i)" using funpow_dist1_prop[OF \s \ orbit f s\] by (auto simp: orbit_list_def funpow_mod_eq) ultimately show ?thesis by (auto simp: list_succ_def fun_eq_iff) qed lemma list_succ_funpow_conv: assumes A: "distinct xs" "x \ set xs" shows "(list_succ xs ^^ n) x = xs ! ((index xs x + n) mod length xs)" proof - have "xs \ []" using assms by auto then show ?thesis by (induct n) (auto simp: hd_conv_nth A index_nth_id list_succ_def mod_simps) qed lemma orbit_list_succ: assumes "distinct xs" "x \ set xs" shows "orbit (list_succ xs) x = set xs" proof (intro set_eqI iffI) fix y assume "y \ orbit (list_succ xs) x" then show "y \ set xs" by induct (auto simp: list_succ_in_conv \x \ set xs\) next fix y assume "y \ set xs" moreover { fix i j have "i < length xs \ j < length xs \ \n. xs ! j = xs ! ((i + n) mod length xs)" using assms by (auto simp: exI[where x="j + (length xs - i)"]) } ultimately show "y \ orbit (list_succ xs) x" using assms by (auto simp: orbit_altdef_permutation permutation_list_succ list_succ_funpow_conv index_nth_id in_set_conv_nth) qed lemma cyclic_on_list_succ: assumes "distinct xs" "xs \ []" shows "cyclic_on (list_succ xs) (set xs)" using assms last_in_set by (auto simp: cyclic_on_def orbit_list_succ) lemma obtain_orbit_list_func: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" obtains xs where "f = list_succ xs" "set xs = orbit f s" "distinct xs" "hd xs = s" proof - { from assms have "f = list_succ (orbit_list f s)" by (simp add: list_succ_orbit_list) moreover have "set (orbit_list f s) = orbit f s" "distinct (orbit_list f s)" by (auto simp: set_orbit_list distinct_orbit_list assms) moreover have "hd (orbit_list f s) = s" by (simp add: orbit_list_def iterate_def hd_map del: upt_Suc) ultimately have "\xs. f = list_succ xs \ set xs = orbit f s \ distinct xs \ hd xs = s" by blast } then show ?thesis by (metis that) qed lemma cyclic_on_obtain_list_succ: assumes "cyclic_on f S" "\x. x \ S \ f x = x" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" proof - from assms obtain s where s: "s \ orbit f s" "\x. x \ orbit f s \ f x = x" "S = orbit f s" by (auto simp: cyclic_on_def) then show ?thesis by (metis that obtain_orbit_list_func) qed lemma cyclic_on_obtain_list_succ': assumes "cyclic_on f S" "f permutes S" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" using assms unfolding permutes_def by (metis cyclic_on_obtain_list_succ) lemma list_succ_unique: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "\!xs. f = list_succ xs \ distinct xs \ hd xs = s \ set xs = orbit f s" proof - from assms obtain xs where xs: "f = list_succ xs" "distinct xs" "hd xs = s" "set xs = orbit f s" by (rule obtain_orbit_list_func) moreover { fix zs assume A: "f = list_succ zs" "distinct zs" "hd zs = s" "set zs = orbit f s" then have "zs \ []" using \s \ orbit f s\ by auto from \distinct xs\ \distinct zs\ \set xs = orbit f s\ \set zs = orbit f s\ have len: "length xs = length zs" by (metis distinct_card) { fix n assume "n < length xs" then have "zs ! n = xs ! n" proof (induct n) case 0 with A xs \zs \ []\ show ?case by (simp add: hd_conv_nth nth_rotate_conv_nth) next case (Suc n) then have "list_succ zs (zs ! n) = list_succ xs (xs! n)" using \f = list_succ xs\ \f = list_succ zs\ by simp with \Suc n < _\ show ?case by (simp add:list_succ_nth len \distinct xs\ \distinct zs\) qed } then have "zs = xs" by (metis len nth_equalityI) } ultimately show ?thesis by metis qed lemma distincts_orbits_list: assumes "distinct as" "permutation f" shows "distincts (orbits_list f as)" using assms(1) proof (induct "length as" arbitrary: as rule: less_induct) case less show ?case proof (cases as) case Nil then show ?thesis by simp next case (Cons a as') let ?as' = "fold remove1 (orbit_list f a) as'" from Cons less.prems have A: "distincts (orbits_list f (fold remove1 (orbit_list f a) as'))" by (intro less) (auto simp: distinct_fold_remove1 length_fold_remove1_le less_Suc_eq_le) have B: "set (orbit_list f a) \ \(sset (orbits_list f (fold remove1 (orbit_list f a) as'))) = {}" proof - have "orbit f a \ set (fold remove1 (orbit_list f a) as') = {}" using assms less.prems Cons by (simp add: set_fold_remove1_distinct set_orbit_list') then have "orbit f a \ \ (orbit f ` set (fold remove1 (orbit_list f a) as')) = {}" by auto (metis assms(2) cyclic_on_orbit disjoint_iff_not_equal permutation_self_in_orbit[OF assms(2)] orbit_cyclic_eq3 permutation_permutes) then show ?thesis using assms by (auto simp: set_orbit_list' sset_orbits_list disjoint_iff_not_equal) qed show ?thesis using A B assms by (auto simp: distincts_Cons Cons distinct_orbit_list') qed qed lemma cyclic_on_lists_succ': assumes "distincts xss" shows "A \ sset xss \ cyclic_on (lists_succ xss) A" using assms proof (induction xss arbitrary: A) case Nil then show ?case by auto next case (Cons xs xss A) then have inter: "set xs \ (\ys\set xss. set ys) = {}" by (auto simp: distincts_Cons) note pcp[OF _ _ inter] = permutes_comp_preserves_cyclic1 permutes_comp_preserves_cyclic2 from Cons show "cyclic_on (lists_succ (xs # xss)) A" by (cases "A = set xs") (auto intro: pcp simp: cyclic_on_list_succ list_succ_permutes lists_succ_permutes lists_succ_Cons_pf distincts_Cons) qed lemma cyclic_on_lists_succ: assumes "distincts xss" shows "\xs. xs \ set xss \ cyclic_on (lists_succ xss) (set xs)" using assms by (auto intro: cyclic_on_lists_succ') lemma permutes_as_lists_succ: assumes "distincts xss" assumes ls_eq: "\xs. xs \ set xss \ list_succ xs = perm_restrict f (set xs)" assumes "f permutes (\(sset xss))" shows "f = lists_succ xss" using assms proof (induct xss arbitrary: f) case Nil then show ?case by simp next case (Cons xs xss) let ?sets = "\xss. \ys \ set xss. set ys" have xs: "distinct xs" "xs \ []" using Cons by (auto simp: distincts_Cons) have f_xs: "perm_restrict f (set xs) = list_succ xs" using Cons by simp have co_xs: "cyclic_on (perm_restrict f (set xs)) (set xs)" unfolding f_xs using xs by (rule cyclic_on_list_succ) have perm_xs: "perm_restrict f (set xs) permutes set xs" unfolding f_xs using \distinct xs\ by (rule list_succ_permutes) have perm_xss: "perm_restrict f (?sets xss) permutes (?sets xss)" proof - have "perm_restrict f (?sets (xs # xss) - set xs) permutes (?sets (xs # xss) - set xs)" using Cons co_xs by (intro perm_restrict_diff_cyclic) (auto simp: cyclic_on_perm_restrict) also have "?sets (xs # xss) - set xs = ?sets xss" using Cons by (auto simp: distincts_Cons) finally show ?thesis . qed have f_xss: "perm_restrict f (?sets xss) = lists_succ xss" proof - have *: "\xs. xs \ set xss \ ((\x\set xss. set x) \ set xs) = set xs" by blast with perm_xss Cons.prems show ?thesis by (intro Cons.hyps) (auto simp: distincts_Cons perm_restrict_perm_restrict *) qed from Cons.prems show "f = lists_succ (xs # xss)" by (simp add: lists_succ_Cons_pf distincts_Cons f_xss[symmetric] perm_restrict_union perm_xs perm_xss) qed lemma cyclic_on_obtain_lists_succ: assumes permutes: "f permutes S" and S: "S = \(sset css)" and dists: "distincts css" and cyclic: "\cs. cs \ set css \ cyclic_on f (set cs)" obtains xss where "f = lists_succ xss" "distincts xss" "map set xss = map set css" "map hd xss = map hd css" proof - let ?fc = "\cs. perm_restrict f (set cs)" define some_list where "some_list cs = (SOME xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs)" for cs { fix cs assume "cs \ set css" then have "cyclic_on (?fc cs) (set cs)" "\x. x \ set cs \ ?fc cs x = x" "hd cs \ set cs" using cyclic dists by (auto simp add: cyclic_on_perm_restrict perm_restrict_def distincts_def) then have "hd cs \ orbit (?fc cs) (hd cs)" "\x. x \ orbit (?fc cs) (hd cs) \ ?fc cs x = x" "hd cs \ set cs" "set cs = orbit (?fc cs) (hd cs)" by (auto simp: cyclic_on_alldef) then have "\xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs" by (metis obtain_orbit_list_func) then have "?fc cs = list_succ (some_list cs) \ set (some_list cs) = set cs \ distinct (some_list cs) \ hd (some_list cs) = hd cs" unfolding some_list_def by (rule someI_ex) then have "?fc cs = list_succ (some_list cs)" "set (some_list cs) = set cs" "distinct (some_list cs)" "hd (some_list cs) = hd cs" by auto } note sl_cs = this have "\cs. cs \ set css \ cs \ []" using dists by (auto simp: distincts_def) then have some_list_ne: "\cs. cs \ set css \ some_list cs \ []" by (metis set_empty sl_cs(2)) have set: "map set (map some_list css) = map set css" "map hd (map some_list css) = map hd css" using sl_cs(2,4) by (auto simp add: map_idI) have distincts: "distincts (map some_list css)" proof - have c_dist: "\xs ys. \xs\set css; ys\set css; xs \ ys\ \ set xs \ set ys = {}" using dists by (auto simp: distincts_def) have "distinct (map some_list css)" proof - have "inj_on some_list (set css)" using sl_cs(2) c_dist by (intro inj_onI) (metis inf.idem set_empty) with \distincts css\ show ?thesis by (auto simp: distincts_distinct distinct_map) qed moreover have "\xs\set (map some_list css). distinct xs \ xs \ []" using sl_cs(3) some_list_ne by auto moreover from c_dist have "(\xs\set (map some_list css). \ys\set (map some_list css). xs \ ys \ set xs \ set ys = {})" using sl_cs(2) by auto ultimately show ?thesis by (simp add: distincts_def) qed have f: "f = lists_succ (map some_list css)" using distincts proof (rule permutes_as_lists_succ) fix xs assume "xs \ set (map some_list css)" then show "list_succ xs = perm_restrict f (set xs)" using sl_cs(1) sl_cs(2) by auto next have "S = (\xs\set (map some_list css). set xs)" using S sl_cs(2) by auto with permutes show "f permutes \(sset (map some_list css))" by simp qed from f distincts set show ?thesis .. qed subsection \Permutations of a List\ lemma length_remove1_less: assumes "x \ set xs" shows "length (remove1 x xs) < length xs" proof - from assms have "0 < length xs" by auto with assms show ?thesis by (auto simp: length_remove1) qed context notes [simp] = length_remove1_less begin fun permutations :: "'a list \ 'a list list" where permutations_Nil: "permutations [] = [[]]" | permutations_Cons: "permutations xs = [y # ys. y <- xs, ys <- permutations (remove1 y xs)]" end declare permutations_Cons[simp del] text \ The function above returns all permutations of a list. The function below computes only those which yield distinct cyclic permutation functions (cf. @{term list_succ}). \ fun cyc_permutations :: "'a list \ 'a list list" where "cyc_permutations [] = [[]]" | "cyc_permutations (x # xs) = map (Cons x) (permutations xs)" lemma nil_in_permutations[simp]: "[] \ set (permutations xs) \ xs = []" by (induct xs) (auto simp: permutations_Cons) lemma permutations_not_nil: assumes "xs \ []" shows "permutations xs = concat (map (\x. map ((#) x) (permutations (remove1 x xs))) xs)" using assms by (cases xs) (auto simp: permutations_Cons) lemma set_permutations_step: assumes "xs \ []" shows "set (permutations xs) = (\x \ set xs. Cons x ` set (permutations (remove1 x xs)))" using assms by (cases xs) (auto simp: permutations_Cons) lemma in_set_permutations: assumes "distinct xs" shows "ys \ set (permutations xs) \ distinct ys \ set xs = set ys" (is "?L xs ys \ ?R xs ys") using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by auto next case (Suc n) then have "xs \ []" by auto show ?case proof assume "?L xs ys" then obtain y ys' where "ys = y # ys'" "y \ set xs" "ys' \ set (permutations (remove1 (hd ys) xs))" using \xs \ []\ by (auto simp: permutations_not_nil) moreover then have "?R (remove1 y xs) ys'" using Suc.prems Suc.hyps(2) by (intro Suc.hyps(1)[THEN iffD1]) (auto simp: length_remove1) ultimately show "?R xs ys" using Suc by auto next assume "?R xs ys" with \xs \ []\ obtain y ys' where "ys = y # ys'" "y \ set xs" by (cases ys) auto moreover then have "ys' \ set (permutations (remove1 y xs))" using Suc \?R xs ys\ by (intro Suc.hyps(1)[THEN iffD2]) (auto simp: length_remove1) ultimately show "?L xs ys" using \xs \ []\ by (auto simp: permutations_not_nil) qed qed lemma in_set_cyc_permutations: assumes "distinct xs" shows "ys \ set (cyc_permutations xs) \ distinct ys \ set xs = set ys \ hd ys = hd xs" (is "?L xs ys \ ?R xs ys") proof (cases xs) case (Cons x xs) with assms show ?thesis by (cases ys) (auto simp: in_set_permutations intro!: imageI) qed auto lemma in_set_cyc_permutations_obtain: assumes "distinct xs" "distinct ys" "set xs = set ys" obtains n where "rotate n ys \ set (cyc_permutations xs)" proof (cases xs) case Nil with assms have "rotate 0 ys \ set (cyc_permutations xs)" by auto then show ?thesis .. next case (Cons x xs') let ?ys' = "rotate (index ys x) ys" have "ys \ []" "x \ set ys" using Cons assms by auto then have "distinct ?ys' \ set xs = set ?ys' \ hd ?ys' = hd xs" using assms Cons by (auto simp add: hd_rotate_conv_nth) with \distinct xs\ have "?ys' \ set (cyc_permutations xs)" by (rule in_set_cyc_permutations[THEN iffD2]) then show ?thesis .. qed lemma list_succ_set_cyc_permutations: assumes "distinct xs" "xs \ []" shows "list_succ ` set (cyc_permutations xs) = {f. f permutes set xs \ cyclic_on f (set xs)}" (is "?L = ?R") proof (intro set_eqI iffI) fix f assume "f \ ?L" moreover have "\ys. set xs = set ys \ xs \ [] \ ys \ []" by auto ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutations list_succ_permutes cyclic_on_list_succ) next fix f assume "f \ ?R" then obtain ys where ys: "list_succ ys = f" "distinct ys" "set ys = set xs" by (auto elim: cyclic_on_obtain_list_succ') moreover with \distinct xs\ obtain n where "rotate n ys \ set (cyc_permutations xs)" by (auto elim: in_set_cyc_permutations_obtain) then have "list_succ (rotate n ys) \ ?L" by simp ultimately show "f \ ?L" by simp qed subsection \Enumerating Permutations from List Orbits\ definition cyc_permutationss :: "'a list list \ 'a list list list" where "cyc_permutationss = product_lists o map cyc_permutations" lemma cyc_permutationss_Nil[simp]: "cyc_permutationss [] = [[]]" by (auto simp: cyc_permutationss_def) lemma in_set_cyc_permutationss: assumes "distincts xss" shows "yss \ set (cyc_permutationss xss) \ distincts yss \ map set xss = map set yss \ map hd xss = map hd yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note Y = this show "?thesis" unfolding cyc_permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma lists_succ_set_cyc_permutationss: assumes "distincts xss" shows "lists_succ ` set (cyc_permutationss xss) = {f. f permutes \(sset xss) \ (\c \ sset xss. cyclic_on f c)}" (is "?L = ?R") using assms proof (intro set_eqI iffI) fix f assume "f \ ?L" then obtain yss where "yss \ set (cyc_permutationss xss)" "f = lists_succ yss" by (rule imageE) moreover from \yss \ _\ assms have "set (map set xss) = set (map set yss)" by (auto simp: in_set_cyc_permutationss) then have "sset xss = sset yss" by simp ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutationss cyclic_on_lists_succ') (metis lists_succ_permutes) next fix f assume "f \ ?R" then have "f permutes \(sset xss)" "\cs. cs \ set xss \ cyclic_on f (set cs)" by auto from this(1) refl assms this(2) obtain yss where "f = lists_succ yss" "distincts yss" "map set yss = map set xss" "map hd yss = map hd xss" by (rule cyclic_on_obtain_lists_succ) with assms show "f \ ?L" by (auto intro!: imageI simp: in_set_cyc_permutationss) qed subsection \Lists of Permutations\ definition permutationss :: "'a list list \ 'a list list list" where "permutationss = product_lists o map permutations" lemma permutationss_Nil[simp]: "permutationss [] = [[]]" by (auto simp: permutationss_def) lemma permutationss_Cons: "permutationss (xs # xss) = concat (map (\ys. map (Cons ys) (permutationss xss)) (permutations xs))" by (auto simp: permutationss_def) lemma in_set_permutationss: assumes "distincts xss" shows "yss \ set (permutationss xss) \ distincts yss \ map set xss = map set yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: in_set_permutations distincts_Cons) } note Y = this show "?thesis" unfolding permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma set_permutationss: assumes "distincts xss" shows "set (permutationss xss) = {yss. distincts yss \ map set xss = map set yss}" using in_set_permutationss[OF assms] by blast lemma permutationss_complete: assumes "distincts xss" "distincts yss" "xss \ []" and "set ` set xss = set ` set yss" shows "set yss \ set ` set (permutationss xss)" proof - have "length xss = length yss" using assms by (simp add: length_distincts) from \sset xss = _\ have "\yss'. set yss' = set yss \ map set yss' = map set xss" using assms(1-2) proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) from \sset (xs # xss) = sset yss\ obtain ys where ys: "ys \ set yss" "set ys = set xs" by auto (metis imageE insertI1) with \distincts yss\ have "set ys \ sset (remove1 ys yss)" by (fastforce simp: distincts_def) moreover from \distincts (xs # xss)\ have "set xs \ sset xss" by (fastforce simp: distincts_def) ultimately have "sset xss = sset (remove1 ys yss)" using \distincts yss\ \sset (xs # xss) = sset yss\ apply (auto simp: distincts_distinct \set ys = set xs\[symmetric]) apply (smt Diff_insert_absorb \ys \ set yss\ image_insert insert_Diff rev_image_eqI) by (metis \ys \ set yss\ image_eqI insert_Diff insert_iff) then obtain yss' where "set yss' = set (remove1 ys yss) \ map set yss' = map set xss" using Cons by atomize_elim (auto simp: distincts_Cons distincts_remove1) then have "set (ys # yss') = set yss \ map set (ys # yss') = map set (xs # xss)" using ys set_remove1_eq \distincts yss\ by (auto simp: distincts_distinct) then show ?case .. qed then obtain yss' where "set yss' = set yss" "map set yss' = map set xss" by blast then have "distincts yss'" using \distincts xss\ \distincts yss\ unfolding distincts_def by simp_all (metis \length xss = length yss\ card_distinct distinct_card length_map) then have "set yss' \ set ` set (permutationss xss)" using \distincts xss\ \map set yss' = _\ by (auto simp: set_permutationss) then show ?thesis using \set yss' = _\ by auto qed lemma permutations_complete: (* could generalize with multi-sets *) assumes "distinct xs" "distinct ys" "set xs = set ys" shows "ys \ set (permutations xs)" using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by simp next case (Suc n) from Suc.hyps have "xs \ []" by auto then obtain y ys' where [simp]: "ys = y # ys'" "y \ set xs" using Suc.prems by (cases ys) auto have "ys' \ set (permutations (remove1 y xs))" using Suc.prems \Suc n = _\ by (intro Suc.hyps) (simp_all add: length_remove1 ) then show ?case using \xs \ []\ by (auto simp: set_permutations_step) qed end diff --git a/thys/Planarity_Certificates/Planarity/Graph_Genus.thy b/thys/Planarity_Certificates/Planarity/Graph_Genus.thy --- a/thys/Planarity_Certificates/Planarity/Graph_Genus.thy +++ b/thys/Planarity_Certificates/Planarity/Graph_Genus.thy @@ -1,538 +1,538 @@ theory Graph_Genus imports "HOL-Combinatorics.Permutations" Graph_Theory.Graph_Theory begin lemma nat_diff_mod_right: fixes a b c :: nat assumes "b < a" shows "(a - b) mod c = (a - b mod c) mod c" proof - from assms have b_mod: "b mod c \ a" by (metis mod_less_eq_dividend linear not_le order_trans) have "int ((a - b) mod c) = (int a - int b mod int c) mod int c" using assms by (simp add: zmod_int of_nat_diff mod_simps) also have "\ = int ((a - b mod c) mod c)" using assms b_mod by (simp add: zmod_int [symmetric] of_nat_diff [symmetric]) finally show ?thesis by simp qed lemma inj_on_f_imageI: assumes "inj_on f S" "\t. t \ T \ t \ S" shows "inj_on ((`) f) T" using assms by (auto simp: inj_on_image_eq_iff intro: inj_onI) section \Combinatorial Maps\ lemma (in bidirected_digraph) has_dom_arev: "has_dom arev (arcs G)" using arev_dom by (auto simp: has_dom_def) record 'b pre_map = edge_rev :: "'b \ 'b" edge_succ :: "'b \ 'b" definition edge_pred :: "'b pre_map \ 'b \ 'b" where "edge_pred M = inv (edge_succ M)" locale pre_digraph_map = pre_digraph + fixes M :: "'b pre_map" locale digraph_map = fin_digraph G + pre_digraph_map G M + bidirected_digraph G "edge_rev M" for G M + assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" lemma (in fin_digraph) digraph_mapI: assumes bidi: "\a. a \ arcs G \ edge_rev M a = a" "\a. a \ arcs G \ edge_rev M a \ a" "\a. a \ arcs G \ edge_rev M (edge_rev M a) = a" "\a. a \ arcs G \ tail G (edge_rev M a) = head G a" assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" shows "digraph_map G M" using assms by unfold_locales auto lemma (in fin_digraph) digraph_mapI_permutes: assumes bidi: "edge_rev M permutes arcs G" "\a. a \ arcs G \ edge_rev M a \ a" "\a. a \ arcs G \ edge_rev M (edge_rev M a) = a" "\a. a \ arcs G \ tail G (edge_rev M a) = head G a" assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" shows "digraph_map G M" proof - interpret bidirected_digraph G "edge_rev M" using bidi by unfold_locales (auto simp: permutes_def) show ?thesis using edge_succ_permutes edge_succ_cyclic by unfold_locales qed context digraph_map begin lemma digraph_map[intro]: "digraph_map G M" by unfold_locales lemma permutation_edge_succ: "permutation (edge_succ M)" by (metis edge_succ_permutes finite_arcs permutation_permutes) lemma edge_pred_succ[simp]: "edge_pred M (edge_succ M a) = a" by (metis edge_pred_def edge_succ_permutes permutes_inverses(2)) lemma edge_succ_pred[simp]: "edge_succ M (edge_pred M a) = a" by (metis edge_pred_def edge_succ_permutes permutes_inverses(1)) lemma edge_pred_permutes: "edge_pred M permutes arcs G" unfolding edge_pred_def using edge_succ_permutes by (rule permutes_inv) lemma permutation_edge_pred: "permutation (edge_pred M)" by (metis edge_pred_permutes finite_arcs permutation_permutes) lemma edge_succ_eq_iff[simp]: "\x y. edge_succ M x = edge_succ M y \ x = y" by (metis edge_pred_succ) lemma edge_rev_in_arcs[simp]: "edge_rev M a \ arcs G \ a \ arcs G" by (metis arev_arev arev_permutes_arcs permutes_not_in) lemma edge_succ_in_arcs[simp]: "edge_succ M a \ arcs G \ a \ arcs G" by (metis edge_pred_succ edge_succ_permutes permutes_not_in) lemma edge_pred_in_arcs[simp]: "edge_pred M a \ arcs G \ a \ arcs G" by (metis edge_succ_pred edge_pred_permutes permutes_not_in) lemma tail_edge_succ[simp]: "tail G (edge_succ M a) = tail G a" proof cases assume "a \ arcs G" then have "tail G a \ verts G" by auto moreover then have "out_arcs G (tail G a) \ {}" using \a \ arcs G\ by auto ultimately have "cyclic_on (edge_succ M) (out_arcs G (tail G a))" by (rule edge_succ_cyclic) moreover have "a \ out_arcs G (tail G a)" using \a \ arcs G\ by simp ultimately have "edge_succ M a \ out_arcs G (tail G a)" by (rule cyclic_on_inI) then show ?thesis by simp next assume "a \ arcs G" then show ?thesis using edge_succ_permutes by (simp add: permutes_not_in) qed lemma tail_edge_pred[simp]: "tail G (edge_pred M a) = tail G a" by (metis edge_succ_pred tail_edge_succ) lemma bij_edge_succ[intro]: "bij (edge_succ M)" using edge_succ_permutes by (simp add: permutes_conv_has_dom) lemma edge_pred_cyclic: assumes "v \ verts G" "out_arcs G v \ {}" shows "cyclic_on (edge_pred M) (out_arcs G v)" proof - obtain a where orb_a_eq: "orbit (edge_succ M) a = out_arcs G v" using edge_succ_cyclic[OF assms] by (auto simp: cyclic_on_def) have "cyclic_on (edge_pred M) (orbit (edge_pred M) a)" using permutation_edge_pred by (rule cyclic_on_orbit') also have "orbit (edge_pred M) a = orbit (edge_succ M) a" unfolding edge_pred_def using permutation_edge_succ by (rule orbit_inv_eq) finally show "cyclic_on (edge_pred M) (out_arcs G v)" by (simp add: orb_a_eq) qed definition (in pre_digraph_map) face_cycle_succ :: "'b \ 'b" where "face_cycle_succ \ edge_succ M o edge_rev M" definition (in pre_digraph_map) face_cycle_pred :: "'b \ 'b" where "face_cycle_pred \ edge_rev M o edge_pred M" lemma face_cycle_pred_succ[simp]: shows "face_cycle_pred (face_cycle_succ a) = a" unfolding face_cycle_pred_def face_cycle_succ_def by simp lemma face_cycle_succ_pred[simp]: shows "face_cycle_succ (face_cycle_pred a) = a" unfolding face_cycle_pred_def face_cycle_succ_def by simp lemma tail_face_cycle_succ: "a \ arcs G \ tail G (face_cycle_succ a) = head G a" by (auto simp: face_cycle_succ_def) lemma funpow_prop: assumes "\x. P (f x) \ P x" shows "P ((f ^^ n) x) \ P x" using assms by (induct n) (auto simp: ) lemma face_cycle_succ_no_arc[simp]: "a \ arcs G \ face_cycle_succ a = a" by (auto simp: face_cycle_succ_def permutes_not_in[OF arev_permutes_arcs] permutes_not_in[OF edge_succ_permutes]) lemma funpow_face_cycle_succ_no_arc[simp]: assumes "a \ arcs G" shows "(face_cycle_succ ^^ n) a = a" using assms by (induct n) auto lemma funpow_face_cycle_pred_no_arc[simp]: assumes "a \ arcs G" shows "(face_cycle_pred ^^ n) a = a" using assms by (induct n) (auto simp: face_cycle_pred_def permutes_not_in[OF arev_permutes_arcs] permutes_not_in[OF edge_pred_permutes]) lemma face_cycle_succ_closed[simp]: "face_cycle_succ a \ arcs G \ a \ arcs G" by (metis comp_apply edge_rev_in_arcs edge_succ_in_arcs face_cycle_succ_def) lemma face_cycle_pred_closed[simp]: "face_cycle_pred a \ arcs G \ a \ arcs G" by (metis face_cycle_succ_closed face_cycle_succ_pred) lemma face_cycle_succ_permutes: "face_cycle_succ permutes arcs G" unfolding face_cycle_succ_def using arev_permutes_arcs edge_succ_permutes by (rule permutes_compose) lemma permutation_face_cycle_succ: "permutation face_cycle_succ" using face_cycle_succ_permutes finite_arcs by (metis permutation_permutes) lemma bij_face_cycle_succ: "bij face_cycle_succ" using face_cycle_succ_permutes by (simp add: permutes_conv_has_dom) lemma face_cycle_pred_permutes: "face_cycle_pred permutes arcs G" unfolding face_cycle_pred_def using edge_pred_permutes arev_permutes_arcs by (rule permutes_compose) definition (in pre_digraph_map) face_cycle_set :: "'b \ 'b set" where "face_cycle_set a = orbit face_cycle_succ a" definition (in pre_digraph_map) face_cycle_sets :: "'b set set" where "face_cycle_sets = face_cycle_set ` arcs G" lemma face_cycle_set_altdef: "face_cycle_set a = {(face_cycle_succ ^^ n) a | n. True}" unfolding face_cycle_set_def by (intro orbit_altdef_self_in permutation_self_in_orbit permutation_face_cycle_succ) lemma face_cycle_set_self[simp, intro]: "a \ face_cycle_set a" unfolding face_cycle_set_def using permutation_face_cycle_succ by (rule permutation_self_in_orbit) lemma empty_not_in_face_cycle_sets: "{} \ face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma finite_face_cycle_set[simp, intro]: "finite (face_cycle_set a)" using face_cycle_set_self unfolding face_cycle_set_def by (simp add: finite_orbit) lemma finite_face_cycle_sets[simp, intro]: "finite face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma face_cycle_set_induct[case_names base step, induct set: face_cycle_set]: assumes consume: "a \ face_cycle_set x" and ih_base: "P x" and ih_step: "\y. y \ face_cycle_set x \ P y \ P (face_cycle_succ y)" shows "P a" using consume unfolding face_cycle_set_def by induct (auto simp: ih_step face_cycle_set_def[symmetric] ih_base ) lemma face_cycle_succ_cyclic: "cyclic_on face_cycle_succ (face_cycle_set a)" unfolding face_cycle_set_def using permutation_face_cycle_succ by (rule cyclic_on_orbit') lemma face_cycle_eq: assumes "b \ face_cycle_set a" shows "face_cycle_set b = face_cycle_set a" using assms unfolding face_cycle_set_def by (auto intro: orbit_swap orbit_trans permutation_face_cycle_succ permutation_self_in_orbit) lemma face_cycle_succ_in_arcsI: "\a. a \ arcs G \ face_cycle_succ a \ arcs G" by (auto simp: face_cycle_succ_def) lemma face_cycle_succ_inI: "\x y. x \ face_cycle_set y \ face_cycle_succ x \ face_cycle_set y" by (metis face_cycle_succ_cyclic cyclic_on_inI) lemma face_cycle_succ_inD: "\x y. face_cycle_succ x \ face_cycle_set y \ x \ face_cycle_set y" by (metis face_cycle_eq face_cycle_set_self face_cycle_succ_inI) lemma face_cycle_set_parts: "face_cycle_set a = face_cycle_set b \ face_cycle_set a \ face_cycle_set b = {}" by (metis disjoint_iff_not_equal face_cycle_eq) definition fc_equiv :: "'b \ 'b \ bool" where "fc_equiv a b \ a \ face_cycle_set b" lemma reflp_fc_equiv: "reflp fc_equiv" by (rule reflpI) (simp add: fc_equiv_def) lemma symp_fc_equiv: "symp fc_equiv" using face_cycle_set_parts by (intro sympI) (auto simp: fc_equiv_def) lemma transp_fc_equiv: "transp fc_equiv" using face_cycle_set_parts by (intro transpI) (auto simp: fc_equiv_def) lemma "equivp fc_equiv" by (intro equivpI reflp_fc_equiv symp_fc_equiv transp_fc_equiv) lemma in_face_cycle_setD: assumes "y \ face_cycle_set x" "x \ arcs G" shows "y \ arcs G" using assms by (auto simp: face_cycle_set_def dest: permutes_orbit_subset[OF face_cycle_succ_permutes]) lemma in_face_cycle_setsD: assumes "x \ face_cycle_sets" shows "x \ arcs G" using assms by (auto simp: face_cycle_sets_def dest: in_face_cycle_setD) end definition (in pre_digraph) isolated_verts :: "'a set" where "isolated_verts \ {v \ verts G. out_arcs G v = {}}" definition (in pre_digraph_map) euler_char :: int where "euler_char \ int (card (verts G)) - int (card (arcs G) div 2) + int (card face_cycle_sets)" definition (in pre_digraph_map) euler_genus :: int where "euler_genus \ (int (2 * card sccs) - int (card isolated_verts) - euler_char) div 2" definition comb_planar :: "('a,'b) pre_digraph \ bool" where "comb_planar G \ \M. digraph_map G M \ pre_digraph_map.euler_genus G M = 0" text \Number of isolated vertices is a graph invariant\ context fixes G hom assumes hom: "pre_digraph.digraph_isomorphism G hom" begin interpretation wf_digraph G using hom by (auto simp: pre_digraph.digraph_isomorphism_def) lemma isolated_verts_app_iso[simp]: "pre_digraph.isolated_verts (app_iso hom G) = iso_verts hom ` isolated_verts" using hom by (auto simp: pre_digraph.isolated_verts_def iso_verts_tail inj_image_mem_iff out_arcs_app_iso_eq) lemma card_isolated_verts_iso[simp]: "card (iso_verts hom ` pre_digraph.isolated_verts G) = card isolated_verts" apply (rule card_image) using hom apply (rule digraph_isomorphism_inj_on_verts[THEN subset_inj_on]) apply (auto simp: isolated_verts_def) done end context digraph_map begin lemma face_cycle_succ_neq: assumes "a \ arcs G" "tail G a \ head G a" shows "face_cycle_succ a \ a " proof - from assms have "edge_rev M a \ arcs G" by (subst edge_rev_in_arcs) simp then have "cyclic_on (edge_succ M) (out_arcs G (tail G (edge_rev M a)))" by (intro edge_succ_cyclic) (auto dest: tail_in_verts simp: out_arcs_def intro: exI[where x="edge_rev M a"]) then have "edge_succ M (edge_rev M a) \ (out_arcs G (tail G (edge_rev M a)))" by (rule cyclic_on_inI) (auto simp: \edge_rev M a \ _\[simplified]) moreover have "tail G (edge_succ M (edge_rev M a)) = head G a" using assms by auto then have "edge_succ M (edge_rev M a) \ a" using assms by metis ultimately show ?thesis using assms by (auto simp: face_cycle_succ_def) qed end section \Maps and Isomorphism\ definition (in pre_digraph) "wrap_iso_arcs hom f = perm_restrict (iso_arcs hom o f o iso_arcs (inv_iso hom)) (arcs (app_iso hom G))" definition (in pre_digraph_map) map_iso :: "('a,'b,'a2,'b2) digraph_isomorphism \ 'b2 pre_map" where "map_iso f \ \ edge_rev = wrap_iso_arcs f (edge_rev M) , edge_succ = wrap_iso_arcs f (edge_succ M) \" lemma funcsetI_permutes: assumes "f permutes S" shows "f \ S \ S" by (metis assms funcsetI permutes_in_image) context fixes G hom assumes hom: "pre_digraph.digraph_isomorphism G hom" begin interpretation wf_digraph G using hom by (auto simp: pre_digraph.digraph_isomorphism_def) lemma wrap_iso_arcs_iso_arcs[simp]: assumes "x \ arcs G" shows "wrap_iso_arcs hom f (iso_arcs hom x) = iso_arcs hom (f x)" using assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) lemma inj_on_wrap_iso_arcs: assumes dom: "\f. f \ F \ has_dom f (arcs G)" assumes funcset: "F \ arcs G \ arcs G" shows "inj_on (wrap_iso_arcs hom) F" proof (rule inj_onI) fix f g assume F: "f \ F" "g \ F" and eq: "wrap_iso_arcs hom f = wrap_iso_arcs hom g" { fix x assume "x \ arcs G" then have "f x = x" "g x = x" using F dom by (auto simp: has_dom_def) then have "f x = g x" by simp } moreover { fix x assume "x \ arcs G" then have "f x \ arcs G" "g x \ arcs G" using F funcset by auto with digraph_isomorphism_inj_on_arcs[OF hom] _ have "iso_arcs hom (f x) = iso_arcs hom (g x) \ f x = g x" by (rule inj_onD) then have "f x = g x" using assms hom \x \ arcs G\ eq by (auto simp: wrap_iso_arcs_def fun_eq_iff perm_restrict_def split: if_splits) } ultimately show "f = g" by auto qed lemma inj_on_wrap_iso_arcs_f: assumes "A \ arcs G" "f \ A \ A" "B = iso_arcs hom ` A" assumes "inj_on f A" shows "inj_on (wrap_iso_arcs hom f) B" proof (rule inj_onI) fix x y assume in_hom_A: "x \ B" "y \ B" and wia_eq: "wrap_iso_arcs hom f x = wrap_iso_arcs hom f y" from in_hom_A \B = _\ obtain x0 where x0: "x = iso_arcs hom x0" "x0 \ A" by auto from in_hom_A \B = _\ obtain y0 where y0: "y = iso_arcs hom y0" "y0 \ A" by auto have arcs_0: "x0 \ arcs G" "y0 \ arcs G" "f x0 \ arcs G" "f y0 \ arcs G" using x0 y0 \A \ _\ \f \ _\ by auto have "(iso_arcs hom o f o iso_arcs (inv_iso hom)) x = (iso_arcs hom o f o iso_arcs (inv_iso hom)) y" using in_hom_A wia_eq assms(1) \B = _\ by (auto simp: wrap_iso_arcs_def perm_restrict_def split: if_splits) then show "x = y" using hom assms digraph_isomorphism_inj_on_arcs[OF hom] x0 y0 arcs_0 \inj_on f A\ \A \ _\ by (auto dest!: inj_onD) qed lemma wrap_iso_arcs_in_funcsetI: assumes "A \ arcs G" "f \ A \ A" shows "wrap_iso_arcs hom f \ iso_arcs hom ` A \ iso_arcs hom ` A" proof fix x assume "x \ iso_arcs hom ` A" then obtain x0 where "x = iso_arcs hom x0" "x0 \ A" by blast then have "f x0 \ A" using \f \ _\ by auto then show "wrap_iso_arcs hom f x \ iso_arcs hom ` A" unfolding \x = _\ using \x0 \ A\ assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) qed lemma wrap_iso_arcs_permutes: assumes "A \ arcs G" "f permutes A" shows "wrap_iso_arcs hom f permutes (iso_arcs hom ` A)" proof - { fix x assume A: "x \ iso_arcs hom ` A" have "wrap_iso_arcs hom f x = x" proof cases assume "x \ iso_arcs hom ` arcs G" then have "iso_arcs (inv_iso hom) x \ A" "x \ arcs (app_iso hom G)" using A hom by (metis arcs_app_iso image_eqI pre_digraph.iso_arcs_iso_inv, simp) then have "f (iso_arcs (inv_iso hom) x) = (iso_arcs (inv_iso hom) x)" using \f permutes A\ by (simp add: permutes_not_in) then show ?thesis using hom assms \x \ arcs _\ by (simp add: wrap_iso_arcs_def perm_restrict_def) next assume "x \ iso_arcs hom ` arcs G" then show ?thesis by (simp add: wrap_iso_arcs_def perm_restrict_def) qed } note not_in_id = this have "f \ A \ A" using assms by (intro funcsetI_permutes) have inj_on_wrap: "inj_on (wrap_iso_arcs hom f) (iso_arcs hom ` A)" using assms \f \ A \ A\ by (intro inj_on_wrap_iso_arcs_f) (auto intro: subset_inj_on permutes_inj) have woa_in_fs: "wrap_iso_arcs hom f \ iso_arcs hom ` A \ iso_arcs hom ` A" using assms \f \ A \ A\ by (intro wrap_iso_arcs_in_funcsetI) { fix x y assume "wrap_iso_arcs hom f x = wrap_iso_arcs hom f y" then have "x = y" apply (cases "x \ iso_arcs hom ` A"; cases "y \ iso_arcs hom ` A") using woa_in_fs inj_on_wrap by (auto dest: inj_onD simp: not_in_id) } note uniqueD = this note \f permutes A\ moreover note not_in_id moreover { fix y have "\x. wrap_iso_arcs hom f x = y" proof cases assume "y \ iso_arcs hom ` A" then obtain y0 where "y0 \ A" "iso_arcs hom y0 = y" by blast with \f permutes A\ obtain x0 where "x0 \ A" "f x0 = y0" unfolding permutes_def by metis moreover then have "\x. x \ arcs G \ iso_arcs hom x0 = iso_arcs hom x \ x0 = x" using assms hom by (auto simp: digraph_isomorphism_def dest: inj_onD) ultimately have "wrap_iso_arcs hom f (iso_arcs hom x0) = y" using \_ = y\ assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) then show ?thesis .. qed (metis not_in_id) } ultimately show ?thesis unfolding permutes_def by (auto simp: dest: uniqueD) qed end lemma (in digraph_map) digraph_map_isoI: assumes "digraph_isomorphism hom" shows "digraph_map (app_iso hom G) (map_iso hom)" proof - interpret iG: fin_digraph "app_iso hom G" using assms by (rule fin_digraphI_app_iso) show ?thesis proof (rule iG.digraph_mapI_permutes) show "edge_rev (map_iso hom) permutes arcs (app_iso hom G)" using assms unfolding map_iso_def by (simp add: wrap_iso_arcs_permutes arev_permutes_arcs) next show "edge_succ (map_iso hom) permutes arcs (app_iso hom G)" using assms unfolding map_iso_def by (simp add: wrap_iso_arcs_permutes edge_succ_permutes) next fix a assume A: "a \ arcs (app_iso hom G)" show "tail (app_iso hom G) (edge_rev (map_iso hom) a) = head (app_iso hom G) a" using A assms by (cases rule: in_arcs_app_iso_cases) (auto simp: map_iso_def iso_verts_tail iso_verts_head) show "edge_rev (map_iso hom) (edge_rev (map_iso hom) a) = a" using A assms by (cases rule: in_arcs_app_iso_cases) (auto simp: map_iso_def) show "edge_rev (map_iso hom) a \ a" using A assms by (auto simp: map_iso_def arev_neq) next fix v assume "v \ verts (app_iso hom G)" and oa_hom: "out_arcs (app_iso hom G) v \ {}" then obtain v0 where "v0 \ verts G" "v = iso_verts hom v0" by auto moreover then have oa: "out_arcs G v0 \ {}" using assms oa_hom by (auto simp: out_arcs_def iso_verts_tail) ultimately have cyclic_on_v0: "cyclic_on (edge_succ M) (out_arcs G v0)" by (intro edge_succ_cyclic) from oa_hom obtain a where "a \ out_arcs (app_iso hom G) v" by blast then obtain a0 where "a0 \ arcs G" "a = iso_arcs hom a0" by auto then have "a0 \ out_arcs G v0" using \v = _\ \v0 \ _\ \a \ _\ assms by (simp add: iso_verts_tail) show "cyclic_on (edge_succ (map_iso hom)) (out_arcs (app_iso hom G) v)" proof (rule cyclic_on_singleI) show "a \ out_arcs (app_iso hom G) v" by fact next have "out_arcs (app_iso hom G) v = iso_arcs hom ` out_arcs G v0" unfolding \v = _\ by (rule out_arcs_app_iso_eq) fact+ also have "out_arcs G v0 = orbit (edge_succ M) a0" using cyclic_on_v0 \a0 \ out_arcs G v0\ unfolding cyclic_on_alldef by simp also have "iso_arcs hom ` \ = orbit (edge_succ (map_iso hom)) a" proof - have "\x. x \ orbit (edge_succ M) a0 \ x \ arcs G" using \out_arcs G v0 = _\ by auto then show ?thesis using \out_arcs G v0 = _\ - unfolding \a = _\ assms using \a0 \ out_arcs G v0\ - by (intro orbit_FOO) (insert assms, auto simp: map_iso_def) + unfolding \a = _\ using \a0 \ out_arcs G v0\ assms + by (intro orbit_inverse) (auto simp: map_iso_def) qed finally show "out_arcs (app_iso hom G) v = orbit (edge_succ (map_iso hom)) a" . qed qed qed end diff --git a/thys/Planarity_Certificates/Planarity/Permutations_2.thy b/thys/Planarity_Certificates/Planarity/Permutations_2.thy --- a/thys/Planarity_Certificates/Planarity/Permutations_2.thy +++ b/thys/Planarity_Certificates/Planarity/Permutations_2.thy @@ -1,153 +1,153 @@ theory Permutations_2 imports "HOL-Combinatorics.Permutations" + Graph_Theory.Auxiliary Executable_Permutations - Graph_Theory.Funpow begin section \Modifying Permutations\ abbreviation funswapid :: "'a \ 'a \ 'a \ 'a" (infix "\\<^sub>F" 90) where "x \\<^sub>F y \ Fun.swap x y id" definition perm_swap :: "'a \ 'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_swap x y f \ x \\<^sub>F y o f o x \\<^sub>F y" definition perm_rem :: "'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_rem x f \ if f x \ x then x \\<^sub>F f x o f else f" text \ An example: @{lemma "perm_rem (2 :: nat) (list_succ [1,2,3,4]) x = list_succ [1,3,4] x" by (auto simp: perm_rem_def Fun.swap_def list_succ_def)} \ lemma perm_swap_id[simp]: "perm_swap a b id = id" by (auto simp: perm_swap_def) lemma perm_rem_permutes: assumes "f permutes S \ {x}" shows "perm_rem x f permutes S" using assms by (auto simp: permutes_def perm_rem_def) (metis swap_id_eq)+ lemma perm_rem_same: assumes "bij f" "f y = y" shows "perm_rem x f y = f y" using assms by (auto simp: perm_rem_def swap_id_eq bij_iff) lemma perm_rem_simps: assumes "bij f" shows "x = y \ perm_rem x f y = x" "f y = x \ perm_rem x f y = f x" "y \ x \ f y \ x \ perm_rem x f y = f y" using assms apply (auto simp: perm_rem_def ) by (metis bij_iff id_apply swap_apply(3)) lemma bij_swap_compose: "bij (x \\<^sub>F y o f) \ bij f" by (metis UNIV_I bij_betw_comp_iff2 bij_betw_id bij_swap_iff subsetI) lemma bij_perm_rem[simp]: "bij (perm_rem x f) \ bij f" by (simp add: perm_rem_def bij_swap_compose) lemma perm_rem_conv: "\f x y. bij f \ perm_rem x f y = ( if x = y then x else if f y = x then f (f y) else f y)" by (auto simp: perm_rem_simps) lemma perm_rem_commutes: assumes "bij f" shows "perm_rem a (perm_rem b f) = perm_rem b (perm_rem a f)" proof - have bij_simp: "\x y. f x = f y \ x = y" using assms by (auto simp: bij_iff) show ?thesis using assms by (auto simp: perm_rem_conv bij_simp fun_eq_iff) qed lemma perm_rem_id[simp]: "perm_rem a id = id" by (simp add: perm_rem_def) lemma bij_eq_iff: assumes "bij f" shows "f x = f y \ x = y" using assms by (metis bij_iff) lemma swap_swap_id[simp]: "(x \\<^sub>F y) ((x \\<^sub>F y) z) = z" by (simp add: swap_id_eq) lemma in_funswapid_image_iff: "\a b x S. x \ (a \\<^sub>F b) ` S \ (a \\<^sub>F b) x \ S" by (metis bij_def bij_id bij_swap_iff inj_image_mem_iff swap_swap_id) lemma perm_swap_comp: "perm_swap a b (f \ g) x = perm_swap a b f (perm_swap a b g x)" by (auto simp: perm_swap_def) lemma bij_perm_swap_iff[simp]: "bij (perm_swap a b f) \ bij f" by (auto simp: perm_swap_def bij_swap_compose bij_comp comp_swap) lemma funpow_perm_swap: "perm_swap a b f ^^ n = perm_swap a b (f ^^ n)" by (induct n) (auto simp: perm_swap_def fun_eq_iff) lemma orbit_perm_swap: "orbit (perm_swap a b f) x = (a \\<^sub>F b) ` orbit f ((a \\<^sub>F b) x)" by (auto simp: orbit_altdef funpow_perm_swap) (auto simp: perm_swap_def) lemma has_dom_perm_swap: "has_dom (perm_swap a b f) S = has_dom f ((a \\<^sub>F b) ` S)" by (auto simp: has_dom_def perm_swap_def inj_image_mem_iff) (metis image_iff swap_swap_id) lemma perm_restrict_dom_subset: assumes "has_dom f A" shows "perm_restrict f A = f" proof - from assms have "\x. x \ A \ f x = x" by (auto simp: has_dom_def) then show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) qed lemma has_domD: "has_dom f S \ x \ S \ f x = x" by (auto simp: has_dom_def) lemma has_domI: "(\x. x \ S \ f x = x) \ has_dom f S" by (auto simp: has_dom_def) lemma perm_swap_permutes2: assumes "f permutes ((x \\<^sub>F y) ` S)" shows "perm_swap x y f permutes S" using assms by (auto simp: perm_swap_def permutes_conv_has_dom has_dom_perm_swap[unfolded perm_swap_def]) (metis bij_swap_iff bij_swap_compose_bij comp_id comp_swap) section \Cyclic Permutations\ lemma cyclic_on_perm_swap: assumes "cyclic_on f S" shows "cyclic_on (perm_swap x y f) ((x \\<^sub>F y) ` S)" - using assms by (rule cyclic_on_FOO) (auto simp: perm_swap_def swap_swap_id) + using assms by (rule cyclic_on_image) (auto simp: perm_swap_def) lemma orbit_perm_rem: assumes "bij f" "x \ y" shows "orbit (perm_rem y f) x = orbit f x - {y}" (is "?L = ?R") proof (intro set_eqI iffI) fix z assume "z \ ?L" then show "z \ ?R" using assms by induct (auto simp: perm_rem_conv bij_iff intro: orbit.intros) next fix z assume A: "z \ ?R" { assume "z \ orbit f x" then have "(z \ y \ z \ ?L) \ (z = y \ f z \ ?L)" proof induct case base with assms show ?case by (auto intro: orbit_eqI(1) simp: perm_rem_conv) next case (step z) then show ?case using assms by (cases "y = z") (auto intro: orbit_eqI simp: perm_rem_conv) qed } with A show "z \ ?L" by auto qed lemma orbit_perm_rem_eq: assumes "bij f" shows "orbit (perm_rem y f) x = (if x = y then {y} else orbit f x - {y})" using assms by (simp add: orbit_eq_singleton_iff orbit_perm_rem perm_rem_simps) lemma cyclic_on_perm_rem: assumes "cyclic_on f S" "bij f" "S \ {x}" shows "cyclic_on (perm_rem x f) (S - {x})" using assms[unfolded cyclic_on_alldef] by (simp add: cyclic_on_def orbit_perm_rem_eq) auto end diff --git a/thys/Planarity_Certificates/Planarity/Planar_Complete.thy b/thys/Planarity_Certificates/Planarity/Planar_Complete.thy --- a/thys/Planarity_Certificates/Planarity/Planar_Complete.thy +++ b/thys/Planarity_Certificates/Planarity/Planar_Complete.thy @@ -1,450 +1,450 @@ theory Planar_Complete imports Digraph_Map_Impl begin section \Kuratowski Graphs are not Combinatorially Planar\ subsection \A concrete K5 graph\ definition "c_K5_list \ ([0..4], [(x,y). x <- [0..4], y <- [0..4], x \ y])" abbreviation c_K5 :: "int pair_pre_digraph" where "c_K5 \ list_digraph c_K5_list" lemma c_K5_not_comb_planar: "\comb_planar c_K5" by (subst comb_planar_impl_correct) eval+ lemma pverts_c_K5: "pverts c_K5 = {0..4}" by (simp add: c_K5_list_def list_digraph_ext_def) lemma parcs_c_K5: "parcs c_K5 = {(u,v). u \ {0..4} \ v \ {0..4} \ u \ v}" by (auto simp: c_K5_list_def list_digraph_ext_def) lemmas c_K5_simps = pverts_c_K5 parcs_c_K5 lemma complete_c_K5: "K\<^bsub>5\<^esub> c_K5" proof - interpret K5: pair_graph c_K5 by eval show ?thesis unfolding complete_digraph_def by (auto simp: c_K5_simps) qed subsection \A concrete K33 graph\ definition "c_K33_list \ ([0..5], [(x,y). x <- [0..5], y <- [0..5], even x \ odd y])" abbreviation c_K33 :: "int pair_pre_digraph" where "c_K33 \ list_digraph c_K33_list" lemma c_K33_not_comb_planar: "\comb_planar c_K33" by (subst comb_planar_impl_correct) eval+ lemma complete_c_K33: "K\<^bsub>3,3\<^esub> c_K33" proof - interpret K33: pair_graph c_K33 by eval show ?thesis unfolding complete_bipartite_digraph_def apply (intro conjI) apply unfold_locales apply (rule exI[of _ "{0,2,4}"]) apply (rule exI[of _ "{1,3,5}"]) unfolding c_K33_list_def list_digraph_simps with_proj_simps apply eval done qed subsection \Generalization to arbitrary Kuratowski Graphs\ subsubsection \Number of Face Cycles is a Graph Invariant\ lemma (in digraph_map) wrap_wrap_iso: assumes hom: "digraph_isomorphism hom" assumes f: "f \ arcs G \ arcs G" and g: "g \ arcs G \ arcs G" shows "wrap_iso_arcs hom f (wrap_iso_arcs hom g x) = wrap_iso_arcs hom (f o g) x" proof - have "\x. x \ arcs G \ g x \ arcs G" using g by auto with hom f show ?thesis by (cases "x \ iso_arcs hom ` arcs G") (auto simp: wrap_iso_arcs_def perm_restrict_simps) qed lemma (in digraph_map) face_cycle_succ_iso: assumes hom: "digraph_isomorphism hom" "x \ iso_arcs hom ` arcs G" shows "pre_digraph_map.face_cycle_succ (map_iso hom) x = wrap_iso_arcs hom face_cycle_succ x" using assms by (simp add: pre_digraph_map.face_cycle_succ_def map_iso_def wrap_wrap_iso) lemma (in digraph_map) face_cycle_set_iso: assumes hom: "digraph_isomorphism hom" "x \ iso_arcs hom ` arcs G" shows "pre_digraph_map.face_cycle_set (map_iso hom) x = iso_arcs hom ` face_cycle_set (iso_arcs (inv_iso hom) x)" proof - have *: "\x y. x \ orbit face_cycle_succ y \ y \ arcs G \ x \ arcs G" "\x. x \ arcs G \ x \ orbit face_cycle_succ x" using face_cycle_set_def by (auto simp: in_face_cycle_setD) show ?thesis using assms unfolding pre_digraph_map.face_cycle_set_def - by (subst orbit_FOO[where g'="pre_digraph_map.face_cycle_succ (map_iso hom)"]) + by (subst orbit_inverse [where g'="pre_digraph_map.face_cycle_succ (map_iso hom)"]) (auto simp: * face_cycle_succ_iso) qed lemma (in digraph_map) face_cycle_sets_iso: assumes hom: "digraph_isomorphism hom" shows "pre_digraph_map.face_cycle_sets (app_iso hom G) (map_iso hom) = (\x. iso_arcs hom ` x) ` face_cycle_sets" using assms by (auto simp: pre_digraph_map.face_cycle_sets_def face_cycle_set_iso) (auto simp: face_cycle_set_iso intro: rev_image_eqI) lemma (in digraph_map) card_face_cycle_sets_iso: assumes hom: "digraph_isomorphism hom" shows "card (pre_digraph_map.face_cycle_sets (app_iso hom G) (map_iso hom)) = card face_cycle_sets" proof - have "inj_on ((`) (iso_arcs hom)) face_cycle_sets" by (rule inj_on_f_imageI digraph_isomorphism_inj_on_arcs hom in_face_cycle_setsD)+ then show ?thesis using hom by (simp add: face_cycle_sets_iso card_image) qed subsubsection \Combinatorial planarity is a Graph Invariant\ lemma (in digraph_map) euler_char_iso: assumes "digraph_isomorphism hom" shows "pre_digraph_map.euler_char (app_iso hom G) (map_iso hom) = euler_char" using assms by (auto simp: pre_digraph_map.euler_char_def card_face_cycle_sets_iso) lemma (in digraph_map) euler_genus_iso: assumes "digraph_isomorphism hom" shows "pre_digraph_map.euler_genus (app_iso hom G) (map_iso hom) = euler_genus" using assms by (auto simp: pre_digraph_map.euler_genus_def euler_char_iso) lemma (in wf_digraph) comb_planar_iso: assumes "digraph_isomorphism hom" shows "comb_planar (app_iso hom G) \ comb_planar G" proof assume "comb_planar G" then obtain M where "digraph_map G M" "pre_digraph_map.euler_genus G M = 0" by (auto simp: comb_planar_def) then have "digraph_map (app_iso hom G) (pre_digraph_map.map_iso G M hom)" "pre_digraph_map.euler_genus (app_iso hom G) (pre_digraph_map.map_iso G M hom) = 0" using assms by (auto intro: digraph_map.digraph_map_isoI simp: digraph_map.euler_genus_iso) then show "comb_planar (app_iso hom G)" by (metis comb_planar_def) next let ?G = "app_iso hom G" assume "comb_planar ?G" then obtain M where "digraph_map ?G M" "pre_digraph_map.euler_genus ?G M = 0" by (auto simp: comb_planar_def) moreover have "pre_digraph.digraph_isomorphism ?G (inv_iso hom)" using assms by (rule digraph_isomorphism_invI) ultimately have "digraph_map (app_iso (inv_iso hom) ?G) (pre_digraph_map.map_iso ?G M (inv_iso hom))" "pre_digraph_map.euler_genus (app_iso (inv_iso hom) ?G) (pre_digraph_map.map_iso ?G M (inv_iso hom)) = 0" using assms by (auto intro: digraph_map.digraph_map_isoI simp only: digraph_map.euler_genus_iso) then show "comb_planar G" using assms by (auto simp: comb_planar_def) qed subsubsection \Completeness is a Graph Invariant\ lemma (in loopfree_digraph) loopfree_digraphI_app_iso: assumes "digraph_isomorphism hom" shows "loopfree_digraph (app_iso hom G)" proof - interpret iG: wf_digraph "app_iso hom G" using assms by (rule wf_digraphI_app_iso) show ?thesis using assms digraph_isomorphism_inj_on_verts[OF assms] by unfold_locales (auto simp: iso_verts_tail iso_verts_head dest: inj_onD no_loops) qed lemma (in nomulti_digraph) nomulti_digraphI_app_iso: assumes "digraph_isomorphism hom" shows "nomulti_digraph (app_iso hom G)" proof - interpret iG: wf_digraph "app_iso hom G" using assms by (rule wf_digraphI_app_iso) show ?thesis using assms by unfold_locales (auto simp: iso_verts_tail iso_verts_head arc_to_ends_def no_multi_arcs dest: inj_onD) qed lemma (in pre_digraph) symmetricI_app_iso: assumes "digraph_isomorphism hom" assumes "symmetric G" shows "symmetric (app_iso hom G)" proof - let ?G = "app_iso hom G" have "sym (arcs_ends ?G)" proof (rule symI) fix u v assume "u \\<^bsub>?G\<^esub> v" then obtain a where a: "a \ arcs ?G" "tail ?G a = u" "head ?G a = v" by auto then obtain a0 where a0: "a0 \ arcs G" "a = iso_arcs hom a0" by auto with \symmetric G\ obtain b0 where "b0 \ arcs G" "tail G b0 = head G a0" "head G b0 = tail G a0" by (auto simp: symmetric_def arcs_ends_conv elim: symE) moreover define b where "b = iso_arcs hom b0" ultimately have "b \ iso_arcs hom ` arcs G" "tail ?G b = v" "head ?G b = u" using a a0 assms by (auto simp: iso_verts_tail iso_verts_head) then show "v \\<^bsub>?G\<^esub> u" by (auto simp: arcs_ends_conv) qed then show ?thesis by (simp add: symmetric_def) qed lemma (in sym_digraph) sym_digraphI_app_iso: assumes "digraph_isomorphism hom" shows "sym_digraph (app_iso hom G)" proof - interpret iG: wf_digraph "app_iso hom G" using assms by (rule wf_digraphI_app_iso) show ?thesis using assms by unfold_locales (intro symmetricI_app_iso sym_arcs) qed lemma (in graph) graphI_app_iso: assumes "digraph_isomorphism hom" shows "graph (app_iso hom G)" proof - interpret iG: fin_digraph "app_iso hom G" using assms by (rule fin_digraphI_app_iso) interpret iG: loopfree_digraph "app_iso hom G" using assms by (rule loopfree_digraphI_app_iso) interpret iG: nomulti_digraph "app_iso hom G" using assms by (rule nomulti_digraphI_app_iso) interpret iG: sym_digraph "app_iso hom G" using assms by (rule sym_digraphI_app_iso) show ?thesis by intro_locales qed lemma (in wf_digraph) graph_app_iso_eq: assumes "digraph_isomorphism hom" shows "graph (app_iso hom G) \ graph G" using assms by (metis app_iso_inv digraph_isomorphism_invI graph.graphI_app_iso) lemma (in pre_digraph) arcs_ends_iso: assumes "digraph_isomorphism hom" shows "arcs_ends (app_iso hom G) = (\(u,v). (iso_verts hom u, iso_verts hom v)) ` arcs_ends G" using assms by (auto simp: arcs_ends_conv image_image iso_verts_tail iso_verts_head cong: image_cong) lemma inj_onI_pair: assumes "inj_on f S" "T \ S \ S" shows "inj_on (\(u,v). (f u, f v)) T" using assms by (intro inj_onI) (auto dest: inj_onD) lemma (in wf_digraph) complete_digraph_iso: assumes "digraph_isomorphism hom" shows "K\<^bsub>n\<^esub> (app_iso hom G) \ K\<^bsub>n\<^esub> G" (is "?L \ ?R") proof assume "?L" then interpret iG: graph "app_iso hom G" by (simp add: complete_digraph_def) { have "{(u, v). u \ iso_verts hom ` verts G \ v \ iso_verts hom ` verts G \ u \ v} = (\(u,v). (iso_verts hom u, iso_verts hom v)) ` {(u,v). u \ verts G \ v \ verts G \ iso_verts hom u \ iso_verts hom v}" (is "?L = _") by auto also have "\ = (\(u,v). (iso_verts hom u, iso_verts hom v)) ` {(u,v). u \ verts G \ v \ verts G \ u \ v}" using digraph_isomorphism_inj_on_verts[OF assms] by (auto dest: inj_onD) finally have "?L = \" . } note X = this { fix A assume A: "A \ verts G \ verts G" then have "inj_on (\(u, v). (iso_verts hom u, iso_verts hom v)) A" using A digraph_isomorphism_inj_on_verts[OF assms] by (intro inj_onI_pair) } note Y = this have "(arcs_ends G \ {(u, v). u \ verts G \ v \ verts G \ u \ v}) \ verts G \ verts G" by auto note Y' = Y[OF this] show ?R using assms \?L\ by (simp add: complete_digraph_def X arcs_ends_iso graph_app_iso_eq inj_on_Un_image_eq_iff Y') next assume "?R" then show ?L using assms by (fastforce simp add: complete_digraph_def arcs_ends_iso graph_app_iso_eq) qed subsubsection \Conclusion\ definition (in pre_digraph) mk_iso :: "('a \ 'c) \ ('b \ 'd) \ ('a, 'b, 'c, 'd) digraph_isomorphism" where "mk_iso fv fa \ \ iso_verts = fv, iso_arcs = fa, iso_head = fv o head G o the_inv_into (arcs G) fa, iso_tail = fv o tail G o the_inv_into (arcs G) fa \" lemma (in pre_digraph) mk_iso_simps[simp]: "iso_verts (mk_iso fv fa) = fv" "iso_arcs (mk_iso fv fa) = fa" by (auto simp: mk_iso_def) lemma (in wf_digraph) digraph_isomorphism_mk_iso: assumes "inj_on fv (verts G)" "inj_on fa (arcs G)" shows "digraph_isomorphism (mk_iso fv fa)" using assms by (auto simp: digraph_isomorphism_def mk_iso_def the_inv_into_f_f wf_digraph) definition "pairself f \ \x. case x of (u,v) \ (f u, f v)" lemma inj_on_pairself: assumes "inj_on f S" and "T \ S \ S" shows "inj_on (pairself f) T" using assms unfolding pairself_def by (rule inj_onI_pair) definition mk_iso_nomulti :: "('a,'b) pre_digraph \ ('c,'d) pre_digraph \ ('a \ 'c) \ ('a, 'b, 'c, 'd) digraph_isomorphism" where "mk_iso_nomulti G H fv \ \ iso_verts = fv, iso_arcs = the_inv_into (arcs H) (arc_to_ends H) o pairself fv o arc_to_ends G, iso_head = head H, iso_tail = tail H \" lemma (in pre_digraph) mk_iso_simps_nomulti[simp]: "iso_verts (mk_iso_nomulti G H fv) = fv" "iso_head (mk_iso_nomulti G H fv) = head H" "iso_tail (mk_iso_nomulti G H fv) = tail H" by (auto simp: mk_iso_nomulti_def) lemma (in nomulti_digraph) assumes "nomulti_digraph H" assumes fv: "inj_on fv (verts G)" "verts H = fv ` verts G" and arcs_ends: "arcs_ends H = pairself fv ` arcs_ends G" shows digraph_isomorphism_mk_iso_nomulti: "digraph_isomorphism (mk_iso_nomulti G H fv)" (is "?t_multi") and ap_iso_mk_iso_nomulti_eq: "app_iso (mk_iso_nomulti G H fv) G = H" (is "?t_app") and digraph_iso_mk_iso_nomulti: "digraph_iso G H" (is "?t_iso") using assms proof - interpret H: nomulti_digraph H by fact let ?fa = "iso_arcs (mk_iso_nomulti G H fv)" have fa: "bij_betw ?fa (arcs G) (arcs H)" proof - have "bij_betw (arc_to_ends G) (arcs G) (arcs_ends G)" by (auto simp: bij_betw_def inj_on_arc_to_ends arcs_ends_def) also have "bij_betw (pairself fv) (arcs_ends G) (arcs_ends H)" using arcs_ends by (auto simp: bij_betw_def arcs_ends_def arc_to_ends_def intro: fv inj_on_pairself) also (bij_betw_trans) have "bij_betw (the_inv_into (arcs H) (arc_to_ends H)) (arcs_ends H) (arcs H)" by (auto simp: bij_betw_def the_inv_into_into H.inj_on_arc_to_ends arcs_ends_def inj_on_the_inv_into) finally (bij_betw_trans) show ?thesis by (simp add: mk_iso_nomulti_def o_assoc) qed moreover { fix a assume "a \ arcs G" then have "pairself fv (arc_to_ends G a) \ arcs_ends H" using arcs_ends by (auto simp: arcs_ends_def) then obtain b where "(pairself fv (arc_to_ends G a)) = arc_to_ends H b" "b \ arcs H" by (auto simp: arcs_ends_def) then have "fv (tail G a) = tail H (?fa a)" "fv (head G a) = head H (?fa a)" by (auto simp: mk_iso_nomulti_def the_inv_into_f_f H.inj_on_arc_to_ends) (auto simp: pairself_def arc_to_ends_def) } ultimately show ?t_multi ?t_app using fv by (auto simp: digraph_isomorphism_def bij_betw_def wf_digraph) then show ?t_iso by (auto simp: digraph_iso_def) qed lemma complete_digraph_are_iso: assumes "K\<^bsub>n\<^esub> G" "K\<^bsub>n\<^esub> H" shows "digraph_iso G H" proof - interpret G: graph G using assms by (simp add: complete_digraph_def) interpret H: graph H using assms by (simp add: complete_digraph_def) from assms have "card (verts G) = n" "card (verts H) = n" by (auto simp: complete_digraph_def) with G.finite_verts H.finite_verts obtain fv where "bij_betw fv (verts G) (verts H)" by (metis finite_same_card_bij) then have fv: "inj_on fv (verts G)" "verts H = fv ` verts G" by (auto simp: bij_betw_def) have "arcs_ends H = {(u,v). u \ verts H \ v \ verts H \ u \ v}" using \K\<^bsub>n\<^esub> H\ by (auto simp: complete_digraph_def) also have "\ = pairself fv ` {(u,v). u \ verts G \ v \ verts G \ u \ v}" (is "?L = ?R") proof (intro set_eqI iffI) fix x assume "x \ ?L" then have "fst x \ fv ` verts G" "snd x \ fv ` verts G" "fst x \ snd x" using fv by auto then obtain u v where "fst x = fv u" "snd x = fv v" "u \ verts G" "v \ verts G" by auto then have "(fst x, snd x) \ ?R" using \x \ ?L\ by (auto simp: pairself_def) then show "x \ ?R" by auto next fix x assume "x \ ?R" then show "x \ ?L" using fv by (auto simp: pairself_def dest: inj_onD) qed also have "\ = pairself fv ` arcs_ends G" using \K\<^bsub>n\<^esub> G\ by (auto simp: complete_digraph_def) finally have arcs_ends: "arcs_ends H = pairself fv ` arcs_ends G" . show ?thesis using H.nomulti_digraph fv arcs_ends by (rule G.digraph_iso_mk_iso_nomulti) qed lemma pairself_image_prod: "pairself f ` (A \ B) = f ` A \ f ` B" by (auto simp: pairself_def) lemma complete_bipartite_digraph_are_iso: assumes "K\<^bsub>m,n\<^esub> G" "K\<^bsub>m,n\<^esub> H" shows "digraph_iso G H" proof - interpret G: graph G using assms by (simp add: complete_bipartite_digraph_def) interpret H: graph H using assms by (simp add: complete_bipartite_digraph_def) from assms obtain GU GV where G_parts: "verts G = GU \ GV" "GU \ GV = {}" "card GU = m" "card GV = n" "arcs_ends G = GU \ GV \ GV \ GU" by (auto simp: complete_bipartite_digraph_def) from assms obtain HU HV where H_parts: "verts H = HU \ HV" "HU \ HV = {}" "card HU = m" "card HV = n" "arcs_ends H = HU \ HV \ HV \ HU" by (auto simp: complete_bipartite_digraph_def) have fin: "finite GU" "finite GV" "finite HU" "finite HV" using G_parts H_parts G.finite_verts H.finite_verts by simp_all obtain fv_U where fv_U: "bij_betw fv_U GU HU" using \card GU = _\ \card HU = _\ \finite GU\ \finite HU\ by (metis finite_same_card_bij) obtain fv_V where fv_V: "bij_betw fv_V GV HV" using \card GV = _\ \card HV = _\ \finite GV\ \finite HV\ by (metis finite_same_card_bij) define fv where "fv x = (if x \ GU then fv_U x else fv_V x)" for x have "\x. x \ GV \ x \ GU" using \GU \ GV = {}\ by blast then have bij_fv_UV: "bij_betw fv GU HU" "bij_betw fv GV HV" using fv_U fv_V by (auto simp: fv_def cong: bij_betw_cong) then have "bij_betw fv (verts G) (verts H)" unfolding \verts G = _\ \verts H = _\ using \HU \ _ = {}\ by (rule bij_betw_combine) then have fv: "inj_on fv (verts G)" "verts H = fv ` verts G" by (auto simp: bij_betw_def) have "arcs_ends H = HU \ HV \ HV \ HU" using \K\<^bsub>m,n\<^esub> H\ H_parts by (auto simp: complete_digraph_def) also have "\ = pairself fv ` (GU \ GV \ GV \ GU)" (is "?L = ?R") proof (intro set_eqI iffI) fix x assume "x \ ?L" then have "(fst x \ fv ` GU \ snd x \ fv ` GV) \ (fst x \ fv ` GV \ snd x \ fv ` GU)" using bij_fv_UV by (auto simp: bij_betw_def) then show "x \ ?R" by (cases x) (auto simp: pairself_image_prod image_Un) next fix x assume "x \ ?R" then show "x \ ?L" using bij_fv_UV by (auto simp: pairself_image_prod image_Un bij_betw_def) qed also have "\ = pairself fv ` arcs_ends G" using \K\<^bsub>m,n\<^esub> G\ G_parts by (auto simp: complete_bipartite_digraph_def) finally have arcs_ends: "arcs_ends H = pairself fv ` arcs_ends G" . show ?thesis using H.nomulti_digraph fv arcs_ends by (rule G.digraph_iso_mk_iso_nomulti) qed lemma K5_not_comb_planar: assumes "K\<^bsub>5\<^esub> G" shows "\comb_planar G" proof - interpret graph G using assms by (auto simp: complete_digraph_def) have "digraph_iso G c_K5" using assms complete_c_K5 by (rule complete_digraph_are_iso) then obtain hom where hom: "digraph_isomorphism hom" "app_iso hom G = c_K5" by (auto simp: digraph_iso_def) then show ?thesis using c_K5_not_comb_planar comb_planar_iso by fastforce qed lemma K33_not_comb_planar: assumes "K\<^bsub>3,3\<^esub> G" shows "\comb_planar G" proof - interpret graph G using assms by (auto simp: complete_bipartite_digraph_def) have "digraph_iso G c_K33" using assms complete_c_K33 by (rule complete_bipartite_digraph_are_iso) then obtain hom where hom: "digraph_isomorphism hom" "app_iso hom G = c_K33" by (auto simp: digraph_iso_def) then show ?thesis using c_K33_not_comb_planar comb_planar_iso by fastforce qed end diff --git a/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy b/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy --- a/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy +++ b/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy @@ -1,1857 +1,1858 @@ theory Planar_Subgraph imports Graph_Genus Permutations_2 "HOL-Library.FuncSet" "HOL-Library.Simps_Case_Conv" begin section \Combinatorial Planarity and Subgraphs\ lemma out_arcs_emptyD_dominates: assumes "out_arcs G x = {}" shows "\x \\<^bsub>G\<^esub> y" using assms by (auto simp: out_arcs_def) lemma (in wf_digraph) reachable_refl_iff: "u \\<^sup>* u \ u \ verts G" by (auto simp: reachable_in_verts) context digraph_map begin lemma face_cycle_set_succ[simp]: "face_cycle_set (face_cycle_succ a) = face_cycle_set a" by (metis face_cycle_eq face_cycle_set_self face_cycle_succ_inD) lemma face_cycle_succ_funpow_in[simp]: "(face_cycle_succ ^^ n) a \ arcs G \ a \ arcs G" by (induct n) auto lemma segment_face_cycle_x_x_eq: "segment face_cycle_succ x x = face_cycle_set x - {x}" unfolding face_cycle_set_def using face_cycle_succ_permutes finite_arcs permutation_permutes by (intro segment_x_x_eq) blast lemma fcs_x_eq_x: "face_cycle_succ x = x \ face_cycle_set x = {x}" (is "?L \ ?R") unfolding face_cycle_set_def orbit_eq_singleton_iff .. end lemma (in bidirected_digraph) bidirected_digraph_del_arc: "bidirected_digraph (pre_digraph.del_arc (pre_digraph.del_arc G (arev a)) a) (perm_restrict arev (arcs G - {a , arev a}))" proof unfold_locales fix b assume A: "b \ arcs (pre_digraph.del_arc (del_arc (arev a)) a)" have "arev b \ b \ b \ arev a \ b \ a \ perm_restrict arev (arcs G - {a, arev a}) (arev b) = b" using bij_arev arev_dom by (subst perm_restrict_simps) (auto simp: bij_iff) then show "perm_restrict arev (arcs G - {a, arev a}) (perm_restrict arev (arcs G - {a, arev a}) b) = b" using A by (auto simp: pre_digraph.del_arc_simps perm_restrict_simps arev_dom) qed (auto simp: pre_digraph.del_arc_simps perm_restrict_simps arev_dom) lemma (in bidirected_digraph) bidirected_digraph_del_vert: "bidirected_digraph (del_vert u) (perm_restrict arev (arcs (del_vert u)))" by unfold_locales (auto simp: del_vert_simps perm_restrict_simps arev_dom) lemma (in pre_digraph) ends_del_arc: "arc_to_ends (del_arc u) = arc_to_ends G" by (simp add: arc_to_ends_def fun_eq_iff) lemma (in pre_digraph) dominates_arcsD: assumes "v \\<^bsub>del_arc u\<^esub> w" shows "v \\<^bsub>G\<^esub> w" using assms by (auto simp: arcs_ends_def ends_del_arc) lemma (in wf_digraph) reachable_del_arcD: assumes "v \\<^sup>*\<^bsub>del_arc u\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" proof - interpret H: wf_digraph "del_arc u" by (rule wf_digraph_del_arc) from assms show ?thesis by (induct) (auto dest: dominates_arcsD intro: adj_reachable_trans) qed lemma (in fin_digraph) finite_isolated_verts[intro!]: "finite isolated_verts" by (auto simp: isolated_verts_def) lemma (in wf_digraph) isolated_verts_in_sccs: assumes "u \ isolated_verts" shows "{u} \ sccs_verts" proof - have "v = u" if "u \\<^sup>*\<^bsub>G\<^esub> v" for v using that assms by induct (auto simp: arcs_ends_def arc_to_ends_def isolated_verts_def) with assms show ?thesis by (auto simp: sccs_verts_def isolated_verts_def) qed lemma (in digraph_map) in_face_cycle_sets: "a \ arcs G \ face_cycle_set a \ face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma (in digraph_map) heads_face_cycle_set: assumes "a \ arcs G" shows "head G ` face_cycle_set a = tail G ` face_cycle_set a" (is "?L = ?R") proof (intro set_eqI iffI) fix u assume "u \ ?L" then obtain b where "b \ face_cycle_set a" "head G b = u" by blast then have "face_cycle_succ b \ face_cycle_set a" "tail G (face_cycle_succ b) = u" using assms by (auto simp: tail_face_cycle_succ face_cycle_succ_inI in_face_cycle_setD) then show "u \ ?R" by auto next fix u assume "u \ ?R" then obtain b where "b \ face_cycle_set a" "tail G b = u" by blast moreover then obtain c where "b = face_cycle_succ c" by (metis face_cycle_succ_pred) ultimately have "c \ face_cycle_set a" "head G c = u" by (auto dest: face_cycle_succ_inD) (metis assms face_cycle_succ_no_arc in_face_cycle_setD tail_face_cycle_succ) then show "u \ ?L" by auto qed lemma (in pre_digraph) casI_nth: assumes "p \ []" "u = tail G (hd p)" "v = head G (last p)" "\i. Suc i < length p \ head G (p ! i) = tail G (p ! Suc i)" shows "cas u p v" using assms proof (induct p arbitrary: u) case Nil then show ?case by simp next case (Cons a p) have "cas (head G a) p v" proof (cases "p = []") case False then show ?thesis using Cons.prems(1-3) Cons.prems(4)[of 0] Cons.prems(4)[of "Suc i" for i] by (intro Cons) (simp_all add: hd_conv_nth) qed (simp add: Cons) with Cons show ?case by simp qed lemma (in digraph_map) obtain_trail_in_fcs: assumes "a \ arcs G" "a0 \ face_cycle_set a" "an \ face_cycle_set a" obtains p where "trail (tail G a0) p (head G an)" "p \ []" "hd p = a0" "last p = an" "set p \ face_cycle_set a" proof - have fcs_a: "face_cycle_set a = orbit face_cycle_succ a0" using assms face_cycle_eq by (simp add: face_cycle_set_def) have "a0 = (face_cycle_succ ^^ 0) a0" by simp have "an = (face_cycle_succ ^^ funpow_dist face_cycle_succ a0 an) a0" using assms by (simp add: fcs_a funpow_dist_prop) define p where "p = map (\n. (face_cycle_succ ^^ n) a0) [0..i. i < length p \ p ! i = (face_cycle_succ ^^ i) a0" by (auto simp: p_def simp del: upt_Suc) have P2: "p \ []" by (simp add: p_def) have P3: "hd p = a0" using \a0 = _\ by (auto simp: p_def hd_map simp del: upt_Suc) have P4: "last p = an" using \an = _\ by (simp add: p_def) have P5: "set p \ face_cycle_set a" unfolding p_def fcs_a orbit_altdef_permutation[OF permutation_face_cycle_succ] by auto have P1: "trail (tail G a0) p (head G an)" proof - have "distinct p" proof - have "an \ orbit face_cycle_succ a0" using assms by (simp add: fcs_a) then have "inj_on (\n. (face_cycle_succ ^^ n) a0) {0..funpow_dist face_cycle_succ a0 an}" by (rule inj_on_funpow_dist) also have "{0..funpow_dist face_cycle_succ a0 an} = (set [0..n. (face_cycle_succ ^^ n) a0) (set [0.. arcs G" by (metis assms(1-2) in_face_cycle_setD) then have "tail G a0 \ verts G" by simp moreover have "set p \ arcs G" using P5 by (metis assms(1) in_face_cycle_setD subset_code(1)) moreover then have "\i. Suc i < length p \ p ! Suc i \ arcs G" by auto then have "\i. Suc i < length p \ head G (p ! i) = tail G (p ! Suc i)" by (auto simp: p_nth tail_face_cycle_succ) ultimately show ?thesis using P2 P3 P4 unfolding trail_def awalk_def by (auto intro: casI_nth) qed from P1 P2 P3 P4 P5 show ?thesis .. qed lemma (in digraph_map) obtain_trail_in_fcs': assumes "a \ arcs G" "u \ tail G ` face_cycle_set a" "v \ tail G ` face_cycle_set a" obtains p where "trail u p v" "set p \ face_cycle_set a" proof - from assms obtain a0 where "tail G a0 = u" "a0 \ face_cycle_set a" by auto moreover from assms obtain an where "head G an = v" "an \ face_cycle_set a" by (auto simp: heads_face_cycle_set[symmetric]) ultimately obtain p where "trail u p v" "set p \ face_cycle_set a" using \a \ arcs G\ by (metis obtain_trail_in_fcs) then show ?thesis .. qed subsection \Deleting an isolated vertex\ locale del_vert_props = digraph_map + fixes u assumes u_in: "u \ verts G" assumes u_isolated: "out_arcs G u = {}" begin lemma u_isolated_in: "in_arcs G u = {}" using u_isolated by (simp add: in_arcs_eq) lemma arcs_dv: "arcs (del_vert u) = arcs G" using u_isolated u_isolated_in by (auto simp: del_vert_simps) lemma out_arcs_dv: "out_arcs (del_vert u) = out_arcs G" by (auto simp: fun_eq_iff arcs_dv tail_del_vert) lemma digraph_map_del_vert: shows "digraph_map (del_vert u) M" proof - have "perm_restrict (edge_rev M) (arcs (del_vert u)) = edge_rev M" using has_dom_arev arcs_dv by (auto simp: perm_restrict_dom_subset) then interpret H: bidirected_digraph "del_vert u" "edge_rev M" using bidirected_digraph_del_vert[of u] by simp show ?thesis by unfold_locales (auto simp: arcs_dv edge_succ_permutes out_arcs_dv edge_succ_cyclic verts_del_vert) qed end sublocale del_vert_props \ H: digraph_map "del_vert u" M by (rule digraph_map_del_vert) context del_vert_props begin lemma card_verts_dv: "card (verts G) = Suc (card (verts (del_vert u)))" by (auto simp: verts_del_vert) (rule card.remove, auto simp: u_in) lemma card_arcs_dv: "card (arcs (del_vert u)) = card (arcs G)" using u_isolated by (auto simp add: arcs_dv in_arcs_eq) lemma isolated_verts_dv: "H.isolated_verts = isolated_verts - {u}" by (auto simp: isolated_verts_def H.isolated_verts_def verts_del_vert out_arcs_dv) lemma u_in_isolated_verts: "u \ isolated_verts" using u_in u_isolated by (auto simp: isolated_verts_def) lemma card_isolated_verts_dv: "card isolated_verts = Suc (card H.isolated_verts)" by (simp add: isolated_verts_dv) (rule card.remove, auto simp: u_in_isolated_verts) lemma face_cycles_dv: "H.face_cycle_sets = face_cycle_sets" unfolding H.face_cycle_sets_def face_cycle_sets_def arcs_dv .. lemma euler_char_dv: "euler_char = 1 + H.euler_char" by (auto simp: euler_char_def H.euler_char_def card_arcs_dv card_verts_dv face_cycles_dv) lemma adj_dv: "v \\<^bsub>del_vert u\<^esub> w \ v \\<^bsub>G\<^esub> w" by (auto simp: arcs_ends_def arcs_dv ends_del_vert) lemma reachable_del_vertD: assumes "v \\<^sup>*\<^bsub>del_vert u\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" using assms by induct (auto simp: verts_del_vert adj_dv intro: adj_reachable_trans) lemma reachable_del_vertI: assumes "v \\<^sup>*\<^bsub>G\<^esub> w" "u \ v \ u \ w" shows "v \\<^sup>*\<^bsub>del_vert u\<^esub> w" using assms proof induct case (step x y) from \x \\<^bsub>G\<^esub> y\ obtain a where "a \ arcs G" "head G a = y" by auto then have "a \ in_arcs G y" by auto then have "y \ u" using u_isolated in_arcs_eq[of u] by auto with step show ?case by (auto simp: adj_dv intro: H.adj_reachable_trans) qed (auto simp: verts_del_vert) lemma G_reach_conv: "v \\<^sup>*\<^bsub>G\<^esub> w \ v \\<^sup>*\<^bsub>del_vert u\<^esub> w \ (v = u \ w = u)" by (auto dest: reachable_del_vertI reachable_del_vertD intro: u_in) lemma sccs_verts_dv: "H.sccs_verts = sccs_verts - {{u}}" (is "?L = ?R") proof - have *:"\S x. S \ sccs_verts \ S \ H.sccs_verts \ x \ S \ x = u" by (simp add: H.in_sccs_verts_conv_reachable in_sccs_verts_conv_reachable G_reach_conv) (meson H.reachable_trans) show ?thesis by (auto dest: *) (auto simp: H.in_sccs_verts_conv_reachable in_sccs_verts_conv_reachable G_reach_conv H.reachable_refl_iff verts_del_vert) qed lemma card_sccs_verts_dv: "card sccs_verts = Suc (card H.sccs_verts)" unfolding sccs_verts_dv by (rule card.remove) (auto simp: isolated_verts_in_sccs u_in_isolated_verts finite_sccs_verts) lemma card_sccs_dv: "card sccs = Suc (card H.sccs)" using card_sccs_verts_dv by (simp add: card_sccs_verts H.card_sccs_verts) lemma euler_genus_eq: "H.euler_genus = euler_genus" by (auto simp: pre_digraph_map.euler_genus_def card_sccs_dv card_isolated_verts_dv euler_char_dv) end subsection \Deleting an arc pair\ locale bidel_arc = G: digraph_map + fixes a assumes a_in: "a \ arcs G" begin abbreviation "a' \ edge_rev M a" definition H :: "('a,'b) pre_digraph" where "H \ pre_digraph.del_arc (pre_digraph.del_arc G a') a" definition HM :: "'b pre_map" where "HM = \ edge_rev = perm_restrict (edge_rev M) (arcs G - {a, a'}) , edge_succ = perm_rem a (perm_rem a' (edge_succ M)) \" lemma verts_H: "verts H = verts G" and arcs_H: "arcs H = arcs G - {a, a'}" and tail_H: "tail H = tail G" and head_H: "head H = head G" and ends_H: "arc_to_ends H = arc_to_ends G"and arcs_in: "{a,a'} \ arcs G" and ends_in: "{tail G a, head G a} \ verts G" by (auto simp: H_def pre_digraph.del_arc_simps a_in arc_to_ends_def) lemma cyclic_on_edge_succ: assumes "x \ verts H" "out_arcs H x \ {}" shows "cyclic_on (edge_succ HM) (out_arcs H x)" proof - have oa_H: "out_arcs H x = (out_arcs G x - {a'}) - {a}" by (auto simp: arcs_H tail_H) have "cyclic_on (perm_rem a (perm_rem a' (edge_succ M))) (out_arcs G x - {a'} - {a})" using assms by (intro cyclic_on_perm_rem G.edge_succ_cyclic) (auto simp: oa_H G.bij_edge_succ G.edge_succ_cyclic) then show ?thesis by (simp add: HM_def oa_H) qed lemma digraph_map: "digraph_map H HM" proof - interpret fin_digraph H unfolding H_def by (rule fin_digraph.fin_digraph_del_arc) (rule G.fin_digraph_del_arc) interpret bidirected_digraph H "edge_rev HM" unfolding H_def using G.bidirected_digraph_del_arc[of a] by (auto simp: HM_def) have *: "insert a' (insert a (arcs H)) = arcs G" using a_in by (auto simp: arcs_H) have "edge_succ HM permutes arcs H" unfolding HM_def by (auto simp: * intro!: perm_rem_permutes G.edge_succ_permutes) moreover { fix v assume "v \ verts H" "out_arcs H v \ {}" then have "cyclic_on (edge_succ HM) (out_arcs H v)" by (rule cyclic_on_edge_succ) } ultimately show ?thesis by unfold_locales qed lemma rev_H: "bidel_arc.H G M a' = H" (is ?t1) and rev_HM: "bidel_arc.HM G M a' = HM" (is ?t2) proof - interpret rev: bidel_arc G M a' using a_in by unfold_locales simp show ?t1 by (rule pre_digraph.equality) (auto simp: rev.verts_H verts_H rev.arcs_H arcs_H rev.tail_H tail_H rev.head_H head_H) show ?t2 using G.edge_succ_permutes by (intro pre_map.equality) (auto simp: HM_def rev.HM_def insert_commute perm_rem_commutes permutes_conv_has_dom) qed end sublocale bidel_arc \ H: digraph_map H HM by (rule digraph_map) context bidel_arc begin lemma a_neq_a': "a \ a'" by (metis G.arev_neq a_in) lemma arcs_G: "arcs G = insert a (insert a' (arcs H))" and arcs_not_in: "{a,a'} \ arcs H = {}" using arcs_in by (auto simp: arcs_H) lemma card_arcs_da: "card (arcs G) = 2 + card (arcs H)" using arcs_G arcs_not_in a_neq_a' by (auto simp: card_insert_if) lemma cas_da: "H.cas = G.cas" proof - { fix u p v have "H.cas u p v = G.cas u p v" by (induct p arbitrary: u) (simp_all add: tail_H head_H) } then show ?thesis by (simp add: fun_eq_iff) qed lemma reachable_daD: assumes "v \\<^sup>*\<^bsub>H\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" apply (rule G.reachable_del_arcD) apply (rule wf_digraph.reachable_del_arcD) apply (rule G.wf_digraph_del_arc) using assms unfolding H_def by assumption lemma not_G_isolated_a: "{tail G a, head G a} \ G.isolated_verts = {}" using a_in G.in_arcs_eq[of "head G a"] by (auto simp: G.isolated_verts_def) lemma isolated_other_da: assumes "u \ {tail G a, head G a}" shows "u \ H.isolated_verts \ u \ G.isolated_verts" using assms by (auto simp: pre_digraph.isolated_verts_def verts_H arcs_H tail_H out_arcs_def) lemma isolated_da_pre: "H.isolated_verts = G.isolated_verts \ (if tail G a \ H.isolated_verts then {tail G a} else {}) \ (if head G a \ H.isolated_verts then {head G a} else {})" (is "?L = ?R") proof (intro set_eqI iffI) fix x assume "x \ ?L" then show "x \ ?R" by (cases "x \ {tail G a, head G a}") (auto simp:isolated_other_da) next fix x assume "x \ ?R" then show "x \ ?L" using not_G_isolated_a by (cases "x \ {tail G a, head G a}") (auto simp:isolated_other_da split: if_splits) qed lemma card_isolated_verts_da0: "card H.isolated_verts = card G.isolated_verts + card ({tail G a, head G a} \ H.isolated_verts)" using not_G_isolated_a by (subst isolated_da_pre) (auto simp: card_insert_if G.finite_isolated_verts) lemma segments_neq: assumes "segment G.face_cycle_succ a' a \ {} \ segment G.face_cycle_succ a a' \ {}" shows "segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" proof - have bij_fcs: "bij G.face_cycle_succ" using G.face_cycle_succ_permutes by (auto simp: permutes_conv_has_dom) show ?thesis using segment_disj[OF a_neq_a' bij_fcs] assms by auto qed lemma H_fcs_eq_G_fcs: assumes "b \ arcs G" "{b,G.face_cycle_succ b} \ {a,a'} = {}" shows "H.face_cycle_succ b = G.face_cycle_succ b" proof - have "edge_rev M b \ {a,a'}" using assms by auto (metis G.arev_arev) then show ?thesis using assms unfolding G.face_cycle_succ_def H.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_simps G.bij_edge_succ) qed lemma face_cycle_set_other_da: assumes "{a, a'} \ G.face_cycle_set b = {}" "b \ arcs G" shows "H.face_cycle_set b = G.face_cycle_set b" proof - have "\s. s \ G.face_cycle_set b \ b \ arcs G \ a \ G.face_cycle_set b \ a' \ G.face_cycle_set b \ pre_digraph_map.face_cycle_succ HM s = G.face_cycle_succ s" by (subst H_fcs_eq_G_fcs) (auto simp: G.in_face_cycle_setD G.face_cycle_succ_inI) then show ?thesis using assms unfolding pre_digraph_map.face_cycle_set_def by (intro orbit_cong) (auto simp add: pre_digraph_map.face_cycle_set_def[symmetric]) qed lemma in_face_cycle_set_other: assumes "S \ G.face_cycle_sets" "{a, a'} \ S = {}" shows "S \ H.face_cycle_sets" proof - from assms obtain b where "S = G.face_cycle_set b" "b \ arcs G" by (auto simp: G.face_cycle_sets_def) with assms have "S = H.face_cycle_set b" by (simp add: face_cycle_set_other_da) moreover with assms have "b \ arcs H" using \b \ arcs G\ by (auto simp: arcs_H) ultimately show ?thesis by (auto simp: H.face_cycle_sets_def) qed lemma H_fcs_in_G_fcs: assumes "b \ arcs H - (G.face_cycle_set a \ G.face_cycle_set a')" shows "H.face_cycle_set b \ G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}" proof - have "H.face_cycle_set b = G.face_cycle_set b" using assms by (intro face_cycle_set_other_da) (auto simp: arcs_H G.face_cycle_eq) moreover have "G.face_cycle_set b \ {G.face_cycle_set a, G.face_cycle_set a'}" "b \ arcs G" using G.face_cycle_eq assms by (auto simp: arcs_H) ultimately show ?thesis by (auto simp: G.face_cycle_sets_def) qed lemma face_cycle_sets_da0: "H.face_cycle_sets = G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'} \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'})" (is "?L = ?R") proof (intro set_eqI iffI) fix S assume "S \ ?L" then obtain b where "S = H.face_cycle_set b" "b \ arcs H" by (auto simp: H.face_cycle_sets_def) then show "S \ ?R" using arcs_not_in H_fcs_in_G_fcs by (cases "b \ G.face_cycle_set a \ G.face_cycle_set a'") auto next fix S assume "S \ ?R" show "S \ ?L" proof (cases "S \ G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}") case True then have "S \ {a,a'} = {}" using G.face_cycle_set_parts by (auto simp: G.face_cycle_sets_def) with True show ?thesis by (intro in_face_cycle_set_other) auto next case False then have "S \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'})" using \S \ ?R\ by blast moreover have "(G.face_cycle_set a \ G.face_cycle_set a') - {a,a'} \ arcs H" using a_in by (auto simp: arcs_H dest: G.in_face_cycle_setD) ultimately show ?thesis by (auto simp: H.face_cycle_sets_def) qed qed lemma card_fcs_aa'_le: "card {G.face_cycle_set a, G.face_cycle_set a'} \ card G.face_cycle_sets" using a_in by (intro card_mono) (auto simp: G.face_cycle_sets_def) lemma card_face_cycle_sets_da0: "card H.face_cycle_sets = card G.face_cycle_sets - card {G.face_cycle_set a, G.face_cycle_set a'} + card (H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'}))" proof - have face_cycle_sets_inter: "(G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}) \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a, a'}) = {}" (is "?L \ ?R = _") proof - define L R P where "L = ?L" and "R = ?R" and "P x \ x \ (G.face_cycle_set a \ G.face_cycle_set a') = {}" for x then have "\x. x \ L \ P x" "\x. x \ R \ \P x" using G.face_cycle_set_parts by (auto simp: G.face_cycle_sets_def) then have "L \ R = {}" by blast then show ?thesis unfolding L_def R_def . qed then show ?thesis using arcs_G by (simp add: card_Diff_subset[symmetric] card_Un_disjoint[symmetric] G.in_face_cycle_sets face_cycle_sets_da0) qed end locale bidel_arc_same_face = bidel_arc + assumes same_face: "G.face_cycle_set a' = G.face_cycle_set a" begin lemma a_in_o: "a \ orbit G.face_cycle_succ a'" unfolding G.face_cycle_set_def[symmetric] by (simp add: same_face) lemma segment_a'_a_in: "segment G.face_cycle_succ a' a \ arcs H" (is "?seg \ _") proof - have "?seg \ G.face_cycle_set a'" by (auto simp: G.face_cycle_set_def segmentD_orbit) moreover have "G.face_cycle_set a' \ arcs G" by (auto simp: G.face_cycle_set_altdef a_in) ultimately show ?thesis using a_in_o by (auto simp: arcs_H a_in not_in_segment1 not_in_segment2) qed lemma segment_a'_a_neD: assumes "segment G.face_cycle_succ a' a \ {}" shows "segment G.face_cycle_succ a' a \ H.face_cycle_sets" (is "?seg \ _") proof - let ?b = "G.face_cycle_succ a'" have fcs_a_neq_a': "G.face_cycle_succ a' \ a" by (metis assms segment1_empty) have in_aG: "\x. x \ segment G.face_cycle_succ a' a \ x \ arcs G - {a,a'}" using not_in_segment1 not_in_segment2 segment_a'_a_in by (auto simp: arcs_H) { fix x assume A: "x \ segment G.face_cycle_succ a' a" and B: "G.face_cycle_succ x \ a" from A have "G.face_cycle_succ x \ a'" proof induct case base then show ?case by (metis a_neq_a' G.face_cycle_set_self not_in_segment1 G.face_cycle_set_def same_face segment.intros) next case step then show ?case by (metis a_in_o a_neq_a' not_in_segment1 segment.step) qed with A B have "{x, G.face_cycle_succ x} \ {a, a'} = {}" using not_in_segment1[OF a_in_o] not_in_segment2[of a G.face_cycle_succ a'] by safe with in_aG have "H.face_cycle_succ x = G.face_cycle_succ x" by (intro H_fcs_eq_G_fcs) (auto intro: A) } note fcs_x_eq = this { fix x assume A: "x \ segment G.face_cycle_succ a' a" and B: "G.face_cycle_succ x = a" have "G.face_cycle_succ a \ a" using B in_aG[OF A] G.bij_face_cycle_succ by (auto simp: bij_eq_iff) then have "edge_succ M a \ edge_rev M a" by (metis a_in_o G.arev_arev comp_apply G.face_cycle_succ_def not_in_segment1 segment.base) then have "H.face_cycle_succ x = G.face_cycle_succ a'" using in_aG[OF A] B G.bij_edge_succ unfolding H.face_cycle_succ_def G.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.arev_eq_iff) } note fcs_last_x_eq = this have "segment G.face_cycle_succ a' a = H.face_cycle_set ?b" proof (intro set_eqI iffI) fix x assume "x \ segment G.face_cycle_succ a' a" then show "x \ H.face_cycle_set ?b" proof induct case base then show ?case by auto next case (step x) then show ?case by (subst fcs_x_eq[symmetric]) (auto simp: H.face_cycle_succ_inI) qed next fix x assume A: "x \ H.face_cycle_set ?b" then show "x \ segment G.face_cycle_succ a' a" proof induct case base then show ?case by (intro segment.base fcs_a_neq_a') next case (step x) then show ?case using fcs_a_neq_a' by (cases "G.face_cycle_succ x = a") (auto simp: fcs_last_x_eq fcs_x_eq intro: segment.intros) qed qed then show ?thesis using segment_a'_a_in by (auto simp: H.face_cycle_sets_def) qed lemma segment_a_a'_neD: assumes "segment G.face_cycle_succ a a' \ {}" shows "segment G.face_cycle_succ a a' \ H.face_cycle_sets" proof - interpret rev: bidel_arc_same_face G M a' using a_in same_face by unfold_locales simp_all from assms show ?thesis using rev.segment_a'_a_neD by (simp add: rev_H rev_HM) qed lemma H_fcs_full: assumes "SS \ H.face_cycle_sets" shows "H.face_cycle_set ` (\SS) = SS" proof - { fix x assume "x \ \SS" then obtain S where "S \ SS" "x \ S" "S \ H.face_cycle_sets" using assms by auto then have "H.face_cycle_set x = S" using H.face_cycle_set_parts by (auto simp: H.face_cycle_sets_def) then have "H.face_cycle_set x \ SS" using \S \ SS\ by auto } moreover { fix S assume "S \ SS" then obtain x where "x \ arcs H" "S = H.face_cycle_set x" "x \ S" using assms by (auto simp: H.face_cycle_sets_def) then have "S \ H.face_cycle_set ` \SS" using \S \ SS\ by auto } ultimately show ?thesis by auto qed lemma card_fcs_gt_0: "0 < card G.face_cycle_sets" using a_in by (auto simp: card_gt_0_iff dest: G.in_face_cycle_sets) lemma card_face_cycle_sets_da': "card H.face_cycle_sets = card G.face_cycle_sets - 1 + card ({segment G.face_cycle_succ a a', segment G.face_cycle_succ a' a, {}} - {{}})" proof - have "G.face_cycle_set a = {a,a'} \ segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" using a_neq_a' same_face by (intro cyclic_split_segment) (auto simp: G.face_cycle_succ_cyclic) then have *: "G.face_cycle_set a \ G.face_cycle_set a' - {a, a'} = segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" by (auto simp: same_face G.face_cycle_set_def[symmetric] not_in_segment1 not_in_segment2) have **: "H.face_cycle_set ` (G.face_cycle_set a \ G.face_cycle_set a' - {a, a'}) = (if segment G.face_cycle_succ a a' \ {} then {segment G.face_cycle_succ a a'} else {}) \ (if segment G.face_cycle_succ a' a \ {} then {segment G.face_cycle_succ a' a} else {})" unfolding * using H_fcs_full[of "{segment G.face_cycle_succ a a', segment G.face_cycle_succ a' a}"] using H_fcs_full[of "{segment G.face_cycle_succ a a'}"] using H_fcs_full[of "{segment G.face_cycle_succ a' a}"] by (auto simp add: segment_a_a'_neD segment_a'_a_neD) show ?thesis unfolding card_face_cycle_sets_da0 ** by (simp add: same_face card_insert_if) qed end locale bidel_arc_diff_face = bidel_arc + assumes diff_face: "G.face_cycle_set a' \ G.face_cycle_set a" begin definition S :: "'b set" where "S \ segment G.face_cycle_succ a a \ segment G.face_cycle_succ a' a'" lemma diff_face_not_in: "a \ G.face_cycle_set a'" "a' \ G.face_cycle_set a" using diff_face G.face_cycle_eq by auto lemma H_fcs_eq_for_a: assumes "b \ arcs H \ G.face_cycle_set a" shows "H.face_cycle_set b = S" (is "?L = ?R") proof (intro set_eqI iffI) fix c assume "c \ ?L" then have "c \ arcs H" using assms by (auto dest: H.in_face_cycle_setD) moreover have "c \ G.face_cycle_set a \ G.face_cycle_set a'" proof (rule ccontr) assume A: "\?thesis" then have "G.face_cycle_set c \ (G.face_cycle_set a \ G.face_cycle_set a') = {}" using G.face_cycle_set_parts by (auto simp: arcs_H) also then have "G.face_cycle_set c = H.face_cycle_set c" using \c \ arcs H\ by (subst face_cycle_set_other_da) (auto simp: arcs_H) also have "\ = H.face_cycle_set b" using \c \ ?L\ using H.face_cycle_set_parts by auto finally show False using assms by auto qed ultimately show "c \ ?R" unfolding S_def arcs_H G.segment_face_cycle_x_x_eq by auto next fix x assume "x \ ?R" from assms have "a \ b" by (auto simp: arcs_H) from assms have b_in: "b \ segment G.face_cycle_succ a a" using G.segment_face_cycle_x_x_eq by (auto simp: arcs_H) have fcs_a_neq_a: "G.face_cycle_succ a \ a" using assms \a \ b\ by (auto simp add: G.segment_face_cycle_x_x_eq G.fcs_x_eq_x) have split_seg: "segment G.face_cycle_succ a a = segment G.face_cycle_succ a b \ {b} \ segment G.face_cycle_succ b a" using b_in by (intro segment_split) have a_in_orb_a: "a \ orbit G.face_cycle_succ a" by (simp add: G.face_cycle_set_def[symmetric]) define c where "c = inv G.face_cycle_succ a" have c_succ: "G.face_cycle_succ c = a" unfolding c_def by (meson bij_inv_eq_iff permutation_bijective G.permutation_face_cycle_succ) have c_in_aa: "c \ segment G.face_cycle_succ a a" unfolding G.segment_face_cycle_x_x_eq c_def using fcs_a_neq_a c_succ c_def by force have c_in: "c \ {b} \ segment G.face_cycle_succ b a" using split_seg b_in c_succ c_in_aa by (auto dest: not_in_segment1[OF segmentD_orbit] intro: segment.intros) from c_in_aa have "c \ arcs H" unfolding G.segment_face_cycle_x_x_eq using arcs_in c_succ diff_face by (auto simp: arcs_H G.face_cycle_eq[of a']) have b_in_L: "b \ ?L" by auto moreover { fix x assume "x \ segment G.face_cycle_succ b a" then have "x \ ?L" proof induct case base then show ?case using assms diff_face_not_in(2) by (subst H_fcs_eq_G_fcs[symmetric]) (auto simp: arcs_H intro: H.face_cycle_succ_inI G.face_cycle_succ_inI) next case (step x) have "G.face_cycle_succ x \ G.face_cycle_set a \ b \ G.face_cycle_set a \ False" using step(1) by (metis G.face_cycle_eq G.face_cycle_succ_inI pre_digraph_map.face_cycle_set_def segmentD_orbit) moreover have "x \ arcs G" using step assms H.in_face_cycle_setD arcs_H by auto moreover then have "(G.face_cycle_succ x \ G.face_cycle_set a \ b \ G.face_cycle_set a \ False) \ {x, G.face_cycle_succ x} \ {a, a'} = {}" using step(2,3) assms diff_face_not_in(2) H.in_face_cycle_setD arcs_H by safe auto ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed } note sba_in_L = this moreover { fix x assume A: "x \ segment G.face_cycle_succ a' a'" then have "x \ ?L" proof - from c_in have "c \ ?L" using b_in_L sba_in_L by blast have "G.face_cycle_succ a' \ a'" using A by (auto simp add: G.segment_face_cycle_x_x_eq G.fcs_x_eq_x) then have *: "G.face_cycle_succ a' = H.face_cycle_succ c" using a_neq_a' c_succ \c \ arcs H\ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.bij_edge_succ G.arev_eq_iff) from A have "x \ H.face_cycle_set c" proof induct case base then show ?case by (simp add: * H.face_cycle_succ_inI) next case (step x) have "x \ arcs G" using\c \ arcs H\ step.hyps(2) by (auto simp: arcs_H dest: H.in_face_cycle_setD) moreover have "G.face_cycle_succ x \ a' \ {x, G.face_cycle_succ x} \ {a, a'} = {}" using step(1) diff_face_not_in(1) G.face_cycle_succ_inI G.segment_face_cycle_x_x_eq by (auto simp: not_in_segment2) ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed also have "H.face_cycle_set c = ?L" using \c \ ?L\ H.face_cycle_set_parts by auto finally show ?thesis . qed } note sa'a'_in_L = this moreover { assume A: "x \ segment G.face_cycle_succ a b" obtain d where "d \ ?L" and d_succ: "H.face_cycle_succ d = G.face_cycle_succ a" proof (cases "G.face_cycle_succ a' = a'") case True from c_in have "c \ ?L" using b_in_L sba_in_L by blast moreover have "H.face_cycle_succ c = G.face_cycle_succ a" using fcs_a_neq_a c_succ a_neq_a' True \c \ arcs H\ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps arcs_H perm_rem_conv G.bij_edge_succ G.arev_eq_iff) ultimately show ?thesis .. next case False define d where "d = inv G.face_cycle_succ a'" have d_succ: "G.face_cycle_succ d = a'" unfolding d_def by (meson bij_inv_eq_iff permutation_bijective G.permutation_face_cycle_succ) have *: "d \ ?L" using sa'a'_in_L False by (metis DiffI d_succ empty_iff G.face_cycle_set_self G.face_cycle_set_succ insert_iff G.permutation_face_cycle_succ pre_digraph_map.face_cycle_set_def segment_x_x_eq) then have "d \ arcs H" using assms by (auto dest: H.in_face_cycle_setD) have "H.face_cycle_succ d = G.face_cycle_succ a" using fcs_a_neq_a a_neq_a' \d \ arcs H\ d_succ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps arcs_H perm_rem_conv G.bij_edge_succ G.arev_eq_iff) with * show ?thesis .. qed then have "d \ arcs H" using assms by - (drule H.in_face_cycle_setD, auto) from A have "x \ H.face_cycle_set d" proof induct case base then show ?case by (simp add: d_succ[symmetric] H.face_cycle_succ_inI) next case (step x) moreover have "x \ arcs G" using \d \ arcs H\ arcs_H digraph_map.in_face_cycle_setD step.hyps(2) by fastforce moreover have " {x, G.face_cycle_succ x} \ {a, a'} = {}" proof - have "a \ x" using step(2) H.in_face_cycle_setD \d \ arcs H\ arcs_not_in by blast moreover have "a \ G.face_cycle_succ x" by (metis b_in not_in_segment1 segment.step segmentD_orbit step(1)) moreover have "a' \ x" "a' \ G.face_cycle_succ x" using step(1) diff_face_not_in(2) by (auto simp: G.face_cycle_set_def dest!: segmentD_orbit intro: orbit.step) ultimately show ?thesis by auto qed ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed also have "H.face_cycle_set d = ?L" using \d \ ?L\ H.face_cycle_set_parts by auto finally have "x \ ?L" . } ultimately show "x \ ?L" using \x \ ?R\ unfolding S_def split_seg by blast qed lemma HJ_fcs_eq_for_a': assumes "b \ arcs H \ G.face_cycle_set a'" shows "H.face_cycle_set b = S" proof - interpret rev: bidel_arc_diff_face G M a' using arcs_in diff_face by unfold_locales simp_all show ?thesis using rev.H_fcs_eq_for_a assms by (auto simp: rev_H rev_HM S_def rev.S_def) qed lemma card_face_cycle_sets_da': "card H.face_cycle_sets = card G.face_cycle_sets - card {G.face_cycle_set a, G.face_cycle_set a'} + (if S = {} then 0 else 1)" proof - have "S = G.face_cycle_set a \ G.face_cycle_set a' - {a, a'}" unfolding S_def using diff_face_not_in by (auto simp: segment_x_x_eq G.permutation_face_cycle_succ G.face_cycle_set_def) moreover { fix x assume "x \ S" then have "x \ arcs H \ (G.face_cycle_set a \ G.face_cycle_set a' - {a, a'})" unfolding \S = _\ arcs_H using a_in by (auto intro: G.in_face_cycle_setD) then have "H.face_cycle_set x = S" using H_fcs_eq_for_a HJ_fcs_eq_for_a' by blast } then have "H.face_cycle_set ` S = (if S = {} then {} else {S})" by auto ultimately show ?thesis by (simp add: card_face_cycle_sets_da0) qed end locale bidel_arc_biconnected = bidel_arc + assumes reach_a: "tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" begin lemma reach_a': "tail G a' \\<^sup>*\<^bsub>H\<^esub> head G a'" using reach_a a_in by (simp add: symmetric_reachable H.sym_arcs) lemma tail_a': "tail G a' = head G a" and head_a': "head G a' = tail G a" using a_in by simp_all lemma reachable_daI: assumes "v \\<^sup>*\<^bsub>G\<^esub> w" shows "v \\<^sup>*\<^bsub>H\<^esub> w" proof - have *: "\v w. v \\<^bsub>G\<^esub> w \ v \\<^sup>*\<^bsub>H\<^esub> w" using reach_a reach_a' by (auto simp: arcs_ends_def ends_H arcs_G arc_to_ends_def tail_a') show ?thesis using assms by induct (auto simp: verts_H intro: * H.reachable_trans) qed lemma reachable_da: "v \\<^sup>*\<^bsub>H\<^esub> w \ v \\<^sup>*\<^bsub>G\<^esub> w" by (metis reachable_daD reachable_daI) lemma sccs_verts_da: "H.sccs_verts = G.sccs_verts" by (auto simp: G.in_sccs_verts_conv_reachable H.in_sccs_verts_conv_reachable reachable_da) lemma card_sccs_da: "card H.sccs = card G.sccs" by (simp add: G.card_sccs_verts[symmetric] H.card_sccs_verts[symmetric] sccs_verts_da) end locale bidel_arc_not_biconnected = bidel_arc + assumes not_reach_a: "\tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" begin lemma H_awalkI: "G.awalk u p v \ {a,a'} \ set p = {} \ H.awalk u p v" by (auto simp: pre_digraph.apath_def pre_digraph.awalk_def verts_H arcs_H cas_da) lemma tail_neq_head: "tail G a \ head G a" using not_reach_a a_in by (metis H.reachable_refl G.head_in_verts verts_H) lemma scc_of_tail_neq_head: "H.scc_of (tail G a) \ H.scc_of (head G a)" proof - have "tail G a \ H.scc_of (tail G a)" "head G a \ H.scc_of (head G a)" using ends_in by (auto simp: H.in_scc_of_self verts_H) with not_reach_a show ?thesis by (auto simp: H.scc_of_def) qed lemma scc_of_G_tail: assumes "u \ G.scc_of (tail G a)" shows "H.scc_of u = H.scc_of (tail G a) \ H.scc_of u = H.scc_of (head G a)" proof - from assms have "u \\<^sup>*\<^bsub>G\<^esub> tail G a" by (auto simp: G.scc_of_def) then obtain p where p: "G.apath u p (tail G a)" by (auto simp: G.reachable_apath) show ?thesis proof (cases "head G a \ set (G.awalk_verts u p)") case True with p obtain p' q where "p = p' @ q" "G.awalk (head G a) q (tail G a)" and p': "G.awalk u p' (head G a)" unfolding G.apath_def by (metis G.awalk_decomp) moreover then have "tail G a \ set (tl (G.awalk_verts (head G a) q))" using tail_neq_head apply (cases q) apply (simp add: G.awalk_Nil_iff ) apply (simp add: G.awalk_Cons_iff) by (metis G.awalkE G.awalk_verts_non_Nil last_in_set) ultimately have "tail G a \ set (G.awalk_verts u p')" using G.apath_decomp_disjoint[OF p, of p' q "tail G a"] by auto with p' have "{a,a'} \ set p' = {}" by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.head_arev) with p' show ?thesis unfolding G.apath_def by (metis H.scc_ofI_awalk H.scc_of_eq H_awalkI) next case False with p have "{a,a'} \ set p = {}" by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.tail_arev) with p show ?thesis unfolding G.apath_def by (metis H.scc_ofI_awalk H.scc_of_eq H_awalkI) qed qed lemma scc_of_other: assumes "u \ G.scc_of (tail G a)" shows "H.scc_of u = G.scc_of u" using assms proof (intro set_eqI iffI) fix v assume "v \ H.scc_of u" then show "v \ G.scc_of u" by (auto simp: H.scc_of_def G.scc_of_def intro: reachable_daD) next fix v assume "v \ G.scc_of u" then obtain p where p: "G.awalk u p v" by (auto simp: G.scc_of_def G.reachable_awalk) moreover have "{a,a'} \ set p = {}" proof - have "\u \\<^sup>*\<^bsub>G\<^esub> tail G a" using assms by (metis G.scc_ofI_reachable) then have "\p. \G.awalk u p (tail G a)" by (metis G.reachable_awalk) then have "tail G a \ set (G.awalk_verts u p)" using p by (auto dest: G.awalk_decomp) with p show ?thesis by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.head_arev) qed ultimately have "H.awalk u p v" by (rule H_awalkI) then show "v \ H.scc_of u" by (metis H.scc_ofI_reachable' H.reachable_awalk) qed lemma scc_of_tail_inter: "tail G a \ G.scc_of (tail G a) \ H.scc_of (tail G a)" using ends_in by (auto simp: G.in_scc_of_self H.in_scc_of_self verts_H) lemma scc_of_head_inter: "head G a \ G.scc_of (tail G a) \ H.scc_of (head G a)" proof - have "tail G a \\<^bsub>G\<^esub> head G a" "head G a \\<^bsub>G\<^esub> tail G a" by (metis a_in G.in_arcs_imp_in_arcs_ends) (metis a_in G.graph_symmetric G.in_arcs_imp_in_arcs_ends) then have "tail G a \\<^sup>*\<^bsub>G\<^esub> head G a" "head G a \\<^sup>*\<^bsub>G\<^esub> tail G a" by auto then show ?thesis using ends_in by (auto simp: G.scc_of_def H.in_scc_of_self verts_H) qed lemma G_scc_of_tail_not_in: "G.scc_of (tail G a) \ H.sccs_verts" proof assume A: "G.scc_of (tail G a) \ H.sccs_verts" from A scc_of_tail_inter have "G.scc_of (tail G a) = H.scc_of (tail G a)" by (metis H.scc_of_in_sccs_verts H.sccs_verts_disjoint a_in empty_iff G.tail_in_verts verts_H) moreover from A scc_of_head_inter have "G.scc_of (tail G a) = H.scc_of (head G a)" by (metis H.scc_of_in_sccs_verts H.sccs_verts_disjoint a_in empty_iff G.head_in_verts verts_H) ultimately show False using scc_of_tail_neq_head by blast qed lemma H_scc_of_a_not_in: "H.scc_of (tail G a) \ G.sccs_verts" "H.scc_of (head G a) \ G.sccs_verts" proof safe assume "H.scc_of (tail G a) \ G.sccs_verts" with scc_of_tail_inter have "G.scc_of (tail G a) = H.scc_of (tail G a)" by (metis G.scc_of_in_sccs_verts G.sccs_verts_disjoint a_in empty_iff G.tail_in_verts) with G_scc_of_tail_not_in show False using ends_in by (auto simp: H.scc_of_in_sccs_verts verts_H) next assume "H.scc_of (head G a) \ G.sccs_verts" with scc_of_head_inter have "G.scc_of (tail G a) = H.scc_of (head G a)" by (metis G.scc_of_in_sccs_verts G.sccs_verts_disjoint a_in empty_iff G.tail_in_verts) with G_scc_of_tail_not_in show False using ends_in by (auto simp: H.scc_of_in_sccs_verts verts_H) qed lemma scc_verts_da: "H.sccs_verts = (G.sccs_verts - {G.scc_of (tail G a)}) \ {H.scc_of (tail G a), H.scc_of (head G a)}" (is "?L = ?R") proof (intro set_eqI iffI) fix S assume "S \ ?L" then obtain u where "u \ verts G" "S = H.scc_of u" by (auto simp: verts_H H.sccs_verts_conv_scc_of) moreover then have "G.scc_of (tail G a) \ H.scc_of u" using \S \ ?L\ G_scc_of_tail_not_in by auto ultimately show "S \ ?R" unfolding G.sccs_verts_conv_scc_of by (cases "u \ G.scc_of (tail G a)") (auto dest: scc_of_G_tail scc_of_other) next fix S assume "S \ ?R" show "S \ ?L" proof (cases "S \ G.sccs_verts") case True with \S \ ?R\ obtain u where u: "u \ verts G" "S = G.scc_of u" and "S \ G.scc_of (tail G a)" using H_scc_of_a_not_in by (auto simp: G.sccs_verts_conv_scc_of) then have "G.scc_of u \ G.scc_of (tail G a) = {}" using ends_in by (intro G.sccs_verts_disjoint) (auto simp: G.scc_of_in_sccs_verts) then have "u \ G.scc_of (tail G a)" using u by (auto dest: G.in_scc_of_self) with u show ?thesis using scc_of_other by (auto simp: H.sccs_verts_conv_scc_of verts_H G.sccs_verts_conv_scc_of) next case False with \S \ ?R\ ends_in show ?thesis by (auto simp: H.sccs_verts_conv_scc_of verts_H) qed qed lemma card_sccs_da: "card H.sccs = Suc (card G.sccs)" using H_scc_of_a_not_in ends_in unfolding G.card_sccs_verts[symmetric] H.card_sccs_verts[symmetric] scc_verts_da by (simp add: card_insert_if G.finite_sccs_verts scc_of_tail_neq_head card_Suc_Diff1 G.scc_of_in_sccs_verts del: card_Diff_insert) end sublocale bidel_arc_not_biconnected \ bidel_arc_same_face proof note a_in moreover from a_in have "head G a \ tail G ` G.face_cycle_set a" by (simp add: G.heads_face_cycle_set[symmetric]) moreover have "tail G a \ tail G ` G.face_cycle_set a" by simp ultimately obtain p where p: "G.trail (head G a) p (tail G a)" "set p \ G.face_cycle_set a" by (rule G.obtain_trail_in_fcs') define p' where "p' = G.awalk_to_apath p" from p have p': "G.apath (head G a) p' (tail G a)" "set p' \ G.face_cycle_set a" by (auto simp: G.trail_def p'_def dest: G.apath_awalk_to_apath G.awalk_to_apath_subset) then have "set p' \ arcs G" using a_in by (blast dest: G.in_face_cycle_setD) have "\set p' \ arcs H" proof assume "set p' \ arcs H" then have "H.awalk (head G a) p' (tail G a)" using p' by (auto simp: G.apath_def arcs_H intro: H_awalkI) then show False using not_reach_a by (metis H.symmetric_reachable' H.reachable_awalk) qed then have "set p' \ {a,a'} \ {}" using \set p' \ arcs G\ by (auto simp: arcs_H) moreover have "a \ set p'" proof assume "a \ set p'" then have "head G a \ set (tl (G.awalk_verts (head G a) p'))" using \G.apath _ p' _\ by (cases p') (auto simp: G.set_awalk_verts G.apath_def G.awalk_Cons_iff, metis imageI) moreover have "head G a \ set (tl (G.awalk_verts (head G a) p'))" using \G.apath _ p' _\ by (cases p') (auto simp: G.apath_def) ultimately show False by contradiction qed ultimately have "a' \ G.face_cycle_set a" using p'(2) by auto then show "G.face_cycle_set a' = G.face_cycle_set a" using G.face_cycle_set_parts by auto qed locale bidel_arc_tail_conn = bidel_arc + assumes conn_tail: "tail G a \ H.isolated_verts" locale bidel_arc_head_conn = bidel_arc + assumes conn_head: "head G a \ H.isolated_verts" locale bidel_arc_tail_isolated = bidel_arc + assumes isolated_tail: "tail G a \ H.isolated_verts" locale bidel_arc_head_isolated = bidel_arc + assumes isolated_head: "head G a \ H.isolated_verts" begin lemma G_edge_succ_a'_no_loop: assumes no_loop_a: "head G a \ tail G a" shows G_edge_succ_a': "edge_succ M a' = a'" (is ?t2) proof - have *: "out_arcs G (tail G a') = {a'}" using a_in isolated_head no_loop_a by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a' \ {a'}" using G.edge_succ_cyclic[of "tail G a'"] apply (rule eq_on_cyclic_on_iff1[where x="a'"]) using * a_in a_neq_a' no_loop_a by simp_all then show ?thesis by auto qed lemma G_face_cycle_succ_a_no_loop: assumes no_loop_a: "head G a \ tail G a" shows "G.face_cycle_succ a = a'" using assms by (auto simp: G.face_cycle_succ_def G_edge_succ_a'_no_loop) end locale bidel_arc_same_face_tail_conn = bidel_arc_same_face + bidel_arc_tail_conn begin definition a_neigh :: 'b where "a_neigh \ SOME b. G.face_cycle_succ b = a" lemma face_cycle_succ_a_neigh: "G.face_cycle_succ a_neigh = a" proof - have "\b. G.face_cycle_succ b = a" by (metis G.face_cycle_succ_pred) then show ?thesis unfolding a_neigh_def by (rule someI_ex) qed lemma a_neigh_in: "a_neigh \ arcs G" using a_in by (metis face_cycle_succ_a_neigh G.face_cycle_succ_closed) lemma a_neigh_neq_a: "a_neigh \ a" proof assume "a_neigh = a" then have "G.face_cycle_set a = {a}" using face_cycle_succ_a_neigh by (simp add: G.fcs_x_eq_x) with a_neq_a' same_face G.face_cycle_set_self[of a'] show False by simp qed lemma a_neigh_neq_a': "a_neigh \ a'" proof assume A: "a_neigh = a'" have a_in_oa: "a \ out_arcs G (tail G a)" using a_in by auto have cyc: "cyclic_on (edge_succ M) (out_arcs G (tail G a))" using a_in by (intro G.edge_succ_cyclic) auto from A have "G.face_cycle_succ a' = a" by (metis face_cycle_succ_a_neigh) then have "edge_succ M a = a " by (auto simp: G.face_cycle_succ_def) then have "card (out_arcs G (tail G a)) = 1" using cyc a_in by (auto elim: eq_on_cyclic_on_iff1) then have "out_arcs G (tail G a) = {a}" using a_in_oa by (auto simp del: in_out_arcs_conv dest: card_eq_SucD) then show False using conn_tail a_in by (auto simp: H.isolated_verts_def arcs_H tail_H verts_H out_arcs_def) qed lemma edge_rev_a_neigh_neq: "edge_rev M a_neigh \ a'" by (metis a_neigh_neq_a G.arev_arev) lemma edge_succ_a_neq: "edge_succ M a \ a'" proof assume "edge_succ M a = a'" then have "G.face_cycle_set a' = {a'}" - using face_cycle_succ_a_neigh by (auto simp: G.face_cycle_set_altdef id_funpow_id G.face_cycle_succ_def) + using face_cycle_succ_a_neigh + by auto (metis G.arev_arev_raw G.face_cycle_succ_def G.fcs_x_eq_x a_in comp_apply singletonD) with a_neq_a' same_face G.face_cycle_set_self[of a] show False by auto qed lemma H_face_cycle_succ_a_neigh: "H.face_cycle_succ a_neigh = G.face_cycle_succ a'" using face_cycle_succ_a_neigh edge_succ_a_neq edge_rev_a_neigh_neq a_neigh_neq_a a_neigh_neq_a' a_neigh_in unfolding H.face_cycle_succ_def G.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.bij_edge_succ) lemma H_fcs_a_neigh: "H.face_cycle_set a_neigh = segment G.face_cycle_succ a' a" (is "?L = ?R") proof - { fix n assume A: "0 < n" "n < funpow_dist1 G.face_cycle_succ a' a" then have *: "(G.face_cycle_succ ^^ n) a' \ segment G.face_cycle_succ a' a" using a_in_o by (auto simp: segment_altdef) then have "(G.face_cycle_succ ^^ n) a' \ {a,a'}" "(G.face_cycle_succ ^^ n) a' \ arcs G" using not_in_segment1[OF a_in_o] not_in_segment2[of a G.face_cycle_succ a'] by (auto simp: segment_altdef a_in_o) } note X = this { fix n assume "0 < n" "n < funpow_dist1 G.face_cycle_succ a' a" then have "(H.face_cycle_succ ^^ n) a_neigh = (G.face_cycle_succ ^^ n) a'" proof (induct n) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "n=0") case True then show ?thesis by (simp add: H_face_cycle_succ_a_neigh) next case False then have "(H.face_cycle_succ ^^ n) a_neigh = (G.face_cycle_succ ^^ n) a'" using Suc by simp then show ?thesis using X[of "Suc n"] X[of n] False Suc by (simp add: H_fcs_eq_G_fcs) qed qed } note Y = this have fcs_a'_neq_a: "G.face_cycle_succ a' \ a" by (metis (no_types) a_neigh_neq_a' G.face_cycle_pred_succ face_cycle_succ_a_neigh) show ?thesis proof (intro set_eqI iffI) fix b assume "b \ ?L" define m where "m = funpow_dist1 G.face_cycle_succ a' a - 1" have b_in0: "b \ orbit H.face_cycle_succ (a_neigh)" using \b \ ?L\ by (simp add: H.face_cycle_set_def[symmetric]) have "0 < m" by (auto simp: m_def) (metis a_neigh_neq_a' G.face_cycle_pred_succ G.face_cycle_set_def G.face_cycle_set_self G.face_cycle_set_succ face_cycle_succ_a_neigh funpow_dist_0_eq neq0_conv same_face) then have pos_dist: "0 < funpow_dist1 H.face_cycle_succ a_neigh b" by (simp add: m_def) have *: "(G.face_cycle_succ ^^ Suc m) a' = a" - using a_in_o by (simp add: m_def funpow_simp_l funpow_dist1_prop del: funpow.simps) + using a_in_o by (simp add: m_def funpow_dist1_prop del: funpow.simps) have "(H.face_cycle_succ ^^ m) a_neigh = a_neigh" proof - have "a = G.face_cycle_succ ((H.face_cycle_succ ^^ m) a_neigh)" using * \0 < m\ by (simp add: Y m_def) then show ?thesis using face_cycle_succ_a_neigh by (metis G.face_cycle_pred_succ) qed then have "funpow_dist1 H.face_cycle_succ a_neigh b \ m" using \0 < m\ b_in0 by (intro funpow_dist1_le_self) simp_all also have "\ < funpow_dist1 G.face_cycle_succ a' a" unfolding m_def by simp finally have dist_less: "funpow_dist1 H.face_cycle_succ a_neigh b < funpow_dist1 G.face_cycle_succ a' a" . have "b = (H.face_cycle_succ ^^ funpow_dist1 H.face_cycle_succ a_neigh b) a_neigh" using b_in0 by (simp add: funpow_dist1_prop del: funpow.simps) also have "\ = (G.face_cycle_succ ^^ funpow_dist1 H.face_cycle_succ a_neigh b) a'" using pos_dist dist_less by (rule Y) also have "\ \ ?R" using pos_dist dist_less by (simp add: segment_altdef a_in_o del: funpow.simps) finally show "b \ ?R" . next fix b assume "b \ ?R" then show "b \ ?L" using Y by (auto simp: segment_altdef a_in_o H.face_cycle_set_altdef Suc_le_eq) metis qed qed end locale bidel_arc_isolated_loop = bidel_arc_biconnected + bidel_arc_tail_isolated begin lemma loop_a[simp]: "head G a = tail G a" using isolated_tail reach_a by (auto simp: H.isolated_verts_def elim: H.converse_reachable_cases dest: out_arcs_emptyD_dominates) end sublocale bidel_arc_isolated_loop \ bidel_arc_head_isolated using isolated_tail loop_a by unfold_locales simp context bidel_arc_isolated_loop begin text \The edges @{term a} and @{term a'} form a loop on an otherwise isolated vertex \ lemma card_isolated_verts_da: "card H.isolated_verts = Suc (card G.isolated_verts)" by (simp add: card_isolated_verts_da0 isolated_tail) lemma G_edge_succ_a[simp]: "edge_succ M a = a'" (is ?t1) and G_edge_succ_a'[simp]: "edge_succ M a' = a" (is ?t2) proof - have *: "out_arcs G (tail G a) = {a,a'}" using a_in isolated_tail by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a' \ {a,a'}" "edge_succ M a' \ a'" using G.edge_succ_cyclic[of "tail G a'"] apply (rule eq_on_cyclic_on_iff1[where x="a'"]) using * a_in a_neq_a' loop_a by auto moreover obtain "edge_succ M a \ {a,a'}" "edge_succ M a \ a" using G.edge_succ_cyclic[of "tail G a"] apply (rule eq_on_cyclic_on_iff1[where x="a"]) using * a_in a_neq_a' loop_a by auto ultimately show ?t1 ?t2 by auto qed lemma G_face_cycle_succ_a[simp]: "G.face_cycle_succ a = a" and G_face_cycle_succ_a'[simp]: "G.face_cycle_succ a' = a'" by (auto simp: G.face_cycle_succ_def) lemma G_face_cycle_set_a[simp]: "G.face_cycle_set a = {a}" and G_face_cycle_set_a'[simp]: "G.face_cycle_set a' = {a'}" unfolding G.fcs_x_eq_x[symmetric] by simp_all end sublocale bidel_arc_isolated_loop \ bidel_arc_diff_face using a_neq_a' by unfold_locales auto context bidel_arc_isolated_loop begin lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (Suc (card H.face_cycle_sets))" unfolding card_face_cycle_sets_da' using diff_face card_fcs_aa'_le by (auto simp: card_insert_if S_def G.segment_face_cycle_x_x_eq) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_two_isolated = bidel_arc_not_biconnected + bidel_arc_tail_isolated + bidel_arc_head_isolated begin text \@{term "tail G a"} and @{term "head G a"} form an SCC with @{term a} and @{term a'} as the only arcs.\ lemma no_loop_a: "head G a \ tail G a" using not_reach_a a_in by (auto simp: verts_H) lemma card_isolated_verts_da: "card H.isolated_verts = Suc (Suc (card G.isolated_verts))" using no_loop_a isolated_tail isolated_head by (simp add: card_isolated_verts_da0 card_insert_if) lemma G_edge_succ_a'[simp]: "edge_succ M a' = a'" using G_edge_succ_a'_no_loop no_loop_a by simp lemma G_edge_succ_a[simp]: "edge_succ M a = a" proof - have *: "out_arcs G (tail G a) = {a}" using a_in isolated_tail isolated_head no_loop_a by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a \ {a}" using G.edge_succ_cyclic[of "tail G a"] apply (rule eq_on_cyclic_on_iff1[where x="a"]) using * a_in a_neq_a' no_loop_a by simp_all then show ?thesis by auto qed lemma G_face_cycle_succ_a[simp]: "G.face_cycle_succ a = a'" and G_face_cycle_succ_a'[simp]: "G.face_cycle_succ a' = a" by (auto simp: G.face_cycle_succ_def) lemma G_face_cycle_set_a[simp]: "G.face_cycle_set a = {a,a'}" (is ?t1) and G_face_cycle_set_a'[simp]: "G.face_cycle_set a' = {a,a'}" (is ?t2) proof - { fix n have "(G.face_cycle_succ ^^ n) a \ {a,a'}" "(G.face_cycle_succ ^^ n) a' \ {a,a'}" by (induct n) auto } then show ?t1 ?t2 by (auto simp: G.face_cycle_set_altdef intro: exI[where x=0] exI[where x=1]) qed lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (card H.face_cycle_sets)" unfolding card_face_cycle_sets_da0 using card_fcs_aa'_le by simp lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_tail_not_isol = bidel_arc_not_biconnected + bidel_arc_tail_conn sublocale bidel_arc_tail_not_isol \ bidel_arc_same_face_tail_conn by unfold_locales locale bidel_arc_only_tail_not_isol = bidel_arc_tail_not_isol + bidel_arc_head_isolated context bidel_arc_only_tail_not_isol begin lemma card_isolated_verts_da: "card H.isolated_verts = Suc (card G.isolated_verts)" using isolated_head conn_tail by (simp add: card_isolated_verts_da0) lemma segment_a'_a_ne: "segment G.face_cycle_succ a' a \ {}" unfolding H_fcs_a_neigh[symmetric] by auto lemma segment_a_a'_e: "segment G.face_cycle_succ a a' = {}" proof - have "a' = G.face_cycle_succ a" using tail_neq_head by (simp add: G_face_cycle_succ_a_no_loop) then show ?thesis by (auto simp: segment1_empty) qed lemma card_face_cycle_sets_da: "card H.face_cycle_sets = card G.face_cycle_sets" unfolding card_face_cycle_sets_da' using segment_a'_a_ne segment_a_a'_e card_fcs_gt_0 by (simp add: card_insert_if) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_only_head_not_isol = bidel_arc_not_biconnected + bidel_arc_head_conn + bidel_arc_tail_isolated begin interpretation rev: bidel_arc G M a' using a_in by unfold_locales simp interpretation rev: bidel_arc_only_tail_not_isol G M a' using a_in not_reach_a by unfold_locales (auto simp: rev_H isolated_tail conn_head dest: H.symmetric_reachable') lemma euler_genus_da: "H.euler_genus = G.euler_genus" using rev.euler_genus_da by (simp add: rev_H rev_HM) end locale bidel_arc_two_not_isol = bidel_arc_tail_not_isol + bidel_arc_head_conn begin lemma isolated_verts_da: "H.isolated_verts = G.isolated_verts" using conn_head conn_tail by (subst isolated_da_pre) simp lemma segment_a'_a_ne': "segment G.face_cycle_succ a' a \ {}" unfolding H_fcs_a_neigh[symmetric] by auto interpretation rev: bidel_arc_tail_not_isol G M a' using arcs_in not_reach_a rev_H conn_head by unfold_locales (auto dest: H.symmetric_reachable') lemma segment_a_a'_ne': "segment G.face_cycle_succ a a' \ {}" using rev.H_fcs_a_neigh[symmetric] rev_H rev_HM by auto lemma card_face_cycle_sets_da: "card H.face_cycle_sets = Suc (card G.face_cycle_sets)" unfolding card_face_cycle_sets_da' using segment_a'_a_ne' segment_a_a'_ne' card_fcs_gt_0 by (simp add: segments_neq card_insert_if) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_biconnected_non_triv = bidel_arc_biconnected + bidel_arc_tail_conn sublocale bidel_arc_biconnected_non_triv \ bidel_arc_head_conn by unfold_locales (metis (mono_tags) G.in_sccs_verts_conv_reachable G.symmetric_reachable' H.isolated_verts_in_sccs conn_tail empty_iff insert_iff reach_a reachable_daD sccs_verts_da) context bidel_arc_biconnected_non_triv begin lemma isolated_verts_da: "H.isolated_verts = G.isolated_verts" using conn_head conn_tail by (subst isolated_da_pre) simp end locale bidel_arc_biconnected_same = bidel_arc_biconnected_non_triv + bidel_arc_same_face sublocale bidel_arc_biconnected_same \ bidel_arc_same_face_tail_conn by unfold_locales context bidel_arc_biconnected_same begin interpretation rev: bidel_arc_same_face_tail_conn G M a' using arcs_in conn_head by unfold_locales (auto simp: same_face rev_H) lemma card_face_cycle_sets_da: "Suc (card H.face_cycle_sets) \ (card G.face_cycle_sets)" unfolding card_face_cycle_sets_da' using card_fcs_gt_0 by linarith lemma euler_genus_da: "H.euler_genus \ G.euler_genus" using card_face_cycle_sets_da unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_sccs_da ) end locale bidel_arc_biconnected_diff = bidel_arc_biconnected_non_triv + bidel_arc_diff_face begin lemma fcs_not_triv: "G.face_cycle_set a \ {a} \ G.face_cycle_set a' \ {a'}" proof (rule ccontr) assume "\?thesis" then have "G.face_cycle_succ a = a" "G.face_cycle_succ a' = a'" by (auto simp: G.fcs_x_eq_x) then have *: "edge_succ M a = a'" "edge_succ M a' = a" by (auto simp: G.face_cycle_succ_def) then have "(edge_succ M ^^ 2) a = a" by (auto simp: eval_nat_numeral) { fix n have "(edge_succ M ^^ 2) a = a" by (auto simp: * eval_nat_numeral) then have "(edge_succ M ^^ n) a = (edge_succ M ^^ (n mod 2)) a" by (auto simp: funpow_mod_eq) moreover have "n mod 2 = 0 \ n mod 2 = 1" by auto ultimately have "(edge_succ M ^^ n) a \ {a, a'}" by (auto simp: *) } then have "orbit (edge_succ M) a = {a, a'}" by (auto simp: orbit_altdef_permutation[OF G.permutation_edge_succ] exI[where x=0] exI[where x=1] *) have "out_arcs G (tail G a) \ {a,a'}" proof - have "cyclic_on (edge_succ M) (out_arcs G (tail G a))" using arcs_in by (intro G.edge_succ_cyclic) auto then have "orbit (edge_succ M) a = out_arcs G (tail G a)" using arcs_in by (intro orbit_cyclic_eq3) auto then show ?thesis using \orbit _ _ = {_, _}\ by auto qed then have "out_arcs H (tail G a) = {}" by (auto simp: arcs_H tail_H) then have "tail G a \ H.isolated_verts" using arcs_in by (simp add: H.isolated_verts_def verts_H) then show False using conn_tail by contradiction qed lemma S_ne: "S \ {}" using fcs_not_triv by (auto simp: S_def G.segment_face_cycle_x_x_eq) lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (card H.face_cycle_sets)" unfolding card_face_cycle_sets_da' using S_ne diff_face card_fcs_aa'_le by simp lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_sccs_da card_face_cycle_sets_da) end context bidel_arc begin lemma euler_genus_da: "H.euler_genus \ G.euler_genus" proof - let ?biconnected = "tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" let ?isol_tail = "tail G a \ H.isolated_verts" let ?isol_head = "head G a \ H.isolated_verts" let ?same_face = "G.face_cycle_set a' = G.face_cycle_set a" { assume ?biconnected ?isol_tail then interpret EG: bidel_arc_isolated_loop by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume ?biconnected "\?isol_tail" ?same_face then interpret EG: bidel_arc_biconnected_same by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume ?biconnected "\?isol_tail" "\?same_face" then interpret EG: bidel_arc_biconnected_diff by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" ?isol_tail ?isol_head then interpret EG: bidel_arc_two_isolated by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" "\?isol_tail" ?isol_head then interpret EG: bidel_arc_only_tail_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" ?isol_tail "\?isol_head" then interpret EG: bidel_arc_only_head_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" "\?isol_tail" "\?isol_head" then interpret EG: bidel_arc_two_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } ultimately show ?thesis by satx qed end subsection \Modifying @{term edge_rev}\ definition (in pre_digraph_map) rev_swap :: "'b \ 'b \ 'b pre_map" where "rev_swap a b = \ edge_rev = perm_swap a b (edge_rev M), edge_succ = perm_swap a b (edge_succ M) \" context digraph_map begin lemma digraph_map_rev_swap: assumes "arc_to_ends G a = arc_to_ends G b" "{a,b} \ arcs G" shows "digraph_map G (rev_swap a b)" proof let ?M' = "rev_swap a b" have tail_swap: "\x. tail G ((a \\<^sub>F b) x) = tail G x" using assms by (case_tac "x \ {a,b}") (auto simp: arc_to_ends_def) have swap_in_arcs: "\x. (a \\<^sub>F b) x \ arcs G \ x \ arcs G" using assms by (case_tac "x \ {a,b}") auto have es_perm: "edge_succ ?M' permutes arcs G" using assms edge_succ_permutes unfolding permutes_conv_has_dom by (auto simp: rev_swap_def has_dom_perm_swap) { fix x show "(x \ arcs G) = (edge_rev (rev_swap a b) x \ x)" using assms(2) by (cases "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def arev_dom swap_id_eq split: if_splits) next fix x assume "x \ arcs G" then show "edge_rev ?M' (edge_rev ?M' x) = x" by (auto simp: rev_swap_def perm_swap_comp[symmetric]) next fix x assume "x \ arcs G" then show "tail G (edge_rev ?M' x) = head G x" using assms by (case_tac "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def tail_swap arc_to_ends_def) next show "edge_succ ?M' permutes arcs G" by fact next fix v assume A: "v \ verts G" "out_arcs G v \ {}" then obtain c where "c \ out_arcs G v" by blast have "inj (perm_swap a b (edge_succ M))" by (simp add: bij_is_inj bij_edge_succ) have "out_arcs G v = (a \\<^sub>F b) ` out_arcs G v" by (auto simp: tail_swap swap_swap_id swap_in_arcs intro: image_eqI[where x="(a \\<^sub>F b) y" for y]) also have "(a \\<^sub>F b) ` out_arcs G v = (a \\<^sub>F b) ` orbit (edge_succ M) ((a \\<^sub>F b) c)" using edge_succ_cyclic using A \c \ _\ by (intro arg_cong[where f="(`) (a \\<^sub>F b)"]) (intro orbit_cyclic_eq3[symmetric], auto simp: swap_in_arcs tail_swap) also have "\ = orbit (edge_succ ?M') c" by (simp add: orbit_perm_swap rev_swap_def) finally have oa_orb: "out_arcs G v = orbit (edge_succ ?M') c" . show "cyclic_on (edge_succ ?M') (out_arcs G v)" unfolding oa_orb using es_perm finite_arcs by (rule cyclic_on_orbit) } qed lemma euler_genus_rev_swap: assumes "arc_to_ends G a = arc_to_ends G b" "{a,b} \ arcs G" shows "pre_digraph_map.euler_genus G (rev_swap a b) = euler_genus" proof - let ?M' = "rev_swap a b" interpret G': digraph_map G ?M' using assms by (rule digraph_map_rev_swap) have swap_in_arcs: "\x. (a \\<^sub>F b) x \ arcs G \ x \ arcs G" using assms by (case_tac "x \ {a,b}") auto have G'_fcs: "G'.face_cycle_succ = perm_swap a b face_cycle_succ" unfolding G'.face_cycle_succ_def face_cycle_succ_def by (auto simp: fun_eq_iff rev_swap_def perm_swap_comp) have "\x. G'.face_cycle_set x = (a \\<^sub>F b) ` face_cycle_set ((a \\<^sub>F b) x)" by (auto simp: face_cycle_set_def G'.face_cycle_set_def orbit_perm_swap G'_fcs imageI) then have "G'.face_cycle_sets = (\S. (a \\<^sub>F b) ` S) ` face_cycle_sets" by (auto simp: pre_digraph_map.face_cycle_sets_def swap_in_arcs) (metis swap_swap_id image_eqI swap_in_arcs) then have "card G'.face_cycle_sets = card ((\S. (a \\<^sub>F b) ` S) ` face_cycle_sets)" by simp also have "\ = card face_cycle_sets" by (rule card_image) (rule inj_on_f_imageI[where S="UNIV"], auto) finally show "pre_digraph_map.euler_genus G ?M' = euler_genus" unfolding pre_digraph_map.euler_genus_def pre_digraph_map.euler_char_def by simp qed end subsection \Conclusion\ lemma bidirected_subgraph_obtain: assumes sg: "subgraph H G" "arcs H \ arcs G" assumes fin: "finite (arcs G)" assumes bidir: "\rev. bidirected_digraph G rev" "\rev. bidirected_digraph H rev" obtains a a' where "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" proof - obtain a where a: "a \ arcs G - arcs H" using sg by blast obtain rev_G rev_H where rev: "bidirected_digraph G rev_G" "bidirected_digraph H rev_H" using bidir by blast interpret G: bidirected_digraph G rev_G by (rule rev) interpret H: bidirected_digraph H rev_H by (rule rev) have sg_props: "arcs H \ arcs G" "tail H = tail G" "head H = head G" using sg by (auto simp: subgraph_def compatible_def) { fix w1 w2 assume A: "tail G a = w1" "head G a = w2" have "in_arcs H w1 \ out_arcs H w2 = rev_H ` (out_arcs H w1 \ in_arcs H w2)" (is "?Sh = _") unfolding H.in_arcs_eq by (simp add: image_Int image_image H.inj_on_arev) then have "card (in_arcs H w1 \ out_arcs H w2) = card (out_arcs H w1 \ in_arcs H w2)" by (metis card_image H.arev_arev inj_on_inverseI) also have "\ < card (out_arcs G w1 \ in_arcs G w2)" (is "card ?Sh1 < card ?Sg1") proof (rule psubset_card_mono) show "finite ?Sg1" using fin by (auto simp: out_arcs_def) show "?Sh1 \ ?Sg1" using A a sg_props by auto qed also have "?Sg1 = rev_G ` (in_arcs G w1 \ out_arcs G w2)" (is "_ = _ ` ?Sg") unfolding G.in_arcs_eq by (simp add: image_Int image_image G.inj_on_arev) also have "card \ = card ?Sg" by (metis card_image G.arev_arev inj_on_inverseI) finally have card_less: "card ?Sh < card ?Sg" . have S_ss: "?Sh \ ?Sg" using sg_props by auto have ?thesis proof (cases "w1 = w2") case True have "card (?Sh - {a}) = card ?Sh" using a by (intro arg_cong[where f=card]) auto also have "\ < card ?Sg - 1" proof - from True have "even (card ?Sg)" "even (card ?Sh)" by (auto simp: G.even_card_loops H.even_card_loops) then show ?thesis using card_less by simp (metis Suc_pred even_Suc le_neq_implies_less lessE less_Suc_eq_le zero_less_Suc) qed also have "\ = card (?Sg - {a})" using fin a A True by (auto simp: out_arcs_def card_Diff_singleton) finally have card_diff_a_less: "card (?Sh - {a}) < card (?Sg - {a})" . moreover from S_ss have "?Sh - {a} \ ?Sg - {a}" using S_ss by blast ultimately have "?Sh - {a} \ ?Sg - {a}" by (intro card_psubset) auto then obtain a' where "a' \ (?Sg - {a})- ?Sh" by blast then have "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" using A a sg_props by auto then show ?thesis .. next case False from card_less S_ss have "?Sh \ ?Sg" by auto then obtain a' where "a' \ ?Sg - ?Sh" by blast then have "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" using A a sg_props False by auto then show ?thesis .. qed } then show ?thesis by simp qed lemma subgraph_euler_genus_le: assumes G: "subgraph H G" "digraph_map G GM" and H: "\rev. bidirected_digraph H rev" obtains HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" proof - let ?d = "\G. card (arcs G) + card (verts G) - card (arcs H) - card (verts H)" from H obtain rev_H where "bidirected_digraph H rev_H" by blast then interpret H: bidirected_digraph H rev_H . from G have "\HM. digraph_map H HM \ pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" proof (induct "?d G" arbitrary: G GM rule: less_induct) case less from less interpret G: digraph_map G GM by - have H_ss: "arcs H \ arcs G" "verts H \ verts G" using \subgraph H G\ by auto then have card_le: "card (arcs H) \ card (arcs G)" "card (verts H) \ card (verts G)" by (auto intro: card_mono) have ends: "tail H = tail G" "head H = head G" using \subgraph H G\ by (auto simp: compatible_def) show ?case proof (cases "?d G = 0") case True then have "card (arcs H) = card (arcs G)" "card (verts H) = card (verts G)" using card_le by linarith+ then have "arcs H = arcs G" "verts H = verts G" using H_ss by (auto simp: card_subset_eq) then have "H = G" using \subgraph H G\ by (auto simp: compatible_def) then have "digraph_map H GM \ pre_digraph_map.euler_genus H GM \ G.euler_genus" by auto then show ?thesis .. next case False then have H_ne: "(arcs G - arcs H) \ {} \ (verts G - verts H) \ {}" using H_ss card_le by auto { assume A: "arcs G - arcs H \ {}" then obtain a a' where aa': "{a, a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a' = tail G a" using H_ss \subgraph H G\ by (auto intro: bidirected_subgraph_obtain) let ?GM' = "G.rev_swap (edge_rev GM a) a'" interpret G': digraph_map G ?GM' using aa' by (intro G.digraph_map_rev_swap) (auto simp: arc_to_ends_def) interpret G': bidel_arc G ?GM' a using aa' by unfold_locales simp have "edge_rev GM a \ a" using aa' by (intro G.arev_neq) auto then have er_a: "edge_rev ?GM' a = a'" using \a' \ a\ by (auto simp: G.rev_swap_def perm_swap_def swap_id_eq dest: G.arev_neq) then have sg: "subgraph H G'.H" using H_ss aa' by (intro subgraphI) (auto simp: G'.verts_H G'.arcs_H G'.tail_H G'.head_H ends compatible_def intro: H.wf_digraph G'.H.wf_digraph) have "card {a,a'} \ card (arcs G)" using aa' by (intro card_mono) auto then obtain HM where HM: "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ G'.H.euler_genus" using aa' False by atomize_elim (rule less, auto simp: G'.verts_H G'.arcs_H card_insert_if sg er_a) have "G'.H.euler_genus \ G'.euler_genus" by (rule G'.euler_genus_da) also have "G'.euler_genus = G.euler_genus" using aa' by (auto simp: G.euler_genus_rev_swap arc_to_ends_def) finally have ?thesis using HM by auto } moreover { assume A: "arcs G - arcs H = {}" then have A': "verts G - verts H \ {}" and arcs_H: "arcs H = arcs G" using H_ss H_ne by auto then obtain v where v: "v \ verts G - verts H" by auto have card_lt: "card (verts H) < card (verts G)" using A' H_ss by (intro psubset_card_mono) auto have "out_arcs G v = out_arcs H v" using A H_ss by (auto simp: ends) then interpret G: del_vert_props G GM v using v by unfold_locales auto have "?d (G.del_vert v) < ?d G" using card_lt by (simp add: arcs_H G.arcs_dv G.card_verts_dv) moreover have "subgraph H (G.del_vert v)" using H_ss v by (auto simp: subgraph_def arcs_H G.arcs_dv G.verts_del_vert H.wf_digraph G.H.wf_digraph compatible_def G.tail_del_vert G.head_del_vert ends) moreover have "bidirected_digraph (G.del_vert v) (edge_rev GM)" using G.arev_dom by (intro G.H.bidirected_digraphI) (auto simp: G.arcs_dv) ultimately have ?thesis unfolding G.euler_genus_eq[symmetric] by (intro less) auto } ultimately show ?thesis by blast qed qed then obtain HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" by atomize_elim then show ?thesis .. qed lemma (in digraph_map) nonneg_euler_genus: "0 \ euler_genus" proof - define H where "H = \ verts = {}, arcs = {}, tail = tail G, head = head G \" then have H_simps: "verts H = {}" "arcs H = {}" "tail H = tail G" "head H = head G" by (simp_all add: H_def) interpret H: bidirected_digraph H id by unfold_locales (auto simp: H_def) have "wf_digraph H" "wf_digraph G" by unfold_locales then have "subgraph H G" by (intro subgraphI) (auto simp: H_def compatible_def) then obtain HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ euler_genus" by (rule subgraph_euler_genus_le) auto then interpret H: digraph_map H HM by - have "H.sccs = {}" proof - { fix x assume *: "x \ H.sccs_verts" then have "x = {}" by (auto dest: H.sccs_verts_subsets simp: H_simps) with * have False by (auto simp: H.in_sccs_verts_conv_reachable) } then show ?thesis by (auto simp: H.sccs_verts_conv) qed then have "H.euler_genus = 0" by (auto simp: H.euler_genus_def H.euler_char_def H.isolated_verts_def H.face_cycle_sets_def H_simps) then show ?thesis using \H.euler_genus \ _\ by simp qed lemma subgraph_comb_planar: assumes "subgraph G H" "comb_planar H" "\rev. bidirected_digraph G rev" shows "comb_planar G" proof - from \comb_planar H\ obtain HM where "digraph_map H HM" and H_genus: "pre_digraph_map.euler_genus H HM = 0" unfolding comb_planar_def by metis obtain GM where G: "digraph_map G GM" "pre_digraph_map.euler_genus G GM \ pre_digraph_map.euler_genus H HM" using assms(1) \digraph_map H HM\ assms(3) by (rule subgraph_euler_genus_le) interpret G: digraph_map G GM by fact show ?thesis using G H_genus G.nonneg_euler_genus unfolding comb_planar_def by auto qed end