Up to index of Isabelle/HOL/SizeChange
theory Correctness(* Title: HOL/Library/SCT_Theorem.thy ID: $Id: Correctness.thy,v 1.1 2007/11/06 16:44:54 krauss Exp $ Author: Alexander Krauss, TU Muenchen *) header "Proof of the Size-Change Principle" theory Correctness imports Main Ramsey Misc_Tools Criterion begin subsection {* Auxiliary definitions *} definition is_thread :: "nat => 'a sequence => ('a, 'a scg) ipath => bool" where "is_thread n ϑ p = (∀i≥n. eqlat p ϑ i)" definition is_fthread :: "'a sequence => ('a, 'a scg) ipath => nat => nat => bool" where "is_fthread ϑ mp i j = (∀k∈{i..<j}. eqlat mp ϑ k)" definition is_desc_fthread :: "'a sequence => ('a, 'a scg) ipath => nat => nat => bool" where "is_desc_fthread ϑ mp i j = (is_fthread ϑ mp i j ∧ (∃k∈{i..<j}. descat mp ϑ k))" definition "has_fth p i j n m = (∃ϑ. is_fthread ϑ p i j ∧ ϑ i = n ∧ ϑ j = m)" definition "has_desc_fth p i j n m = (∃ϑ. is_desc_fthread ϑ p i j ∧ ϑ i = n ∧ ϑ j = m)" subsection {* Everything is finite *} lemma finite_range: fixes f :: "nat => 'a" assumes fin: "finite (range f)" shows "∃x. ∃∞i. f i = x" proof (rule classical) assume "¬(∃x. ∃∞i. f i = x)" hence "∀x. ∃j. ∀i>j. f i ≠ x" unfolding INF_nat by blast with choice have "∃j. ∀x. ∀i>(j x). f i ≠ x" . then obtain j where neq: "!!x i. j x < i ==> f i ≠ x" by blast from fin have "finite (range (j o f))" by (auto simp:comp_def) with finite_nat_bounded obtain m where "range (j o f) ⊆ {..<m}" by blast hence "j (f m) < m" unfolding comp_def by auto with neq[of "f m" m] show ?thesis by blast qed lemma finite_range_ignore_prefix: fixes f :: "nat => 'a" assumes fA: "finite A" assumes inA: "∀x≥n. f x ∈ A" shows "finite (range f)" proof - have a: "UNIV = {0 ..< (n::nat)} ∪ { x. n ≤ x }" by auto have b: "range f = f ` {0 ..< n} ∪ f ` { x. n ≤ x }" (is "… = ?A ∪ ?B") by (unfold a) (simp add:image_Un) have "finite ?A" by (rule finite_imageI) simp moreover from inA have "?B ⊆ A" by auto from this fA have "finite ?B" by (rule finite_subset) ultimately show ?thesis using b by simp qed definition "finite_graph G = finite (dest_graph G)" definition "all_finite G = (∀n H m. has_edge G n H m --> finite_graph H)" definition "finite_acg A = (finite_graph A ∧ all_finite A)" definition "nodes G = fst ` dest_graph G ∪ snd ` snd ` dest_graph G" definition "edges G = fst ` snd ` dest_graph G" definition "smallnodes G = \<Union>(nodes ` edges G)" lemma thread_image_nodes: assumes th: "is_thread n ϑ p" shows "∀i≥n. ϑ i ∈ nodes (snd (p i))" using prems unfolding is_thread_def has_edge_def nodes_def by force lemma finite_nodes: "finite_graph G ==> finite (nodes G)" unfolding finite_graph_def nodes_def by auto lemma nodes_subgraph: "A ≤ B ==> nodes A ⊆ nodes B" unfolding graph_leq_def nodes_def by auto lemma finite_edges: "finite_graph G ==> finite (edges G)" unfolding finite_graph_def edges_def by auto lemma edges_sum[simp]: "edges (A + B) = edges A ∪ edges B" unfolding edges_def graph_plus_def by auto lemma nodes_sum[simp]: "nodes (A + B) = nodes A ∪ nodes B" unfolding nodes_def graph_plus_def by auto lemma finite_acg_subset: "A ≤ B ==> finite_acg B ==> finite_acg A" unfolding finite_acg_def finite_graph_def all_finite_def has_edge_def graph_leq_def by (auto elim:finite_subset) lemma scg_finite: fixes G :: "'a scg" assumes fin: "finite (nodes G)" shows "finite_graph G" unfolding finite_graph_def proof (rule finite_subset) show "dest_graph G ⊆ nodes G × UNIV × nodes G" (is "_ ⊆ ?P") unfolding nodes_def by force show "finite ?P" by (intro finite_cartesian_product fin finite) qed lemma smallnodes_sum[simp]: "smallnodes (A + B) = smallnodes A ∪ smallnodes B" unfolding smallnodes_def by auto lemma in_smallnodes: fixes A :: "'a acg" assumes e: "has_edge A x G y" shows "nodes G ⊆ smallnodes A" proof - have "fst (snd (x, G, y)) ∈ fst ` snd ` dest_graph A" unfolding has_edge_def by (rule imageI)+ (rule e[unfolded has_edge_def]) then have "G ∈ edges A" unfolding edges_def by simp thus ?thesis unfolding smallnodes_def by blast qed lemma finite_smallnodes: assumes fA: "finite_acg A" shows "finite (smallnodes A)" unfolding smallnodes_def edges_def proof from fA show "finite (nodes ` fst ` snd ` dest_graph A)" unfolding finite_acg_def finite_graph_def by simp fix M assume "M ∈ nodes ` fst ` snd ` dest_graph A" then obtain n G m where M: "M = nodes G" and nGm: "(n,G,m) ∈ dest_graph A" by auto from fA have "all_finite A" unfolding finite_acg_def by simp with nGm have "finite_graph G" unfolding all_finite_def has_edge_def by auto with finite_nodes show "finite M" unfolding finite_graph_def M . qed lemma nodes_tcl: "nodes (tcl A) = nodes A" proof show "nodes A ⊆ nodes (tcl A)" apply (rule nodes_subgraph) by (subst tcl_unfold_right) simp show "nodes (tcl A) ⊆ nodes A" proof fix x assume "x ∈ nodes (tcl A)" then obtain z G y where z: "z ∈ dest_graph (tcl A)" and dis: "z = (x, G, y) ∨ z = (y, G, x)" unfolding nodes_def by auto force+ from dis show "x ∈ nodes A" proof assume "z = (x, G, y)" with z have "has_edge (tcl A) x G y" unfolding has_edge_def by simp then obtain n where "n > 0 " and An: "has_edge (A ^ n) x G y" unfolding in_tcl by auto then obtain n' where "n = Suc n'" by (cases n, auto) hence "A ^ n = A * A ^ n'" by (simp add:power_Suc) with An obtain e k where "has_edge A x e k" by (auto simp:in_grcomp) thus "x ∈ nodes A" unfolding has_edge_def nodes_def by force next assume "z = (y, G, x)" with z have "has_edge (tcl A) y G x" unfolding has_edge_def by simp then obtain n where "n > 0 " and An: "has_edge (A ^ n) y G x" unfolding in_tcl by auto then obtain n' where "n = Suc n'" by (cases n, auto) hence "A ^ n = A ^ n' * A" by (simp add:power_Suc power_commutes) with An obtain e k where "has_edge A k e x" by (auto simp:in_grcomp) thus "x ∈ nodes A" unfolding has_edge_def nodes_def by force qed qed qed lemma smallnodes_tcl: fixes A :: "'a acg" shows "smallnodes (tcl A) = smallnodes A" proof (intro equalityI subsetI) fix n assume "n ∈ smallnodes (tcl A)" then obtain x G y where edge: "has_edge (tcl A) x G y" and "n ∈ nodes G" unfolding smallnodes_def edges_def has_edge_def by auto from `n ∈ nodes G` have "n ∈ fst ` dest_graph G ∨ n ∈ snd ` snd ` dest_graph G" (is "?A ∨ ?B") unfolding nodes_def by blast thus "n ∈ smallnodes A" proof assume ?A then obtain m e where A: "has_edge G n e m" unfolding has_edge_def by auto have "tcl A = A * star A" unfolding tcl_def by (simp add: star_commute[of A A A, simplified]) with edge have "has_edge (A * star A) x G y" by simp then obtain H H' z where AH: "has_edge A x H z" and G: "G = H * H'" by (auto simp:in_grcomp) from A obtain m' e' where "has_edge H n e' m'" by (auto simp:G in_grcomp) hence "n ∈ nodes H" unfolding nodes_def has_edge_def by force with in_smallnodes[OF AH] show "n ∈ smallnodes A" .. next assume ?B then obtain m e where B: "has_edge G m e n" unfolding has_edge_def by auto with edge have "has_edge (star A * A) x G y" by (simp add:tcl_def) then obtain H H' z where AH': "has_edge A z H' y" and G: "G = H * H'" by (auto simp:in_grcomp) from B obtain m' e' where "has_edge H' m' e' n" by (auto simp:G in_grcomp) hence "n ∈ nodes H'" unfolding nodes_def has_edge_def by force with in_smallnodes[OF AH'] show "n ∈ smallnodes A" .. qed next fix x assume "x ∈ smallnodes A" then show "x ∈ smallnodes (tcl A)" by (subst tcl_unfold_right) simp qed lemma finite_nodegraphs: assumes F: "finite F" shows "finite { G::'a scg. nodes G ⊆ F }" (is "finite ?P") proof (rule finite_subset) show "?P ⊆ Graph ` (Pow (F × UNIV × F))" (is "?P ⊆ ?Q") proof fix x assume xP: "x ∈ ?P" obtain S where x[simp]: "x = Graph S" by (cases x) auto from xP show "x ∈ ?Q" apply (simp add:nodes_def) apply (rule imageI) apply (rule PowI) apply force done qed show "finite ?Q" by (auto intro:finite_imageI finite_cartesian_product F finite) qed lemma finite_graphI: fixes A :: "'a acg" assumes fin: "finite (nodes A)" "finite (smallnodes A)" shows "finite_graph A" proof - obtain S where A[simp]: "A = Graph S" by (cases A) auto have "finite S" proof (rule finite_subset) show "S ⊆ nodes A × { G::'a scg. nodes G ⊆ smallnodes A } × nodes A" (is "S ⊆ ?T") proof fix x assume xS: "x ∈ S" obtain a b c where x[simp]: "x = (a, b, c)" by (cases x) auto then have edg: "has_edge A a b c" unfolding has_edge_def using xS by simp hence "a ∈ nodes A" "c ∈ nodes A" unfolding nodes_def has_edge_def by force+ moreover from edg have "nodes b ⊆ smallnodes A" by (rule in_smallnodes) hence "b ∈ { G :: 'a scg. nodes G ⊆ smallnodes A }" by simp ultimately show "x ∈ ?T" by simp qed show "finite ?T" by (intro finite_cartesian_product fin finite_nodegraphs) qed thus ?thesis unfolding finite_graph_def by simp qed lemma smallnodes_allfinite: fixes A :: "'a acg" assumes fin: "finite (smallnodes A)" shows "all_finite A" unfolding all_finite_def proof (intro allI impI) fix n H m assume "has_edge A n H m" then have "nodes H ⊆ smallnodes A" by (rule in_smallnodes) then have "finite (nodes H)" by (rule finite_subset) (rule fin) thus "finite_graph H" by (rule scg_finite) qed lemma finite_tcl: fixes A :: "'a acg" shows "finite_acg (tcl A) <-> finite_acg A" proof assume f: "finite_acg A" from f have g: "finite_graph A" and "all_finite A" unfolding finite_acg_def by auto from g have "finite (nodes A)" by (rule finite_nodes) then have "finite (nodes (tcl A))" unfolding nodes_tcl . moreover from f have "finite (smallnodes A)" by (rule finite_smallnodes) then have fs: "finite (smallnodes (tcl A))" unfolding smallnodes_tcl . ultimately have "finite_graph (tcl A)" by (rule finite_graphI) moreover from fs have "all_finite (tcl A)" by (rule smallnodes_allfinite) ultimately show "finite_acg (tcl A)" unfolding finite_acg_def .. next assume a: "finite_acg (tcl A)" have "A ≤ tcl A" by (rule less_tcl) thus "finite_acg A" using a by (rule finite_acg_subset) qed lemma finite_acg_empty: "finite_acg (Graph {})" unfolding finite_acg_def finite_graph_def all_finite_def has_edge_def by simp lemma finite_acg_ins: assumes fA: "finite_acg (Graph A)" assumes fG: "finite G" shows "finite_acg (Graph (insert (a, Graph G, b) A))" using fA fG unfolding finite_acg_def finite_graph_def all_finite_def has_edge_def by auto lemmas finite_acg_simps = finite_acg_empty finite_acg_ins finite_graph_def subsection {* Contraction and more *} abbreviation "pdesc P == (fst P, prod P, end_node P)" lemma pdesc_acgplus: assumes "has_ipath \<A> p" and "i < j" shows "has_edge (tcl \<A>) (fst (p〈i,j〉)) (prod (p〈i,j〉)) (end_node (p〈i,j〉))" unfolding plus_paths apply (rule exI) apply (insert prems) by (auto intro:sub_path_is_path[of "\<A>" p i j] simp:sub_path_def) lemma combine_fthreads: assumes range: "i < j" "j ≤ k" shows "has_fth p i k m r = (∃n. has_fth p i j m n ∧ has_fth p j k n r)" (is "?L = ?R") proof (intro iffI) assume "?L" then obtain ϑ where "is_fthread ϑ p i k" and [simp]: "ϑ i = m" "ϑ k = r" by (auto simp:has_fth_def) with range have "is_fthread ϑ p i j" and "is_fthread ϑ p j k" by (auto simp:is_fthread_def) hence "has_fth p i j m (ϑ j)" and "has_fth p j k (ϑ j) r" by (auto simp:has_fth_def) thus "?R" by auto next assume "?R" then obtain n ϑ1 ϑ2 where ths: "is_fthread ϑ1 p i j" "is_fthread ϑ2 p j k" and [simp]: "ϑ1 i = m" "ϑ1 j = n" "ϑ2 j = n" "ϑ2 k = r" by (auto simp:has_fth_def) let ?ϑ = "(λi. if i < j then ϑ1 i else ϑ2 i)" have "is_fthread ?ϑ p i k" unfolding is_fthread_def proof fix l assume range: "l ∈ {i..<k}" show "eqlat p ?ϑ l" proof (cases rule:three_cases) assume "Suc l < j" with ths range show ?thesis unfolding is_fthread_def Ball_def by simp next assume "Suc l = j" hence "l < j" "ϑ2 (Suc l) = ϑ1 (Suc l)" by auto with ths range show ?thesis unfolding is_fthread_def Ball_def by simp next assume "j ≤ l" with ths range show ?thesis unfolding is_fthread_def Ball_def by simp qed arith qed moreover have "?ϑ i = m" "?ϑ k = r" using range by auto ultimately show "has_fth p i k m r" by (auto simp:has_fth_def) qed lemma desc_is_fthread: "is_desc_fthread ϑ p i k ==> is_fthread ϑ p i k" unfolding is_desc_fthread_def by simp lemma combine_dfthreads: assumes range: "i < j" "j ≤ k" shows "has_desc_fth p i k m r = (∃n. (has_desc_fth p i j m n ∧ has_fth p j k n r) ∨ (has_fth p i j m n ∧ has_desc_fth p j k n r))" (is "?L = ?R") proof assume "?L" then obtain ϑ where desc: "is_desc_fthread ϑ p i k" and [simp]: "ϑ i = m" "ϑ k = r" by (auto simp:has_desc_fth_def) hence "is_fthread ϑ p i k" by (simp add: desc_is_fthread) with range have fths: "is_fthread ϑ p i j" "is_fthread ϑ p j k" unfolding is_fthread_def by auto hence hfths: "has_fth p i j m (ϑ j)" "has_fth p j k (ϑ j) r" by (auto simp:has_fth_def) from desc obtain l where "i ≤ l" "l < k" and "descat p ϑ l" by (auto simp:is_desc_fthread_def) with fths have "is_desc_fthread ϑ p i j ∨ is_desc_fthread ϑ p j k" unfolding is_desc_fthread_def by (cases "l < j") auto hence "has_desc_fth p i j m (ϑ j) ∨ has_desc_fth p j k (ϑ j) r" by (auto simp:has_desc_fth_def) with hfths show ?R by auto next assume "?R" then obtain n ϑ1 ϑ2 where "(is_desc_fthread ϑ1 p i j ∧ is_fthread ϑ2 p j k) ∨ (is_fthread ϑ1 p i j ∧ is_desc_fthread ϑ2 p j k)" and [simp]: "ϑ1 i = m" "ϑ1 j = n" "ϑ2 j = n" "ϑ2 k = r" by (auto simp:has_fth_def has_desc_fth_def) hence ths2: "is_fthread ϑ1 p i j" "is_fthread ϑ2 p j k" and dths: "is_desc_fthread ϑ1 p i j ∨ is_desc_fthread ϑ2 p j k" by (auto simp:desc_is_fthread) let ?ϑ = "(λi. if i < j then ϑ1 i else ϑ2 i)" have "is_fthread ?ϑ p i k" unfolding is_fthread_def proof fix l assume range: "l ∈ {i..<k}" show "eqlat p ?ϑ l" proof (cases rule:three_cases) assume "Suc l < j" with ths2 range show ?thesis unfolding is_fthread_def Ball_def by simp next assume "Suc l = j" hence "l < j" "ϑ2 (Suc l) = ϑ1 (Suc l)" by auto with ths2 range show ?thesis unfolding is_fthread_def Ball_def by simp next assume "j ≤ l" with ths2 range show ?thesis unfolding is_fthread_def Ball_def by simp qed arith qed moreover from dths have "∃l. i ≤ l ∧ l < k ∧ descat p ?ϑ l" proof assume "is_desc_fthread ϑ1 p i j" then obtain l where range: "i ≤ l" "l < j" and "descat p ϑ1 l" unfolding is_desc_fthread_def Bex_def by auto hence "descat p ?ϑ l" by (cases "Suc l = j", auto) with `j ≤ k` and range show ?thesis by (rule_tac x="l" in exI, auto) next assume "is_desc_fthread ϑ2 p j k" then obtain l where range: "j ≤ l" "l < k" and "descat p ϑ2 l" unfolding is_desc_fthread_def Bex_def by auto with `i < j` have "descat p ?ϑ l" "i ≤ l" by auto with range show ?thesis by (rule_tac x="l" in exI, auto) qed ultimately have "is_desc_fthread ?ϑ p i k" by (simp add: is_desc_fthread_def Bex_def) moreover have "?ϑ i = m" "?ϑ k = r" using range by auto ultimately show "has_desc_fth p i k m r" by (auto simp:has_desc_fth_def) qed lemma fth_single: "has_fth p i (Suc i) m n = eql (snd (p i)) m n" (is "?L = ?R") proof assume "?L" thus "?R" unfolding is_fthread_def Ball_def has_fth_def by auto next let ?ϑ = "λk. if k = i then m else n" assume edge: "?R" hence "is_fthread ?ϑ p i (Suc i) ∧ ?ϑ i = m ∧ ?ϑ (Suc i) = n" unfolding is_fthread_def Ball_def by auto thus "?L" unfolding has_fth_def by auto qed lemma desc_fth_single: "has_desc_fth p i (Suc i) m n = dsc (snd (p i)) m n" (is "?L = ?R") proof assume "?L" thus "?R" unfolding is_desc_fthread_def has_desc_fth_def is_fthread_def Bex_def by (elim exE conjE) (case_tac "k = i", auto) next let ?ϑ = "λk. if k = i then m else n" assume edge: "?R" hence "is_desc_fthread ?ϑ p i (Suc i) ∧ ?ϑ i = m ∧ ?ϑ (Suc i) = n" unfolding is_desc_fthread_def is_fthread_def Ball_def Bex_def by auto thus "?L" unfolding has_desc_fth_def by auto qed lemma mk_eql: "(G \<turnstile> m \<leadsto>e n) ==> eql G m n" by (cases e, auto) lemma eql_scgcomp: "eql (G * H) m r = (∃n. eql G m n ∧ eql H n r)" (is "?L = ?R") proof show "?L ==> ?R" by (auto simp:in_grcomp intro!:mk_eql) assume "?R" then obtain n where l: "eql G m n" and r:"eql H n r" by auto thus ?L by (cases "dsc G m n") (auto simp:in_grcomp mult_sedge_def) qed lemma desc_scgcomp: "dsc (G * H) m r = (∃n. (dsc G m n ∧ eql H n r) ∨ (eq G m n ∧ dsc H n r))" (is "?L = ?R") proof show "?R ==> ?L" by (auto simp:in_grcomp mult_sedge_def) assume "?L" thus ?R by (auto simp:in_grcomp mult_sedge_def) (case_tac "e", auto, case_tac "e'", auto) qed lemma has_fth_unfold: assumes "i < j" shows "has_fth p i j m n = (∃r. has_fth p i (Suc i) m r ∧ has_fth p (Suc i) j r n)" by (rule combine_fthreads) (insert `i < j`, auto) lemma has_dfth_unfold: assumes range: "i < j" shows "has_desc_fth p i j m r = (∃n. (has_desc_fth p i (Suc i) m n ∧ has_fth p (Suc i) j n r) ∨ (has_fth p i (Suc i) m n ∧ has_desc_fth p (Suc i) j n r))" by (rule combine_dfthreads) (insert `i < j`, auto) lemma Lemma7a: "i ≤ j ==> has_fth p i j m n = eql (prod (p〈i,j〉)) m n" proof (induct i arbitrary: m rule:inc_induct) case base show ?case unfolding has_fth_def is_fthread_def sub_path_def by (auto simp:in_grunit one_sedge_def) next case (step i) note IH = `!!m. has_fth p (Suc i) j m n = eql (prod (p〈Suc i,j〉)) m n` have "has_fth p i j m n = (∃r. has_fth p i (Suc i) m r ∧ has_fth p (Suc i) j r n)" by (rule has_fth_unfold[OF `i < j`]) also have "… = (∃r. has_fth p i (Suc i) m r ∧ eql (prod (p〈Suc i,j〉)) r n)" by (simp only:IH) also have "… = (∃r. eql (snd (p i)) m r ∧ eql (prod (p〈Suc i,j〉)) r n)" by (simp only:fth_single) also have "… = eql (snd (p i) * prod (p〈Suc i,j〉)) m n" by (simp only:eql_scgcomp) also have "… = eql (prod (p〈i,j〉)) m n" by (simp only:prod_unfold[OF `i < j`]) finally show ?case . qed lemma Lemma7b: assumes "i ≤ j" shows "has_desc_fth p i j m n = dsc (prod (p〈i,j〉)) m n" using prems proof (induct i arbitrary: m rule:inc_induct) case base show ?case unfolding has_desc_fth_def is_desc_fthread_def sub_path_def by (auto simp:in_grunit one_sedge_def) next case (step i) thus ?case by (simp only:prod_unfold desc_scgcomp desc_fth_single has_dfth_unfold fth_single Lemma7a) auto qed lemma descat_contract: assumes [simp]: "increasing s" shows "descat (contract s p) ϑ i = has_desc_fth p (s i) (s (Suc i)) (ϑ i) (ϑ (Suc i))" by (simp add:Lemma7b increasing_weak contract_def) lemma eqlat_contract: assumes [simp]: "increasing s" shows "eqlat (contract s p) ϑ i = has_fth p (s i) (s (Suc i)) (ϑ i) (ϑ (Suc i))" by (auto simp:Lemma7a increasing_weak contract_def) subsubsection {* Connecting threads *} definition "connect s ϑs = (λk. ϑs (section_of s k) k)" lemma next_in_range: assumes [simp]: "increasing s" assumes a: "k ∈ section s i" shows "(Suc k ∈ section s i) ∨ (Suc k ∈ section s (Suc i))" proof - from a have "k < s (Suc i)" by simp hence "Suc k < s (Suc i) ∨ Suc k = s (Suc i)" by arith thus ?thesis proof assume "Suc k < s (Suc i)" with a have "Suc k ∈ section s i" by simp thus ?thesis .. next assume eq: "Suc k = s (Suc i)" with increasing_strict have "Suc k < s (Suc (Suc i))" by simp with eq have "Suc k ∈ section s (Suc i)" by simp thus ?thesis .. qed qed lemma connect_threads: assumes [simp]: "increasing s" assumes connected: "ϑs i (s (Suc i)) = ϑs (Suc i) (s (Suc i))" assumes fth: "is_fthread (ϑs i) p (s i) (s (Suc i))" shows "is_fthread (connect s ϑs) p (s i) (s (Suc i))" unfolding is_fthread_def proof fix k assume krng: "k ∈ section s i" with fth have eqlat: "eqlat p (ϑs i) k" unfolding is_fthread_def by simp from krng and next_in_range have "(Suc k ∈ section s i) ∨ (Suc k ∈ section s (Suc i))" by simp thus "eqlat p (connect s ϑs) k" proof assume "Suc k ∈ section s i" with krng eqlat show ?thesis unfolding connect_def by (simp only:section_of_known `increasing s`) next assume skrng: "Suc k ∈ section s (Suc i)" with krng have "Suc k = s (Suc i)" by auto with krng skrng eqlat show ?thesis unfolding connect_def by (simp only:section_of_known connected[symmetric] `increasing s`) qed qed lemma connect_dthreads: assumes inc[simp]: "increasing s" assumes connected: "ϑs i (s (Suc i)) = ϑs (Suc i) (s (Suc i))" assumes fth: "is_desc_fthread (ϑs i) p (s i) (s (Suc i))" shows "is_desc_fthread (connect s ϑs) p (s i) (s (Suc i))" unfolding is_desc_fthread_def proof show "is_fthread (connect s ϑs) p (s i) (s (Suc i))" apply (rule connect_threads) apply (insert fth) by (auto simp:connected is_desc_fthread_def) from fth obtain k where dsc: "descat p (ϑs i) k" and krng: "k ∈ section s i" unfolding is_desc_fthread_def by blast from krng and next_in_range have "(Suc k ∈ section s i) ∨ (Suc k ∈ section s (Suc i))" by simp hence "descat p (connect s ϑs) k" proof assume "Suc k ∈ section s i" with krng dsc show ?thesis unfolding connect_def by (simp only:section_of_known inc) next assume skrng: "Suc k ∈ section s (Suc i)" with krng have "Suc k = s (Suc i)" by auto with krng skrng dsc show ?thesis unfolding connect_def by (simp only:section_of_known connected[symmetric] inc) qed with krng show "∃k∈section s i. descat p (connect s ϑs) k" .. qed lemma mk_inf_thread: assumes [simp]: "increasing s" assumes fths: "!!i. i > n ==> is_fthread ϑ p (s i) (s (Suc i))" shows "is_thread (s (Suc n)) ϑ p" unfolding is_thread_def proof (intro allI impI) fix j assume st: "s (Suc n) ≤ j" let ?k = "section_of s j" from in_section_of st have rs: "j ∈ section s ?k" by simp with st have "s (Suc n) < s (Suc ?k)" by simp with increasing_bij have "n < ?k" by simp with rs and fths[of ?k] show "eqlat p ϑ j" by (simp add:is_fthread_def) qed lemma mk_inf_desc_thread: assumes [simp]: "increasing s" assumes fths: "!!i. i > n ==> is_fthread ϑ p (s i) (s (Suc i))" assumes fdths: "∃∞i. is_desc_fthread ϑ p (s i) (s (Suc i))" shows "is_desc_thread ϑ p" unfolding is_desc_thread_def proof (intro exI conjI) from mk_inf_thread[of s n ϑ p] fths show "∀i≥s (Suc n). eqlat p ϑ i" by (fold is_thread_def) simp show "∃∞l. descat p ϑ l" unfolding INF_nat proof fix i let ?k = "section_of s i" from fdths obtain j where "?k < j" "is_desc_fthread ϑ p (s j) (s (Suc j))" unfolding INF_nat by auto then obtain l where "s j ≤ l" and desc: "descat p ϑ l" unfolding is_desc_fthread_def by auto have "i < s (Suc ?k)" by (rule section_of2) simp also have "… ≤ s j" by (rule increasing_weak [OF `increasing s`]) (insert `?k < j`, arith) also note `… ≤ l` finally have "i < l" . with desc show "∃l. i < l ∧ descat p ϑ l" by blast qed qed lemma desc_ex_choice: assumes A: "((∃n.∀i≥n. ∃x. P x i) ∧ (∃∞i. ∃x. Q x i))" and imp: "!!x i. Q x i ==> P x i" shows "∃xs. ((∃n.∀i≥n. P (xs i) i) ∧ (∃∞i. Q (xs i) i))" (is "∃xs. ?Ps xs ∧ ?Qs xs") proof let ?w = "λi. (if (∃x. Q x i) then (SOME x. Q x i) else (SOME x. P x i))" from A obtain n where P: "!!i. n ≤ i ==> ∃x. P x i" by auto { fix i::'a assume "n ≤ i" have "P (?w i) i" proof (cases "∃x. Q x i") case True hence "Q (?w i) i" by (auto intro:someI) with imp show "P (?w i) i" . next case False with P[OF `n ≤ i`] show "P (?w i) i" by (auto intro:someI) qed } hence "?Ps ?w" by (rule_tac x=n in exI) auto moreover from A have "∃∞i. (∃x. Q x i)" .. hence "?Qs ?w" by (rule INF_mono) (auto intro:someI) ultimately show "?Ps ?w ∧ ?Qs ?w" .. qed lemma dthreads_join: assumes [simp]: "increasing s" assumes dthread: "is_desc_thread ϑ (contract s p)" shows "∃ϑs. desc (λi. is_fthread (ϑs i) p (s i) (s (Suc i)) ∧ ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i)) (λi. is_desc_fthread (ϑs i) p (s i) (s (Suc i)) ∧ ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i))" apply (rule desc_ex_choice) apply (insert dthread) apply (simp only:is_desc_thread_def) apply (simp add:eqlat_contract) apply (simp add:descat_contract) apply (simp only:has_fth_def has_desc_fth_def) by (auto simp:is_desc_fthread_def) lemma INF_drop_prefix: "(∃∞i::nat. i > n ∧ P i) = (∃∞i. P i)" apply (auto simp:INF_nat) apply (drule_tac x = "max m n" in spec) apply (elim exE conjE) apply (rule_tac x = "na" in exI) by auto lemma contract_keeps_threads: assumes inc[simp]: "increasing s" shows "(∃ϑ. is_desc_thread ϑ p) <-> (∃ϑ. is_desc_thread ϑ (contract s p))" (is "?A <-> ?B") proof assume "?A" then obtain ϑ n where fr: "∀i≥n. eqlat p ϑ i" and ds: "∃∞i. descat p ϑ i" unfolding is_desc_thread_def by auto let ?cϑ = "λi. ϑ (s i)" have "is_desc_thread ?cϑ (contract s p)" unfolding is_desc_thread_def proof (intro exI conjI) show "∀i≥n. eqlat (contract s p) ?cϑ i" proof (intro allI impI) fix i assume "n ≤ i" also have "i ≤ s i" using increasing_inc by auto finally have "n ≤ s i" . with fr have "is_fthread ϑ p (s i) (s (Suc i))" unfolding is_fthread_def by auto hence "has_fth p (s i) (s (Suc i)) (ϑ (s i)) (ϑ (s (Suc i)))" unfolding has_fth_def by auto with less_imp_le[OF increasing_strict] have "eql (prod (p〈s i,s (Suc i)〉)) (ϑ (s i)) (ϑ (s (Suc i)))" by (simp add:Lemma7a) thus "eqlat (contract s p) ?cϑ i" unfolding contract_def by auto qed show "∃∞i. descat (contract s p) ?cϑ i" unfolding INF_nat proof fix i let ?K = "section_of s (max (s (Suc i)) n)" from `∃∞i. descat p ϑ i` obtain j where "s (Suc ?K) < j" "descat p ϑ j" unfolding INF_nat by blast let ?L = "section_of s j" { fix x assume r: "x ∈ section s ?L" have e1: "max (s (Suc i)) n < s (Suc ?K)" by (rule section_of2) simp note `s (Suc ?K) < j` also have "j < s (Suc ?L)" by (rule section_of2) simp finally have "Suc ?K ≤ ?L" by (simp add:increasing_bij) with increasing_weak have "s (Suc ?K) ≤ s ?L" by simp with e1 r have "max (s (Suc i)) n < x" by simp hence "(s (Suc i)) < x" "n < x" by auto } note range_est = this have "is_desc_fthread ϑ p (s ?L) (s (Suc ?L))" unfolding is_desc_fthread_def is_fthread_def proof show "∀m∈section s ?L. eqlat p ϑ m" proof fix m assume "m∈section s ?L" with range_est(2) have "n < m" . with fr show "eqlat p ϑ m" by simp qed from in_section_of inc less_imp_le[OF `s (Suc ?K) < j`] have "j ∈ section s ?L" . with `descat p ϑ j` show "∃m∈section s ?L. descat p ϑ m" .. qed with less_imp_le[OF increasing_strict] have a: "descat (contract s p) ?cϑ ?L" unfolding contract_def Lemma7b[symmetric] by (auto simp:Lemma7b[symmetric] has_desc_fth_def) have "i < ?L" proof (rule classical) assume "¬ i < ?L" hence "s ?L < s (Suc i)" by (simp add:increasing_bij) also have "… < s ?L" by (rule range_est(1)) (simp add:increasing_strict) finally show ?thesis . qed with a show "∃l. i < l ∧ descat (contract s p) ?cϑ l" by blast qed qed with exI show "?B" . next assume "?B" then obtain ϑ where dthread: "is_desc_thread ϑ (contract s p)" .. with dthreads_join inc obtain ϑs where ths_spec: "desc (λi. is_fthread (ϑs i) p (s i) (s (Suc i)) ∧ ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i)) (λi. is_desc_fthread (ϑs i) p (s i) (s (Suc i)) ∧ ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i))" (is "desc ?alw ?inf") by blast then obtain n where fr: "∀i≥n. ?alw i" by blast hence connected: "!!i. n < i ==> ϑs i (s (Suc i)) = ϑs (Suc i) (s (Suc i))" by auto let ?jϑ = "connect s ϑs" from fr ths_spec have ths_spec2: "!!i. i > n ==> is_fthread (ϑs i) p (s i) (s (Suc i))" "∃∞i. is_desc_fthread (ϑs i) p (s i) (s (Suc i))" by (auto intro:INF_mono) have p1: "!!i. i > n ==> is_fthread ?jϑ p (s i) (s (Suc i))" by (rule connect_threads) (auto simp:connected ths_spec2) from ths_spec2(2) have "∃∞i. n < i ∧ is_desc_fthread (ϑs i) p (s i) (s (Suc i))" unfolding INF_drop_prefix . hence p2: "∃∞i. is_desc_fthread ?jϑ p (s i) (s (Suc i))" apply (rule INF_mono) apply (rule connect_dthreads) by (auto simp:connected) with `increasing s` p1 have "is_desc_thread ?jϑ p" by (rule mk_inf_desc_thread) with exI show "?A" . qed lemma repeated_edge: assumes "!!i. i > n ==> dsc (snd (p i)) k k" shows "is_desc_thread (λi. k) p" proof- have th: "∀ m. ∃na>m. n < na" by arith show ?thesis using prems unfolding is_desc_thread_def apply (auto) apply (rule_tac x="Suc n" in exI, auto) apply (rule INF_mono[where P="λi. n < i"]) apply (simp only:INF_nat) by (auto simp add: th) qed lemma fin_from_inf: assumes "is_thread n ϑ p" assumes "n < i" assumes "i < j" shows "is_fthread ϑ p i j" using prems unfolding is_thread_def is_fthread_def by auto subsection {* Ramsey's Theorem *} definition "set2pair S = (THE (x,y). x < y ∧ S = {x,y})" lemma set2pair_conv: fixes x y :: nat assumes "x < y" shows "set2pair {x, y} = (x, y)" unfolding set2pair_def proof (rule the_equality, simp_all only:split_conv split_paired_all) from `x < y` show "x < y ∧ {x,y}={x,y}" by simp next fix a b assume a: "a < b ∧ {x, y} = {a, b}" hence "{a, b} = {x, y}" by simp_all hence "(a, b) = (x, y) ∨ (a, b) = (y, x)" by (cases "x = y") auto thus "(a, b) = (x, y)" proof assume "(a, b) = (y, x)" with a and `x < y` show ?thesis by auto (* contradiction *) qed qed definition "set2list = inv set" lemma finite_set2list: assumes "finite S" shows "set (set2list S) = S" unfolding set2list_def proof (rule f_inv_f) from `finite S` have "∃l. set l = S" by (rule finite_list) thus "S ∈ range set" unfolding image_def by auto qed corollary RamseyNatpairs: fixes S :: "'a set" and f :: "nat × nat => 'a" assumes "finite S" and inS: "!!x y. x < y ==> f (x, y) ∈ S" obtains T :: "nat set" and s :: "'a" where "infinite T" and "s ∈ S" and "!!x y. [|x ∈ T; y ∈ T; x < y|] ==> f (x, y) = s" proof - from `finite S` have "set (set2list S) = S" by (rule finite_set2list) then obtain l where S: "S = set l" by auto also from set_conv_nth have "… = {l ! i |i. i < length l}" . finally have "S = {l ! i |i. i < length l}" . let ?s = "length l" from inS have index_less: "!!x y. x ≠ y ==> index_of l (f (set2pair {x, y})) < ?s" proof - fix x y :: nat assume neq: "x ≠ y" have "f (set2pair {x, y}) ∈ S" proof (cases "x < y") case True hence "set2pair {x, y} = (x, y)" by (rule set2pair_conv) with True inS show ?thesis by simp next case False with neq have y_less: "y < x" by simp have x:"{x,y} = {y,x}" by auto with y_less have "set2pair {x, y} = (y, x)" by (simp add:set2pair_conv) with y_less inS show ?thesis by simp qed thus "index_of l (f (set2pair {x, y})) < length l" by (simp add: S index_of_length) qed have "∃Y. infinite Y ∧ (∃t. t < ?s ∧ (∀x∈Y. ∀y∈Y. x ≠ y --> index_of l (f (set2pair {x, y})) = t))" by (rule Ramsey2[of "UNIV::nat set", simplified]) (auto simp:index_less) then obtain T i where inf: "infinite T" and i: "i < length l" and d: "!!x y. [|x ∈ T; y∈T; x ≠ y|] ==> index_of l (f (set2pair {x, y})) = i" by auto have "l ! i ∈ S" unfolding S using i by (rule nth_mem) moreover have "!!x y. x ∈ T ==> y∈T ==> x < y ==> f (x, y) = l ! i" proof - fix x y assume "x ∈ T" "y ∈ T" "x < y" with d have "index_of l (f (set2pair {x, y})) = i" by auto with `x < y` have "i = index_of l (f (x, y))" by (simp add:set2pair_conv) with `i < length l` show "f (x, y) = l ! i" by (auto intro:index_of_member[symmetric] iff:index_of_length) qed moreover note inf ultimately show ?thesis using prems by blast qed subsection {* Main Result *} theorem LJA_Theorem4: assumes "finite_acg A" shows "SCT A <-> SCT' A" proof assume "SCT A" show "SCT' A" proof (rule classical) assume "¬ SCT' A" then obtain n G where in_closure: "(tcl A) \<turnstile> n \<leadsto>G n" and idemp: "G * G = G" and no_strict_arc: "∀p. ¬(G \<turnstile> p \<leadsto>\<down> p)" unfolding SCT'_def no_bad_graphs_def by auto from in_closure obtain k where k_pow: "A ^ k \<turnstile> n \<leadsto>G n" and "0 < k" unfolding in_tcl by auto from power_induces_path k_pow obtain loop where loop_props: "has_fpath A loop" "n = fst loop" "n = end_node loop" "G = prod loop" "k = length (snd loop)" . with `0 < k` and path_loop_graph have "has_ipath A (omega loop)" by blast with `SCT A` have thread: "∃ϑ. is_desc_thread ϑ (omega loop)" by (auto simp:SCT_def) let ?s = "λi. k * i" let ?cp = "λi::nat. (n, G)" from loop_props have "fst loop = end_node loop" by auto with `0 < k` `k = length (snd loop)` have "!!i. (omega loop)〈?s i,?s (Suc i)〉 = loop" by (rule sub_path_loop) with `n = fst loop` `G = prod loop` `k = length (snd loop)` have a: "contract ?s (omega loop) = ?cp" unfolding contract_def by (simp add:path_loop_def split_def fst_p0) from `0 < k` have "increasing ?s" by (auto simp:increasing_def) with thread have "∃ϑ. is_desc_thread ϑ ?cp" unfolding a[symmetric] by (unfold contract_keeps_threads[symmetric]) then obtain ϑ where desc: "is_desc_thread ϑ ?cp" by auto then obtain n where thr: "is_thread n ϑ ?cp" unfolding is_desc_thread_def is_thread_def by auto have "finite (range ϑ)" proof (rule finite_range_ignore_prefix) from `finite_acg A` have "finite_acg (tcl A)" by (simp add:finite_tcl) with in_closure have "finite_graph G" unfolding finite_acg_def all_finite_def by blast thus "finite (nodes G)" by (rule finite_nodes) from thread_image_nodes[OF thr] show "∀i≥n. ϑ i ∈ nodes G" by simp qed with finite_range obtain p where inf_visit: "∃∞i. ϑ i = p" by auto then obtain i where "n < i" "ϑ i = p" by (auto simp:INF_nat) from desc have "∃∞i. descat ?cp ϑ i" unfolding is_desc_thread_def by auto then obtain j where "i < j" and "descat ?cp ϑ j" unfolding INF_nat by auto from inf_visit obtain k where "j < k" "ϑ k = p" by (auto simp:INF_nat) from `i < j` `j < k` `n < i` thr fin_from_inf[of n ϑ ?cp] `descat ?cp ϑ j` have "is_desc_fthread ϑ ?cp i k" unfolding is_desc_fthread_def by auto with `ϑ k = p` `ϑ i = p` have dfth: "has_desc_fth ?cp i k p p" unfolding has_desc_fth_def by auto from `i < j` `j < k` have "i < k" by auto hence "prod (?cp〈i, k〉) = G" proof (induct i rule:strict_inc_induct) case base thus ?case by (simp add:sub_path_def) next case (step i) thus ?case by (simp add:sub_path_def upt_rec[of i k] idemp) qed with `i < j` `j < k` dfth Lemma7b[of i k ?cp p p] have "dsc G p p" by auto with no_strict_arc have False by auto thus ?thesis .. qed next assume "SCT' A" show "SCT A" proof (rule classical) assume "¬ SCT A" with SCT_def obtain p where ipath: "has_ipath A p" and no_desc_th: "¬ (∃ϑ. is_desc_thread ϑ p)" by blast from `finite_acg A` have "finite_acg (tcl A)" by (simp add: finite_tcl) hence "finite (dest_graph (tcl A))" (is "finite ?AG") by (simp add: finite_acg_def finite_graph_def) from pdesc_acgplus[OF ipath] have a: "!!x y. x<y ==> pdesc p〈x,y〉 ∈ dest_graph (tcl A)" unfolding has_edge_def . obtain S G where "infinite S" "G ∈ dest_graph (tcl A)" and all_G: "!!x y. [| x ∈ S; y ∈ S; x < y|] ==> pdesc (p〈x,y〉) = G" apply (rule RamseyNatpairs[of ?AG "λ(x,y). pdesc p〈x, y〉"]) apply (rule `finite ?AG`) by (simp only:split_conv, rule a, auto) obtain n H m where G_struct: "G = (n, H, m)" by (cases G) let ?s = "enumerate S" let ?q = "contract ?s p" note all_in_S[simp] = enumerate_in_set[OF `infinite S`] from `infinite S` have inc[simp]: "increasing ?s" unfolding increasing_def by (simp add:enumerate_mono) note increasing_bij[OF this, simp] from ipath_contract inc ipath have "has_ipath (tcl A) ?q" . from all_G G_struct have all_H: "!!i. (snd (?q i)) = H" unfolding contract_def by simp have loop: "(tcl A) \<turnstile> n \<leadsto>H n" and idemp: "H * H = H" proof - let ?i = "?s 0" and ?j = "?s (Suc 0)" and ?k = "?s (Suc (Suc 0))" have "pdesc (p〈?i,?j〉) = G" and "pdesc (p〈?j,?k〉) = G" and "pdesc (p〈?i,?k〉) = G" using all_G by auto with G_struct have "m = end_node (p〈?i,?j〉)" "n = fst (p〈?j,?k〉)" and Hs: "prod (p〈?i,?j〉) = H" "prod (p〈?j,?k〉) = H" "prod (p〈?i,?k〉) = H" by auto hence "m = n" by simp thus "tcl A \<turnstile> n \<leadsto>H n" using G_struct `G ∈ dest_graph (tcl A)` by (simp add:has_edge_def) from sub_path_prod[of ?i ?j ?k p] show "H * H = H" unfolding Hs by simp qed moreover have "!!k. ¬dsc H k k" proof fix k :: 'a assume "dsc H k k" with all_H repeated_edge have "∃ϑ. is_desc_thread ϑ ?q" by fast with inc have "∃ϑ. is_desc_thread ϑ p" by (subst contract_keeps_threads) with no_desc_th show False .. qed ultimately have False using `SCT' A`[unfolded SCT'_def no_bad_graphs_def] by blast thus ?thesis .. qed qed end
lemma finite_range:
finite (range f) ==> ∃x. ∃∞i. f i = x
lemma finite_range_ignore_prefix:
[| finite A; ∀x≥n. f x ∈ A |] ==> finite (range f)
lemma thread_image_nodes:
is_thread n ϑ p ==> ∀i≥n. ϑ i ∈ nodes (snd (p i))
lemma finite_nodes:
finite_graph G ==> finite (nodes G)
lemma nodes_subgraph:
A ≤ B ==> nodes A ⊆ nodes B
lemma finite_edges:
finite_graph G ==> finite (edges G)
lemma edges_sum:
edges (A + B) = edges A ∪ edges B
lemma nodes_sum:
nodes (A + B) = nodes A ∪ nodes B
lemma finite_acg_subset:
[| A ≤ B; finite_acg B |] ==> finite_acg A
lemma scg_finite:
finite (nodes G) ==> finite_graph G
lemma smallnodes_sum:
smallnodes (A + B) = smallnodes A ∪ smallnodes B
lemma in_smallnodes:
A \<turnstile> x \<leadsto>G y ==> nodes G ⊆ smallnodes A
lemma finite_smallnodes:
finite_acg A ==> finite (smallnodes A)
lemma nodes_tcl:
nodes (tcl A) = nodes A
lemma smallnodes_tcl:
smallnodes (tcl A) = smallnodes A
lemma finite_nodegraphs:
finite F ==> finite {G. nodes G ⊆ F}
lemma finite_graphI:
[| finite (nodes A); finite (smallnodes A) |] ==> finite_graph A
lemma smallnodes_allfinite:
finite (smallnodes A) ==> all_finite A
lemma finite_tcl:
finite_acg (tcl A) = finite_acg A
lemma finite_acg_empty:
finite_acg (Graph {})
lemma finite_acg_ins:
[| finite_acg (Graph A); finite G |]
==> finite_acg (Graph (insert (a, Graph G, b) A))
lemma finite_acg_simps:
finite_acg (Graph {})
[| finite_acg (Graph A); finite G |]
==> finite_acg (Graph (insert (a, Graph G, b) A))
finite_graph G = finite (dest_graph G)
lemma pdesc_acgplus:
[| has_ipath \<A> p; i < j |]
==> tcl \<A> \<turnstile> fst p〈i,j〉 \<leadsto>prod p〈i,j〉 end_node p〈i,j〉
lemma combine_fthreads:
[| i < j; j ≤ k |]
==> has_fth p i k m r = (∃n. has_fth p i j m n ∧ has_fth p j k n r)
lemma desc_is_fthread:
is_desc_fthread ϑ p i k ==> is_fthread ϑ p i k
lemma combine_dfthreads:
[| i < j; j ≤ k |]
==> has_desc_fth p i k m r =
(∃n. has_desc_fth p i j m n ∧ has_fth p j k n r ∨
has_fth p i j m n ∧ has_desc_fth p j k n r)
lemma fth_single:
has_fth p i (Suc i) m n = snd (p i) \<turnstile> m \<leadsto> n
lemma desc_fth_single:
has_desc_fth p i (Suc i) m n = snd (p i) \<turnstile> m \<leadsto>\<down> n
lemma mk_eql:
G \<turnstile> m \<leadsto>e n ==> G \<turnstile> m \<leadsto> n
lemma eql_scgcomp:
G * H \<turnstile> m \<leadsto> r =
(∃n. G \<turnstile> m \<leadsto> n ∧ H \<turnstile> n \<leadsto> r)
lemma desc_scgcomp:
G * H \<turnstile> m \<leadsto>\<down> r =
(∃n. G \<turnstile> m \<leadsto>\<down> n ∧ H \<turnstile> n \<leadsto> r ∨
G \<turnstile> m \<leadsto>\<Down> n ∧
H \<turnstile> n \<leadsto>\<down> r)
lemma has_fth_unfold:
i < j
==> has_fth p i j m n = (∃r. has_fth p i (Suc i) m r ∧ has_fth p (Suc i) j r n)
lemma has_dfth_unfold:
i < j
==> has_desc_fth p i j m r =
(∃n. has_desc_fth p i (Suc i) m n ∧ has_fth p (Suc i) j n r ∨
has_fth p i (Suc i) m n ∧ has_desc_fth p (Suc i) j n r)
lemma Lemma7a:
i ≤ j ==> has_fth p i j m n = prod p〈i,j〉 \<turnstile> m \<leadsto> n
lemma Lemma7b:
i ≤ j
==> has_desc_fth p i j m n = prod p〈i,j〉 \<turnstile> m \<leadsto>\<down> n
lemma descat_contract:
increasing s
==> snd (contract s p i) \<turnstile> ϑ i \<leadsto>\<down> ϑ (Suc i) =
has_desc_fth p (s i) (s (Suc i)) (ϑ i) (ϑ (Suc i))
lemma eqlat_contract:
increasing s
==> snd (contract s p i) \<turnstile> ϑ i \<leadsto> ϑ (Suc i) =
has_fth p (s i) (s (Suc i)) (ϑ i) (ϑ (Suc i))
lemma next_in_range:
[| increasing s; k ∈ section s i |]
==> Suc k ∈ section s i ∨ Suc k ∈ section s (Suc i)
lemma connect_threads:
[| increasing s; ϑs i (s (Suc i)) = ϑs (Suc i) (s (Suc i));
is_fthread (ϑs i) p (s i) (s (Suc i)) |]
==> is_fthread (connect s ϑs) p (s i) (s (Suc i))
lemma connect_dthreads:
[| increasing s; ϑs i (s (Suc i)) = ϑs (Suc i) (s (Suc i));
is_desc_fthread (ϑs i) p (s i) (s (Suc i)) |]
==> is_desc_fthread (connect s ϑs) p (s i) (s (Suc i))
lemma mk_inf_thread:
[| increasing s; !!i. n < i ==> is_fthread ϑ p (s i) (s (Suc i)) |]
==> is_thread (s (Suc n)) ϑ p
lemma mk_inf_desc_thread:
[| increasing s; !!i. n < i ==> is_fthread ϑ p (s i) (s (Suc i));
∃∞i. is_desc_fthread ϑ p (s i) (s (Suc i)) |]
==> is_desc_thread ϑ p
lemma desc_ex_choice:
[| (∃n. ∀i≥n. ∃x. P x i) ∧ (∃∞i. ∃x. Q x i); !!x i. Q x i ==> P x i |]
==> ∃xs. (∃n. ∀i≥n. P (xs i) i) ∧ (∃∞i. Q (xs i) i)
lemma dthreads_join:
[| increasing s; is_desc_thread ϑ (contract s p) |]
==> ∃ϑs. (∃n. ∀i≥n. is_fthread (ϑs i) p (s i) (s (Suc i)) ∧
ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i)) ∧
(∃∞i. is_desc_fthread (ϑs i) p (s i) (s (Suc i)) ∧
ϑs i (s i) = ϑ i ∧ ϑs i (s (Suc i)) = ϑ (Suc i))
lemma INF_drop_prefix:
(∃∞i. n < i ∧ P i) = (∃∞i. P i)
lemma contract_keeps_threads:
increasing s
==> (∃ϑ. is_desc_thread ϑ p) = (∃ϑ. is_desc_thread ϑ (contract s p))
lemma repeated_edge:
(!!i. n < i ==> snd (p i) \<turnstile> k \<leadsto>\<down> k)
==> is_desc_thread (λi. k) p
lemma fin_from_inf:
[| is_thread n ϑ p; n < i; i < j |] ==> is_fthread ϑ p i j
lemma set2pair_conv:
x < y ==> set2pair {x, y} = (x, y)
lemma finite_set2list:
finite S ==> set (set2list S) = S
corollary RamseyNatpairs:
[| finite S; !!x y. x < y ==> f (x, y) ∈ S;
!!T s. [| infinite T; s ∈ S;
!!x y. [| x ∈ T; y ∈ T; x < y |] ==> f (x, y) = s |]
==> thesis |]
==> thesis
theorem LJA_Theorem4:
finite_acg A ==> SCT A = SCT' A