Theory TemporalOrderOnPath

(*  Title:      Schutz_Spacetime/TemporalOrderOnPath.thy
    Authors:    Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot
                University of Edinburgh, 2021          
*)
theory TemporalOrderOnPath
imports Minkowski "HOL-Library.Disjoint_Sets"
begin

text ‹
  In Schutz cite‹pp.~18-30› in "schutz1997", this is ``Chapter 3: Temporal order on a path''.
  All theorems are from Schutz, all lemmas are either parts of the Schutz proofs extracted, or
  additional lemmas which needed to be added, with the exception of the three transitivity lemmas
  leading to Theorem 9, which are given by Schutz as well.
  Much of what we'd like to prove about chains with respect to injectivity, surjectivity,
  bijectivity, is proved in TernaryOrdering.thy›.
  Some more things are proved in interlude sections.
›


section "Preliminary Results for Primitives"
text ‹
  First some proofs that belong in this section but aren't proved in the book or are covered but
  in a different form or off-handed remark.
›

context MinkowskiPrimitive begin

lemma cross_once_notin:
  assumes "Q  𝒫"
      and "R  𝒫"
      and "a  Q"
      and "b  Q"
      and "b  R"
      and "a  b"
      and "Q  R"
  shows "a  R"
using assms paths_cross_once eq_paths by meson

lemma paths_cross_at:
  assumes path_Q: "Q  𝒫" and path_R: "R  𝒫"
      and Q_neq_R: "Q  R"
      and QR_nonempty: "Q  R  {}"
      and x_inQ: "x  Q" and x_inR: "x  R"
  shows "Q  R = {x}"
proof (rule equalityI)
  show "Q  R  {x}"
  proof (rule subsetI, rule ccontr)
    fix y
    assume y_in_QR: "y  Q  R"
       and y_not_in_just_x: "y  {x}"
    then have y_neq_x: "y  x" by simp
    then have "¬ (z. Q  R = {z})"
        by (meson Q_neq_R path_Q path_R x_inQ x_inR y_in_QR cross_once_notin IntD1 IntD2)
    thus False using paths_cross_once by (meson QR_nonempty Q_neq_R path_Q path_R)
  qed
  show "{x}  Q  R" using x_inQ x_inR by simp
qed

lemma events_distinct_paths:
  assumes a_event: "a  "
      and b_event: "b  "
      and a_neq_b: "a  b"
  shows "R𝒫. S𝒫. a  R  b  S  (R  S  (∃!c. R  S = {c}))"
  by (metis events_paths assms paths_cross_once)

end (* context MinkowskiPrimitive *)
context MinkowskiBetweenness begin

lemma assumes "[a;b;c]" shows "f. local_long_ch_by_ord f {a,b,c}"
  using abc_abc_neq[OF assms] unfolding chain_defs
  by (simp add: assms ord_ordered_loc)

lemma between_chain: "[a;b;c]  ch {a,b,c}"
proof -
  assume "[a;b;c]"
  hence "f. local_ordering f betw {a,b,c}"
    by (simp add: abc_abc_neq ord_ordered_loc)
  hence "f. local_long_ch_by_ord f {a,b,c}"
    using [a;b;c] abc_abc_neq local_long_ch_by_ord_def by auto
  thus ?thesis
    by (simp add: chain_defs)
qed

end (* context MinkowskiBetweenness *)


section "3.1 Order on a finite chain"
context MinkowskiBetweenness begin

subsection ‹Theorem 1›
text ‹
  See Minkowski.abc_only_cba›.
  Proving it again here to show it can be done following the prose in Schutz.
›
theorem theorem1 [no_atp]:
  assumes abc: "[a;b;c]"
  shows "[c;b;a]  ¬ [b;c;a]  ¬ [c;a;b]" 
proof -
  (* (i) *)
  have part_i: "[c;b;a]" using abc abc_sym by simp
  (* (ii) *)
  have part_ii: "¬ [b;c;a]"
  proof (rule notI)
    assume "[b;c;a]"
    then have "[a;b;a]" using abc abc_bcd_abd by blast
    thus False using abc_ac_neq by blast
  qed
  (* (iii) *)
  have part_iii: "¬ [c;a;b]"
  proof (rule notI)
    assume "[c;a;b]"
    then have "[c;a;c]" using abc abc_bcd_abd by blast
    thus False using abc_ac_neq by blast
  qed
  thus ?thesis using part_i part_ii part_iii by auto
qed

subsection ‹Theorem 2›
text ‹
  The lemma abc_bcd_acd›, equal to the start of Schutz's proof, is given in Minkowski› in order
  to prove some equivalences.
  We're splitting up Theorem 2 into two named results:
  \begin{itemize}
    \item[order_finite_chain›] there is a betweenness relation for each triple of adjacent events, and
    \item[index_injective›] all events of a chain are distinct.
  \end{itemize}
  We will be following Schutz' proof for both.
  Distinctness of chain events is interpreted as injectivity of the indexing function
  (see index_injective›): we assume that this corresponds to what Schutz means by distinctness
  of elements in a sequence.
›

text ‹
  For the case of two-element chains: the elements are distinct by definition,
  and the statement on termlocal_ordering is void (respectively, termFalse  P for any termP).
  We exclude this case from our proof of order_finite_chain›. Two helper lemmas are provided,
  each capturing one of the proofs by induction in Schutz' writing.
›

lemma thm2_ind1:
  assumes chX: "local_long_ch_by_ord f X"
      and finiteX: "finite X"
    shows "j i. ((i::nat) < j  j < card X - 1)  [f i; f j; f (j + 1)]"
proof (rule allI)+
  let ?P = "λ i j. [f i; f j; f (j+1)]"
  fix i j
  show "(i<j  j<card X -1)  ?P i j"
  proof (induct j)
    case 0 (* this assumption is False, so easy *)
    show ?case by blast
  next
    case (Suc j)
    show ?case
    proof (clarify)
      assume asm: "i<Suc j" "Suc j<card X -1"
      have pj: "?P j (Suc j)"
        using asm(2) chX less_diff_conv local_long_ch_by_ord_def local_ordering_def
        by (metis Suc_eq_plus1)
      have "i<j  i=j" using asm(1)
        by linarith
      thus "?P i (Suc j)"
      proof
        assume "i=j"
        hence "Suc i = Suc j  Suc (Suc j) = Suc (Suc j)"
          by simp
        thus "?P i (Suc j)"
          using pj by auto
      next
        assume "i<j"
        have "j < card X - 1"
          using asm(2) by linarith
        thus "?P i (Suc j)"
          using i<j Suc.hyps asm(1) asm(2) chX finiteX Suc_eq_plus1 abc_bcd_acd pj
          by presburger 
      qed
    qed
  qed
qed

lemma thm2_ind2:
  assumes chX: "local_long_ch_by_ord f X"
      and finiteX: "finite X"
    shows "m l. (0<(l-m)  (l-m) < l  l < card X)  [f (l-m-1); f (l-m); (f l)]"
proof (rule allI)+
  fix l m
  let ?P = "λ k l. [f (k-1); f k; f l]"
  let ?n = "card X"
  let ?k = "(l::nat)-m"
  show "0 < ?k  ?k < l  l < ?n  ?P ?k l"
  proof (induct m)
    case 0 (* yet again, this assumption is False, so easy *)
    show ?case by simp
  next
    case (Suc m)
    show ?case
    proof (clarify)
      assume asm: "0 < l - Suc m" "l - Suc m < l" "l < ?n"
      have "Suc m = 1  Suc m > 1" by linarith
      thus "[f (l - Suc m - 1); f (l - Suc m); f l]" (is ?goal)
      proof
        assume "Suc m = 1"
        show ?goal
        proof -
          have "l - Suc m < card X"
            using asm(2) asm(3) less_trans by blast
          then show ?thesis
            using Suc m = 1 asm finiteX thm2_ind1 chX
            using Suc_eq_plus1 add_diff_inverse_nat diff_Suc_less
                  gr_implies_not_zero less_one plus_1_eq_Suc
            by (smt local_long_ch_by_ord_def ordering_ord_ijk_loc)
        qed
      next
        assume "Suc m > 1"
        show ?goal
          apply (rule_tac a="f l" and c="f(l - Suc m - 1)" in abc_sym)
          apply (rule_tac a="f l" and c="f(l-Suc m)" and d="f(l-Suc m-1)" and b="f(l-m)" in abc_bcd_acd)
        proof -
          have "[f(l-m-1); f(l-m); f l]"
            using Suc.hyps 1 < Suc m asm(1,3) by force
          thus "[f l; f(l - m); f(l - Suc m)]"
            using abc_sym One_nat_def diff_zero minus_nat.simps(2)
            by metis
          have "Suc(l - Suc m - 1) = l - Suc m" "Suc(l - Suc m) = l-m"
            using Suc_pred asm(1) by presburger+
          hence "[f(l - Suc m - 1); f(l - Suc m); f(l - m)]"
            using chX unfolding local_long_ch_by_ord_def local_ordering_def
            by (metis asm(2,3) less_trans_Suc)
          thus "[f(l - m); f(l - Suc m); f(l - Suc m - 1)]"
            using abc_sym by blast
        qed
      qed
    qed
  qed
qed

lemma thm2_ind2b:
  assumes chX: "local_long_ch_by_ord f X"
      and finiteX: "finite X"
      and ordered_nats: "0<k  k<l  l < card X"
    shows "[f (k-1); f k; f l]"
  using thm2_ind2 finiteX chX ordered_nats
  by (metis diff_diff_cancel less_imp_le)


text ‹
  This is Theorem 2 properly speaking, except for the "chain elements are distinct" part
  (which is proved as injectivity of the index later). Follows Schutz fairly well!
  The statement Schutz proves under (i) is given in MinkowskiBetweenness.abc_bcd_acd› instead.
›
theorem (*2*) order_finite_chain:
  assumes chX: "local_long_ch_by_ord f X"
      and finiteX: "finite X"
      and ordered_nats: "0  (i::nat)  i < j  j < l  l < card X"
    shows "[f i; f j; f l]"
proof -
  let ?n = "card X - 1"
  have ord1: "0i  i<j  j<?n"
    using ordered_nats by linarith
  have e2: "[f i; f j; f (j+1)]" using thm2_ind1
    using Suc_eq_plus1 chX finiteX ord1
    by presburger
  have e3: "k. 0<k  k<l  [f (k-1); f k; f l]"
    using thm2_ind2b chX finiteX ordered_nats
    by blast
  have "j<l-1  j=l-1"
    using ordered_nats by linarith
  thus ?thesis
  proof
    assume "j<l-1"
    have  "[f j; f (j+1); f l]"
      using e3 abc_abc_neq ordered_nats
      using j < l - 1 less_diff_conv by auto
    thus ?thesis
      using e2 abc_bcd_abd
      by blast
  next
    assume "j=l-1"
    thus ?thesis using e2
      using ordered_nats by auto
  qed
qed


corollary (*2*) order_finite_chain2:
  assumes chX: "[fX]"
      and finiteX: "finite X"
      and ordered_nats: "0  (i::nat)  i < j  j < l  l < card X"
    shows "[f i; f j; f l]"
proof -
  have "card X > 2" using ordered_nats by (simp add: eval_nat_numeral)
  thus ?thesis using order_finite_chain chain_defs short_ch_card(1) by (metis assms nat_neq_iff)
qed


theorem (*2ii*) index_injective:
  fixes i::nat and j::nat
  assumes chX: "local_long_ch_by_ord f X"
      and finiteX: "finite X"
      and indices: "i<j" "j<card X"
    shows "f i  f j"
proof (cases)
  assume "Suc i < j"
  then have "[f i; f (Suc(i)); f j]"
    using order_finite_chain chX finiteX indices(2) by blast
  then show ?thesis
    using abc_abc_neq by blast
next
  assume "¬Suc i < j"
  hence "Suc i = j"
    using Suc_lessI indices(1) by blast
  show ?thesis
  proof (cases)
    assume "Suc j = card X"
    then have "0<i"
    proof -
      have "card X  3"
        using assms(1) finiteX long_chain_card_geq by blast
      thus ?thesis
        using Suc i = j Suc j = card X by linarith
    qed
    then have "[f 0; f i; f j]"
      using assms order_finite_chain by blast
    thus ?thesis
      using abc_abc_neq by blast
  next
    assume "¬Suc j = card X"
    then have "Suc j < card X"
      using Suc_lessI indices(2) by blast
    then have "[f i; f j; f(Suc j)]"
      using chX finiteX indices(1) order_finite_chain by blast
    thus ?thesis
      using abc_abc_neq by blast
  qed
qed

theorem (*2ii*) index_injective2:
  fixes i::nat and j::nat
  assumes chX: "[fX]"
      and finiteX: "finite X"
      and indices: "i<j" "j<card X"
    shows "f i  f j"
  using assms(1) unfolding ch_by_ord_def
proof (rule disjE)
  assume asm: "short_ch_by_ord f X"
  hence "card X = 2" using short_ch_card(1) by simp
  hence "j=1" "i=0" using indices plus_1_eq_Suc by auto
  thus ?thesis using asm unfolding chain_defs by force
next
  assume "local_long_ch_by_ord f X" thus ?thesis using index_injective assms by presburger
qed

text ‹
  Surjectivity of the index function is easily derived from the definition of local_ordering›,
  so we obtain bijectivity as an easy corollary to the second part of Theorem 2.
›

corollary index_bij_betw:
  assumes chX: "local_long_ch_by_ord f X"
    and finiteX: "finite X"
  shows "bij_betw f {0..<card X} X"
proof (unfold bij_betw_def, (rule conjI))
  show "inj_on f {0..<card X}"
    using index_injective[OF assms] by (metis (mono_tags) atLeastLessThan_iff inj_onI nat_neq_iff)
  {
    fix n assume "n  {0..<card X}"
    then have "f n  X"
      using assms unfolding chain_defs local_ordering_def by auto
  } moreover {
    fix x assume "x  X"
    then have "n  {0..<card X}. f n = x"
      using assms unfolding chain_defs local_ordering_def
      using atLeastLessThan_iff bot_nat_0.extremum by blast
  } ultimately show "f ` {0..<card X} = X" by blast
qed

corollary index_bij_betw2:
  assumes chX: "[fX]"
    and finiteX: "finite X"
  shows "bij_betw f {0..<card X} X"
  using assms(1) unfolding ch_by_ord_def
proof (rule disjE)
  assume "local_long_ch_by_ord f X"
  thus "bij_betw f {0..<card X} X" using index_bij_betw assms by presburger
next
  assume asm: "short_ch_by_ord f X"
  show "bij_betw f {0..<card X} X"
  proof (unfold bij_betw_def, (rule conjI))
    show "inj_on f {0..<card X}"
      using index_injective2[OF assms] by (metis (mono_tags) atLeastLessThan_iff inj_onI nat_neq_iff)
    {
      fix n assume asm2: "n  {0..<card X}"
      have "f n  X"
        using asm asm2 short_ch_card(1) apply (simp add: eval_nat_numeral)
        by (metis One_nat_def less_Suc0 less_antisym short_ch_ord_in)
    } moreover {
      fix x assume asm2: "x  X"
      have "n  {0..<card X}. f n = x"
        using short_ch_card(1) short_ch_by_ord_def asm asm2 atLeast0_lessThan_Suc by (auto simp: eval_nat_numeral)[1]
    } ultimately show "f ` {0..<card X} = X" by blast
  qed
qed

subsection "Additional lemmas about chains"

lemma first_neq_last:
  assumes "[fQ|x..z]"
  shows "xz"
  apply (cases rule: finite_chain_with_cases[OF assms])
  using chain_defs apply (metis Suc_1 card_2_iff diff_Suc_1)
  using index_injective[of f Q 0 "card Q - 1"]
  by (metis card.infinite diff_is_0_eq diff_less gr0I le_trans less_imp_le_nat
    less_numeral_extra(1) numeral_le_one_iff semiring_norm(70))

lemma index_middle_element:
  assumes "[fX|a..b..c]"
  shows "n. 0<n  n<(card X - 1)  f n = b"
proof -
  obtain n where n_def: "n < card X" "f n = b"
    using local_ordering_def assms chain_defs by (metis two_ordered_loc)
  have "0<n  n<(card X - 1)  f n = b"
    using assms chain_defs n_def
    by (metis (no_types, lifting) Suc_pred' gr_implies_not0 less_SucE not_gr_zero)
  thus ?thesis by blast
qed

text ‹
  Another corollary to Theorem 2, without mentioning indices.
›
corollary fin_ch_betw: "[fX|a..b..c]  [a;b;c]"
  using order_finite_chain2 index_middle_element
  using finite_chain_def finite_chain_with_def finite_long_chain_with_def
  by (metis (no_types, lifting) card_gt_0_iff diff_less empty_iff le_eq_less_or_eq less_one)


lemma long_chain_2_imp_3: "[fX|a..c]; card X > 2  b. [fX|a..b..c]"
  using points_in_chain first_neq_last finite_long_chain_with_def
  by (metis card_2_iff' numeral_less_iff semiring_norm(75,78))


lemma finite_chain2_betw: "[fX|a..c]; card X > 2  b. [a;b;c]"
  using fin_ch_betw long_chain_2_imp_3 by metis


lemma finite_long_chain_with_alt [chain_alts]: "[fQ|x..y..z]  [fQ|x..z]  [x;y;z]  yQ"
proof
  { 
    assume "[fQ|x .. z]  [x;y;z]  yQ"
    thus "[fQ|x..y..z]"
      using abc_abc_neq finite_long_chain_with_def by blast
  } {
    assume asm: "[fQ|x..y..z]"
    show "[fQ|x .. z]  [x;y;z]  yQ"
    using asm fin_ch_betw finite_long_chain_with_def by blast
  }
qed


lemma finite_long_chain_with_card: "[fQ|x..y..z]  card Q  3"
  unfolding chain_defs numeral_3_eq_3 by fastforce


lemma finite_long_chain_with_alt2:
  assumes "finite Q" "local_long_ch_by_ord f Q" "f 0 = x" "f (card Q - 1) = z" "[x;y;z]  yQ"
  shows "[fQ|x..y..z]"
  using assms finite_chain_alt finite_chain_with_def finite_long_chain_with_alt by blast


lemma finite_long_chain_with_alt3:
  assumes "finite Q" "local_long_ch_by_ord f Q" "f 0 = x" "f (card Q - 1) = z" "yx  yz  yQ"
  shows "[fQ|x..y..z]"
  using assms finite_chain_alt finite_chain_with_def finite_long_chain_with_def by auto


lemma chain_sym_obtain:
  assumes "[fX|a..b..c]"
  obtains g where "[gX|c..b..a]" and "g=(λn. f (card X - 1 - n))"
  using ordering_sym_loc[of betw X f] abc_sym assms unfolding chain_defs
  using  first_neq_last points_in_long_chain(3)
  by (metis assms diff_self_eq_0 empty_iff finite_long_chain_with_def insert_iff minus_nat.diff_0)

lemma chain_sym:
  assumes "[fX|a..b..c]"
    shows "[λn. f (card X - 1 - n)X|c..b..a]"
  using chain_sym_obtain [where f=f and a=a and b=b and c=c and X=X]
  using assms(1) by blast

lemma chain_sym2:
  assumes "[fX|a..c]"
    shows "[λn. f (card X - 1 - n)X|c..a]"
proof -
  {
    assume asm: "a = f 0" "c = f (card X - 1)"
      and asm_short: "short_ch_by_ord f X"
    hence cardX: "card X = 2"
      using short_ch_card(1) by auto
    hence ac: "f 0 = a" "f 1 = c"
      by (simp add: asm)+
    have "n=1  n=0" if "n<card X" for n
      using cardX that by linarith
    hence fn_eq: "(λn. if n = 0 then f 1 else f 0) = (λn. f (card X - Suc n))" if "n<card X" for n
      by (metis One_nat_def Zero_not_Suc ac(2) asm(2) not_gr_zero old.nat.inject zero_less_diff)
    have "c = f (card X - 1 - 0)" and "a = f (card X - 1 - (card X - 1))"
    and "short_ch_by_ord (λn. f (card X - 1 - n)) X"
      apply (simp add: asm)+
      using short_ch_sym[OF asm_short] fn_eq f 1 = c asm(2) short_ch_by_ord_def by fastforce
  }
  consider "short_ch_by_ord f X"|"b. [fX|a..b..c]"
    using assms long_chain_2_imp_3 finite_chain_with_alt by fastforce
  thus ?thesis
    apply cases
    using a=f 0; c=f (card X-1); short_ch_by_ord f X  short_ch_by_ord (λn. f (card X -1-n)) X
      assms finite_chain_alt finite_chain_with_def apply auto[1]
    using chain_sym finite_long_chain_with_alt by blast
qed

lemma chain_sym_obtain2:
  assumes "[fX|a..c]"
  obtains g where "[gX|c..a]" and "g=(λn. f (card X - 1 - n))"
  using assms chain_sym2 by auto


end (* context MinkowskiBetweenness *)


section "Preliminary Results for Kinematic Triangles and Paths/Betweenness"

text ‹
  Theorem 3 (collinearity)
  First we prove some lemmas that will be very helpful.
›


context MinkowskiPrimitive begin

lemma triangle_permutes [no_atp]:
  assumes " a b c" 
  shows " a c b" " b a c" " b c a" " c a b" " c b a"
  using assms by (auto simp add: kinematic_triangle_def)+

lemma triangle_paths [no_atp]:
  assumes tri_abc: " a b c"
  shows "path_ex a b" "path_ex a c" "path_ex b c"
using tri_abc by (auto simp add: kinematic_triangle_def)+


lemma triangle_paths_unique:
  assumes tri_abc: " a b c"
  shows "∃!ab. path ab a b"      
  using path_unique tri_abc triangle_paths(1) by auto

text ‹
  The definition of the kinematic triangle says that there exist paths that $a$ and $b$ pass through,
  and $a$ and $c$ pass through etc that are not equal. But we can show there is a \emph{unique} $ab$ that $a$
  and $b$ pass through, and assuming there is a path $abc$  that $a, b, c$ pass through, it must be
  unique. Therefore ab = abc› and ac = abc›, but ab ≠ ac›, therefore False›.
  Lemma tri_three_paths› is not in the books but might simplify some path obtaining.
›

lemma triangle_diff_paths:
  assumes tri_abc: " a b c"
  shows "¬ (Q𝒫. a  Q  b  Q  c  Q)"
proof (rule notI)
  assume not_thesis: "Q𝒫. a  Q  b  Q  c  Q"
  (* First show [abc] or similar so I can show the path through abc is unique. *)
  then obtain abc where path_abc: "abc  𝒫  a  abc  b  abc  c  abc" by auto
  have abc_neq: "a  b  a  c  b  c" using tri_abc kinematic_triangle_def by simp
  (* Now extract some information from △ a b c. *)
  have "ab𝒫. ac𝒫. ab  ac  a  ab  b  ab  a  ac  c  ac"
    using tri_abc kinematic_triangle_def by metis
  then obtain ab ac where ab_ac_relate: "ab  𝒫  ac  𝒫  ab  ac  {a,b}  ab  {a,c}  ac"
    by blast
  have "∃!ab𝒫. a  ab  b  ab" using tri_abc triangle_paths_unique by blast
  then have ab_eq_abc: "ab = abc" using path_abc ab_ac_relate by auto
  have "∃!ac𝒫. a  ac  b  ac" using tri_abc triangle_paths_unique by blast
  then have ac_eq_abc: "ac = abc" using path_abc ab_ac_relate eq_paths abc_neq by auto
  have "ab = ac" using ab_eq_abc ac_eq_abc by simp
  thus False using ab_ac_relate by simp
qed

lemma tri_three_paths [elim]:
  assumes tri_abc: " a b c"
  shows "ab bc ca. path ab a b  path bc b c  path ca c a  ab  bc  ab  ca  bc  ca"
using tri_abc triangle_diff_paths triangle_paths(2,3) triangle_paths_unique
by fastforce

lemma triangle_paths_neq:
  assumes tri_abc: " a b c"
      and path_ab: "path ab a b"
      and path_ac: "path ac a c"
  shows "ab  ac"
using assms triangle_diff_paths by blast

end (*context MinkowskiPrimitive*)
context MinkowskiBetweenness begin

lemma abc_ex_path_unique:
  assumes abc: "[a;b;c]"
  shows "∃!Q𝒫. a  Q  b  Q  c  Q"
proof -
  have a_neq_c: "a  c" using abc_ac_neq abc by simp
  have "Q𝒫. a  Q  b  Q  c  Q" using abc_ex_path abc by simp
  then obtain P Q where path_P: "P  𝒫" and abc_inP: "a  P  b  P  c  P"
                    and path_Q: "Q  𝒫" and abc_in_Q: "a  Q  b  Q  c  Q" by auto
  then have "P = Q" using a_neq_c eq_paths by blast
  thus ?thesis using eq_paths a_neq_c using abc_inP path_P by auto
qed

lemma betw_c_in_path:
  assumes abc: "[a;b;c]"
      and path_ab: "path ab a b"
  shows "c  ab"
(* By abc_ex_path, there is a path through a b c. By eq_paths there can be only one, which must be ab. *)
using eq_paths abc_ex_path assms by blast

lemma betw_b_in_path:
  assumes abc: "[a;b;c]"
      and path_ab: "path ac a c"
  shows "b  ac"
using assms abc_ex_path_unique path_unique by blast

lemma betw_a_in_path:
  assumes abc: "[a;b;c]"
      and path_ab: "path bc b c"
  shows "a  bc"
using assms abc_ex_path_unique path_unique by blast

lemma triangle_not_betw_abc:
  assumes tri_abc: " a b c"
  shows "¬ [a;b;c]"
using tri_abc abc_ex_path triangle_diff_paths by blast

lemma triangle_not_betw_acb:
  assumes tri_abc: " a b c"
  shows "¬ [a;c;b]"
by (simp add: tri_abc triangle_not_betw_abc triangle_permutes(1))

lemma triangle_not_betw_bac:
  assumes tri_abc: " a b c"
  shows "¬ [b;a;c]"
by (simp add: tri_abc triangle_not_betw_abc triangle_permutes(2))

lemma triangle_not_betw_any:
  assumes tri_abc: " a b c"
  shows "¬ (d{a,b,c}. e{a,b,c}. f{a,b,c}. [d;e;f])"
  by (metis abc_ex_path abc_abc_neq empty_iff insertE tri_abc triangle_diff_paths)

end (*context MinkowskiBetweenness*)


section "3.2 First collinearity theorem"

theorem (*3*) (in MinkowskiChain) collinearity_alt2:
  assumes tri_abc: " a b c"
      and path_de: "path de d e"
      (* This follows immediately from tri_abc without it as an assumption, but we need ab in scope
         to refer to it in the conclusion. *)
      and path_ab: "path ab a b"
      and bcd: "[b;c;d]"
      and cea: "[c;e;a]"
  shows "fdeab. [a;f;b]"
proof -
  have "fab  de. X ord. [ordX|a..f..b]"
  proof -
    have "path_ex a c" using tri_abc triangle_paths(2) by auto
    then obtain ac where path_ac: "path ac a c" by auto
    have "path_ex b c" using tri_abc triangle_paths(3) by auto
    then obtain bc where path_bc: "path bc b c" by auto
    have ab_neq_ac: "ab  ac" using triangle_paths_neq path_ab path_ac tri_abc by fastforce
    have ab_neq_bc: "ab  bc" using eq_paths ab_neq_ac path_ab path_ac path_bc by blast
    have ac_neq_bc: "ac  bc" using eq_paths ab_neq_bc path_ab path_ac path_bc by blast
    have d_in_bc: "d  bc" using bcd betw_c_in_path path_bc by blast
    have e_in_ac: "e  ac" using betw_b_in_path cea path_ac by blast 
    show ?thesis
      using O6_old [where Q = ab and R = ac and S = bc and T = de and a = a and b = b and c = c] 
            ab_neq_ac ab_neq_bc ac_neq_bc path_ab path_bc path_ac path_de bcd cea d_in_bc e_in_ac
      by auto
  qed
  thus ?thesis using fin_ch_betw by blast
qed


theorem (*3*) (in MinkowskiChain) collinearity_alt:
  assumes tri_abc: " a b c"
      and path_de: "path de d e"
      and bcd: "[b;c;d]"
      and cea: "[c;e;a]"
  shows "ab. path ab a b  (fdeab. [a;f;b])"
proof -
  have ex_path_ab: "path_ex a b"
    using tri_abc triangle_paths_unique by blast
  then obtain ab where path_ab: "path ab a b"
    by blast
  have "fab  de. X g. [gX|a..f..b]"
  proof -
    have "path_ex a c" using tri_abc triangle_paths(2) by auto
    then obtain ac where path_ac: "path ac a c" by auto
    have "path_ex b c" using tri_abc triangle_paths(3) by auto
    then obtain bc where path_bc: "path bc b c" by auto
    have ab_neq_ac: "ab  ac" using triangle_paths_neq path_ab path_ac tri_abc by fastforce
    have ab_neq_bc: "ab  bc" using eq_paths ab_neq_ac path_ab path_ac path_bc by blast
    have ac_neq_bc: "ac  bc" using eq_paths ab_neq_bc path_ab path_ac path_bc by blast
    have d_in_bc: "d  bc" using bcd betw_c_in_path path_bc by blast
    have e_in_ac: "e  ac" using betw_b_in_path cea path_ac by blast 
    show ?thesis
      using O6_old [where Q = ab and R = ac and S = bc and T = de and a = a and b = b and c = c] 
            ab_neq_ac ab_neq_bc ac_neq_bc path_ab path_bc path_ac path_de bcd cea d_in_bc e_in_ac
      by auto
  qed
  thus ?thesis using fin_ch_betw path_ab by fastforce
qed


theorem (*3*) (in MinkowskiChain) collinearity:
  assumes tri_abc: " a b c"
      and path_de: "path de d e"
      and bcd: "[b;c;d]"
      and cea: "[c;e;a]"
    shows "(fde(path_of a b). [a;f;b])"
proof - 
  let ?ab = "path_of a b"
  have path_ab: "path ?ab a b"
    using tri_abc theI' [OF triangle_paths_unique] by blast
  have "f?ab  de. X ord. [ordX|a..f..b]"
  proof -
    have "path_ex a c" using tri_abc triangle_paths(2) by auto
    then obtain ac where path_ac: "path ac a c" by auto
    have "path_ex b c" using tri_abc triangle_paths(3) by auto
    then obtain bc where path_bc: "path bc b c" by auto
    have ab_neq_ac: "?ab  ac" using triangle_paths_neq path_ab path_ac tri_abc by fastforce
    have ab_neq_bc: "?ab  bc" using eq_paths ab_neq_ac path_ab path_ac path_bc by blast
    have ac_neq_bc: "ac  bc" using eq_paths ab_neq_bc path_ab path_ac path_bc by blast
    have d_in_bc: "d  bc" using bcd betw_c_in_path path_bc by blast
    have e_in_ac: "e  ac" using betw_b_in_path cea path_ac by blast 
    show ?thesis
      using O6_old [where Q = ?ab and R = ac and S = bc and T = de and a = a and b = b and c = c] 
            ab_neq_ac ab_neq_bc ac_neq_bc path_ab path_bc path_ac path_de bcd cea d_in_bc e_in_ac
            IntI Int_commute
      by (metis (no_types, lifting))
  qed
  thus ?thesis using fin_ch_betw by blast
qed


section "Additional results for Paths and Unreachables"

context MinkowskiPrimitive begin

text ‹The degenerate case.›
lemma big_bang:
  assumes no_paths: "𝒫 = {}"
  shows "a.  = {a}"
proof -
  have "a. a  " using nonempty_events by blast
  then obtain a where a_event: "a  " by auto
  have "¬ (b. b  a)"
  proof (rule notI)
    assume "b. b  a"
    then have "Q. Q  𝒫" using events_paths a_event by auto
    thus False using no_paths by simp
  qed
  then have "b. b = a" by simp
  thus ?thesis using a_event by auto
qed

lemma two_events_then_path:
  assumes two_events: "a. b. a  b"
  shows "Q. Q  𝒫"
proof -
  have "(a.   {a})  𝒫  {}" using big_bang by blast
  then have "𝒫  {}" using two_events by blast
  thus ?thesis by blast
qed

lemma paths_are_events: "Q𝒫. aQ. a  "
  by simp

lemma same_empty_unreach:
  "Q  𝒫; a  Q  unreach-on Q from a = {}"
apply (unfold unreachable_subset_def)
by simp

lemma same_path_reachable:
  "Q  𝒫; a  Q; b  Q  a  Q - unreach-on Q from b"
by (simp add: same_empty_unreach)

text ‹
  If we have two paths crossing and $a$ is on the crossing point, and $b$ is on one of the paths,
  then $a$ is in the reachable part of the path $b$ is on.
›

lemma same_path_reachable2:
  "Q  𝒫; R  𝒫; a  Q; a  R; b  Q  a  R - unreach-on R from b"
  unfolding unreachable_subset_def by blast

(* This will never be used without R ∈ 𝒫 but we might as well leave it off as the proof doesn't
   need it. *)
lemma cross_in_reachable:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and b_inR: "b  R"
  shows "b  R - unreach-on R from a"
unfolding unreachable_subset_def using a_inQ b_inQ b_inR path_Q by auto

lemma reachable_path:
  assumes path_Q: "Q  𝒫"
      and b_event: "b  "
      and a_reachable: "a  Q - unreach-on Q from b"
  shows "R𝒫. a  R  b  R"
proof -
  have a_inQ: "a  Q" using a_reachable by simp
  have "Q  𝒫  b    b  Q  (R𝒫. b  R  a  R)"
      using a_reachable unreachable_subset_def by auto
  then have "b  Q  (R𝒫. b  R  a  R)" using path_Q b_event by simp
  thus ?thesis
  proof (rule disjE)
    assume "b  Q"
    thus ?thesis using a_inQ path_Q by auto
  next
    assume "R𝒫. b  R  a  R"
    thus ?thesis using conj_commute by simp
  qed
qed

end (* context MinkowskiPrimitive *)
context MinkowskiBetweenness begin


lemma ord_path_of:
  assumes "[a;b;c]"
  shows "a  path_of b c" "b  path_of a c" "c  path_of a b"
    and "path_of a b = path_of a c" "path_of a b = path_of b c"
proof -
  show "a  path_of b c"
    using betw_a_in_path[of a b c "path_of b c"] path_of_ex abc_ex_path_unique abc_abc_neq assms
    by (smt (z3) betw_a_in_path the1_equality)
  show "b  path_of a c"
    using betw_b_in_path[of a b c "path_of a c"] path_of_ex abc_ex_path_unique abc_abc_neq assms
    by (smt (z3) betw_b_in_path the1_equality)
  show "c  path_of a b"
    using betw_c_in_path[of a b c "path_of a b"] path_of_ex abc_ex_path_unique abc_abc_neq assms
    by (smt (z3) betw_c_in_path the1_equality)
  show "path_of a b = path_of a c"
    by (metis (mono_tags) abc_ac_neq assms betw_b_in_path betw_c_in_path ends_notin_segment seg_betw)
  show "path_of a b = path_of b c"
    by (metis (mono_tags) assms betw_a_in_path betw_c_in_path ends_notin_segment seg_betw)
qed

text ‹
  Schutz defines chains as subsets of paths. The result below proves that even though we do not
  include this fact in our definition, it still holds, at least for finite chains.
›
text ‹
  Notice that this whole proof would be unnecessary if including path-belongingness in the
  definition, as Schutz does. This would also keep path-belongingness independent of axiom O1 and O4,
  thus enabling an independent statement of axiom O6, which perhaps we now lose. In exchange,
  our definition is slightly weaker (for card X ≥ 3› and infinite X›).
›

lemma obtain_index_fin_chain:
  assumes "[fX]" "xX" "finite X"
  obtains i where "f i = x" "i<card X"
proof -
  have "i<card X. f i = x"
    using assms(1) unfolding ch_by_ord_def
  proof (rule disjE)
    assume asm: "short_ch_by_ord f X"
    hence "card X = 2"
      using short_ch_card(1) by auto
    thus "i<card X. f i = x"
      using asm assms(2) unfolding chain_defs by force
  next
    assume asm: "local_long_ch_by_ord f X"
    thus "i<card X. f i = x"
      using asm assms(2,3) unfolding chain_defs local_ordering_def by blast
  qed
  thus ?thesis using that by blast
qed

lemma obtain_index_inf_chain:
  assumes "[fX]" "xX" "infinite X"
  obtains i where "f i = x"
  using assms unfolding chain_defs local_ordering_def by blast


lemma fin_chain_on_path2:
  assumes "[fX]" "finite X"
  shows "P𝒫. XP"
  using assms(1) unfolding ch_by_ord_def
proof (rule disjE)
  assume "short_ch_by_ord f X"
  thus ?thesis
    using short_ch_by_ord_def by auto
next
  assume asm: "local_long_ch_by_ord f X"
  have "[f 0;f 1;f 2]"
    using order_finite_chain asm assms(2) local_long_ch_by_ord_def by auto
  then obtain P where "P𝒫" "{f 0,f 1,f 2}  P"
    by (meson abc_ex_path empty_subsetI insert_subset)
  then have "path P (f 0) (f 1)"
    using [f 0;f 1;f 2] by (simp add: abc_abc_neq)
  { 
    fix x assume "xX"
    then obtain i where i: "f i = x" "i<card X"
      using obtain_index_fin_chain assms by blast
    consider "i=0i=1"|"i>1" by linarith
    hence "xP"
    proof (cases)
      case 1 thus ?thesis
      using i(1) {f 0, f 1, f 2}  P by auto
    next
      case 2
      hence "[f 0;f 1;f i]"
        using assms i(2) order_finite_chain2 by auto
      hence "{f 0,f 1,f i}P"
        using path P (f 0) (f 1) betw_c_in_path by blast
      thus ?thesis by (simp add: i(1))
    qed
  }
  thus ?thesis
    using P𝒫 by auto
qed


lemma fin_chain_on_path:
  assumes "[fX]" "finite X"
  shows "∃!P𝒫. XP"
proof -
  obtain P where P: "P𝒫" "XP"
    using fin_chain_on_path2[OF assms] by auto
  obtain a b where ab: "aX" "bX" "ab"
    using assms(1) unfolding chain_defs by (metis assms(2) insertCI three_in_set3)
  thus ?thesis using P ab by (meson eq_paths in_mono)
qed


lemma fin_chain_on_path3:
  assumes "[fX]" "finite X" "aX" "bX" "ab"
  shows "X  path_of a b"
proof -
  let ?ab = "path_of a b"
  obtain P where P: "P𝒫" "XP" using fin_chain_on_path2[OF assms(1,2)] by auto
  have "path P a b" using P assms(3-5) by auto
  then have "path ?ab a b" using path_of_ex by blast
  hence "?ab = P" using eq_paths path P a b by auto
  thus "X  path_of a b" using P by simp
qed


end (* context MinkowskiBetweenness *)
context MinkowskiUnreachable begin

text ‹
  First some basic facts about the primitive notions, which seem to belong here.
  I don't think any/all of these are explicitly proved in Schutz.
›

lemma no_empty_paths [simp]:
  assumes "Q𝒫"
  shows "Q{}"
  (*using assms nonempty_events two_in_unreach unreachable_subset_def by blast*)
proof -
  obtain a where "a" using nonempty_events by blast
  have "aQ  aQ" by auto
  thus ?thesis
  proof
    assume "aQ"
    thus ?thesis by blast
  next
    assume "aQ"
    then obtain b where "bunreach-on Q from a"
      using two_in_unreach a   assms
      by blast
    thus ?thesis
      using unreachable_subset_def by auto
  qed
qed

lemma events_ex_path:
  assumes ge1_path: "𝒫  {}"
  shows "x. Q𝒫. x  Q"
  (*using ex_in_conv no_empty_paths in_path_event assms events_paths
  by metis*)
proof
  fix x
  assume x_event: "x  "
  have "Q. Q  𝒫" using ge1_path using ex_in_conv by blast
  then obtain Q where path_Q: "Q  𝒫" by auto
  then have "y. y  Q" using no_empty_paths by blast
  then obtain y where y_inQ: "y  Q" by auto
  then have y_event: "y  " using in_path_event path_Q by simp
  have "P𝒫. x  P"
  proof cases
    assume "x = y"
    thus ?thesis using y_inQ path_Q by auto
  next
    assume "x  y"
    thus ?thesis using events_paths x_event y_event by auto
  qed
  thus "Q𝒫. x  Q" by simp
qed

lemma unreach_ge2_then_ge2:
  assumes "xunreach-on Q from b. yunreach-on Q from b. x  y"
  shows "xQ. yQ. x  y"
using assms unreachable_subset_def by auto


text ‹
  This lemma just proves that the chain obtained to bound the unreachable set of a path
  is indeed on that path. Extends I6; requires Theorem 2; used in Theorem 13.
  Seems to be assumed in Schutz' chain notation in I6.
›

lemma chain_on_path_I6:
  assumes path_Q: "Q𝒫"
      and event_b: "bQ" "b"
      and unreach: "Qx  unreach-on Q from b" "Qz  unreach-on Q from b" "Qx  Qz"
      and X_def: "[fX|Qx..Qz]"
            "(i{1 .. card X - 1}. (f i)  unreach-on Q from b  (Qy. [(f(i-1)); Qy; f i]  Qy  unreach-on Q from b))"
  shows "XQ"
proof -
  have 1: "path Q Qx Qz" using unreachable_subset_def unreach path_Q by simp
  then have 2: "Q = path_of Qx Qz" using path_of_ex[of Qx Qz] by (meson eq_paths)
  have "Xpath_of Qx Qz"
  proof (rule fin_chain_on_path3[of f])
    from unreach(3) show "Qx  Qz" by simp
    from X_def chain_defs show "[fX]" "finite X" by metis+
    from assms(7) points_in_chain show "Qx  X" "Qz  X" by auto
  qed
  thus ?thesis using 2 by simp
qed

end (* context MinkowskiUnreachable *)


section "Results about Paths as Sets"

text ‹
  Note several of the following don't need MinkowskiPrimitive, they are just Set lemmas;
  nevertheless I'm naming them and writing them this way for clarity.
›

context MinkowskiPrimitive begin

lemma distinct_paths:
  assumes "Q  𝒫"
      and "R  𝒫"
      and "d  Q"
      and "d  R"
  shows "R  Q"
using assms by auto

lemma distinct_paths2:
  assumes "Q  𝒫"
      and "R  𝒫"
      and "d. d  Q  d  R"
  shows "R  Q"
using assms by auto

lemma external_events_neq:
  "Q  𝒫; a  Q; b  ; b  Q  a  b"
by auto

lemma notin_cross_events_neq:
  "Q  𝒫; R  𝒫; Q  R; a  Q; b  R; a  RQ  a  b"
by blast

lemma nocross_events_neq:
  "Q  𝒫; R  𝒫; a  Q; b  R; RQ = {}  a  b"
by auto

text ‹
  Given a nonempty path $Q$, and an external point $d$, we can find another path
  $R$ passing through $d$ (by I2 aka events_paths›). This path is distinct
  from $Q$, as it passes through a point external to it.
›
lemma external_path:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and d_notinQ: "d  Q"
      and d_event: "d  "
  shows "R𝒫. d  R"
proof -
  have a_neq_d: "a  d" using a_inQ d_notinQ by auto
  thus "R𝒫. d  R" using events_paths by (meson a_inQ d_event in_path_event path_Q)
qed

lemma distinct_path:
  assumes "Q  𝒫"
      and "a  Q"
      and "d  Q"
      and "d  "
  shows "R𝒫. R  Q"
using assms external_path by metis

lemma external_distinct_path:
  assumes "Q  𝒫"
      and "a  Q"
      and "d  Q"
      and "d  "
  shows "R𝒫. R  Q  d  R"
  using assms external_path by fastforce

end (* context MinkowskiPrimitive *)


section "3.3 Boundedness of the unreachable set"

subsection ‹Theorem 4 (boundedness of the unreachable set)›
text ‹
  The same assumptions as I7, different conclusion.
  This doesn't just give us boundedness, it gives us another event outside of the unreachable
  set, as long as we have one already.
  I7 conclusion:  ∃g X Qn. [g↝X|Qx..Qy..Qn] ∧ Qn ∈ Q - unreach-on Q from b›

theorem (*4*) (in MinkowskiUnreachable) unreachable_set_bounded:
  assumes path_Q: "Q  𝒫"
      and b_nin_Q: "b  Q"
      and b_event: "b  "
      and Qx_reachable: "Qx  Q - unreach-on Q from b"
      and Qy_unreachable: "Qy  unreach-on Q from b"
  shows "QzQ - unreach-on Q from b. [Qx;Qy;Qz]  Qx  Qz"
  using assms I7_old finite_long_chain_with_def fin_ch_betw
  by (metis first_neq_last)

subsection ‹Theorem 5 (first existence theorem)›
text ‹
  The lemma below is used in the contradiction in external_event›,
  which is the essential part to Theorem 5(i).
›
lemma (in MinkowskiUnreachable) only_one_path:
  assumes path_Q: "Q  𝒫"
      and all_inQ: "a. a  Q"
      and path_R: "R  𝒫"
  shows "R = Q"
proof (rule ccontr)
  assume "¬ R = Q"
  then have R_neq_Q: "R  Q" by simp
  have " = Q"
    by (simp add: all_inQ antisym path_Q path_sub_events subsetI)
  hence "RQ"
    using R_neq_Q path_R path_sub_events by auto
  obtain c where "cR" "cQ"
    using R  Q by blast
  then obtain a b where "path R a b"
    using  = Q path_R two_in_unreach unreach_ge2_then_ge2 by blast
  have "aQ" "bQ"
    using  = Q path R a b in_path_event by blast+
  thus False using eq_paths
    using R_neq_Q path R a b path_Q by blast
qed

context MinkowskiSpacetime begin

text ‹Unfortunately, we cannot assume that a path exists without the axiom of dimension.›
lemma external_event:
  assumes path_Q: "Q  𝒫"
  shows "d. d  Q"
proof (rule ccontr)
  assume "¬ (d. d  Q)"
  then have all_inQ: "d. d  Q" by simp
  then have only_one_path: "P𝒫. P = Q" by (simp add: only_one_path path_Q) 
  thus False using ex_3SPRAY three_SPRAY_ge4 four_paths by auto
qed

text ‹
  Now we can prove the first part of the theorem's conjunction.
  This follows pretty much exactly the same pattern as the book, except it relies on more
  intermediate lemmas.
›
theorem (*5i*) ge2_events:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
  shows "bQ. b  a"
proof -
  have d_notinQ: "d. d  Q" using path_Q external_event by blast 
  then obtain d where "d  " and "d  Q" by auto
  thus ?thesis using two_in_unreach [where Q = Q and b = d] path_Q unreach_ge2_then_ge2 by metis
qed

text ‹
  Simple corollary which is easier to use when we don't have one event on a path yet.
  Anything which uses this implicitly used no_empty_paths› on top of ge2_events›.
›
lemma ge2_events_lax:
  assumes path_Q: "Q  𝒫"
  shows "aQ. bQ. a  b"
proof -
  have "a. a  Q" using path_Q no_empty_paths by (meson ex_in_conv in_path_event)
  thus ?thesis using path_Q ge2_events by blast
qed

lemma ex_crossing_path:
  assumes path_Q: "Q  𝒫"
  shows "R𝒫. R  Q  (c. c  R  c  Q)"
proof -
  obtain a where a_inQ: "a  Q" using ge2_events_lax path_Q by blast
  obtain d where d_event: "d  "
             and d_notinQ: "d  Q" using external_event path_Q by auto
  then have "a  d" using a_inQ by auto
  then have ex_through_d: "R𝒫. S𝒫. a  R  d  S  R  S  {}"
      using events_paths [where a = a and b = d]
            path_Q a_inQ in_path_event d_event by simp
  then obtain R S where path_R: "R  𝒫"
                    and path_S: "S  𝒫"
                    and a_inR: "a  R"
                    and d_inS: "d  S"
                    and R_crosses_S: "R  S  {}" by auto
  have S_neq_Q: "S  Q" using d_notinQ d_inS by auto
  show ?thesis
  proof cases
    assume "R = Q"
    then have "Q  S  {}" using R_crosses_S by simp
    thus ?thesis using S_neq_Q path_S by blast
  next
    assume "R  Q"
    thus ?thesis using a_inQ a_inR path_R by blast
  qed
qed

text ‹
  If we have two paths $Q$ and $R$ with $a$ on $Q$ and $b$ at the intersection of $Q$ and $R$, then by
  two_in_unreach› (I5) and Theorem 4 (boundedness of the unreachable set), there is an unreachable
  set from $a$ on one side of $b$ on $R$, and on the other side of that there is an event which is
  reachable from $a$ by some path, which is the path we want.
›

lemma path_past_unreach:
  assumes path_Q: "Q  𝒫"
      and path_R: "R  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and b_inR: "b  R"
      and Q_neq_R: "Q  R"
      and a_neq_b: "a  b"
  shows "S𝒫. S  Q  a  S  (c. c  S  c  R)"
proof -
  obtain d where d_event: "d  "
             and d_notinR: "d  R" using external_event path_R by blast
  have b_reachable: "b  R - unreach-on R from a" using cross_in_reachable path_R a_inQ b_inQ b_inR path_Q by simp
  have a_notinR: "a  R" using cross_once_notin
                               Q_neq_R a_inQ a_neq_b b_inQ b_inR path_Q path_R by blast
  then obtain u where "u  unreach-on R from a"
      using two_in_unreach a_inQ in_path_event path_Q path_R by blast
  then obtain c where c_reachable: "c  R - unreach-on R from a"
                  and c_neq_b: "b  c" using unreachable_set_bounded
                                        [where Q = R and Qx = b and b = a and Qy = u]
                                        path_R d_event d_notinR
      using a_inQ a_notinR b_reachable in_path_event path_Q by blast
  then obtain S where S_facts: "S  𝒫  a  S  (c  S  c  R)" using reachable_path
      by (metis Diff_iff a_inQ in_path_event path_Q path_R)
  then have "S  Q" using Q_neq_R b_inQ b_inR c_neq_b eq_paths path_R by blast
  thus ?thesis using S_facts by auto
qed

theorem (*5ii*) ex_crossing_at:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
  shows "ac𝒫. ac  Q  (c. c  Q  a  ac  c  ac)"
proof -
  obtain b where b_inQ: "b  Q"
             and a_neq_b: "a  b" using a_inQ ge2_events path_Q by blast
  have "R𝒫. R  Q  (e. e  R  e  Q)" by (simp add: ex_crossing_path path_Q)
  then obtain R e where path_R: "R  𝒫"
                    and R_neq_Q: "R  Q"
                    and e_inR: "e  R"
                    and e_inQ: "e  Q" by auto
  thus ?thesis
  proof cases
    assume e_eq_a: "e = a"
    then have "c. c  unreach-on R from b" using R_neq_Q a_inQ a_neq_b b_inQ e_inR path_Q path_R
                                    two_in_unreach path_unique in_path_event by metis
    thus ?thesis using R_neq_Q e_eq_a e_inR path_Q path_R
                       eq_paths ge2_events_lax by metis
  next
    assume e_neq_a: "e  a"
    (* We know the whole of R isn't unreachable from a because e is on R and both a and e are on Q.
       We also know there is some point after e, and after the unreachable set, which is reachable
       from a (because there are at least two events in the unreachable set, and it is bounded). *)
    (* This does follow Schutz, if you unfold the proof for path_past_unreach here, though it's a
       little trickier than Schutz made it seem. *)
    then have "S𝒫. S  Q  a  S  (c. c  S  c  R)"
        using path_past_unreach
              R_neq_Q a_inQ e_inQ e_inR path_Q path_R by auto
    thus ?thesis by (metis R_neq_Q e_inR e_neq_a eq_paths path_Q path_R) 
  qed
qed

(* Alternative formulation using the path function *)
lemma (*5ii_alt*) ex_crossing_at_alt:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
    shows "ac. c. path ac a c  ac  Q  c  Q"
  using ex_crossing_at assms by fastforce

end (* context MinkowskiSpacetime *)


section "3.4 Prolongation"
context MinkowskiSpacetime begin

lemma (in MinkowskiPrimitive) unreach_on_path:
  "a  unreach-on Q from b  a  Q"
using unreachable_subset_def by simp

lemma (in MinkowskiUnreachable) unreach_equiv:
  "Q  𝒫; R  𝒫; a  Q; b  R; a  unreach-on Q from b  b  unreach-on R from a"
  unfolding unreachable_subset_def by auto

theorem (*6i*) prolong_betw:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and ab_neq: "a  b"
  shows "c. [a;b;c]"
proof -
  obtain e ae where e_event: "e  "
                and e_notinQ: "e  Q"
                and path_ae: "path ae a e"
    using ex_crossing_at a_inQ path_Q in_path_event by blast
  have "b  ae" using a_inQ ab_neq b_inQ e_notinQ eq_paths path_Q path_ae by blast
  then obtain f where f_unreachable: "f  unreach-on ae from b"
    using two_in_unreach b_inQ in_path_event path_Q path_ae by blast
  then have b_unreachable: "b  unreach-on Q from f" using unreach_equiv
      by (metis (mono_tags, lifting) CollectD b_inQ path_Q unreachable_subset_def)
  have a_reachable: "a  Q - unreach-on Q from f"
      using same_path_reachable2 [where Q = ae and R = Q and a = a and b = f]
            path_ae a_inQ path_Q f_unreachable unreach_on_path by blast
  thus ?thesis
      using unreachable_set_bounded [where Qy = b and Q = Q and b = f and Qx = a]
            b_unreachable unreachable_subset_def by auto
qed

lemma (in MinkowskiSpacetime) prolong_betw2:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and ab_neq: "a  b"
  shows "cQ. [a;b;c]"
  by (metis assms betw_c_in_path prolong_betw)

lemma (in MinkowskiSpacetime) prolong_betw3:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and ab_neq: "a  b"
  shows "cQ. dQ. [a;b;c]  [a;b;d]  cd"
  by (metis (full_types) abc_abc_neq abc_bcd_abd a_inQ ab_neq b_inQ path_Q prolong_betw2)

lemma finite_path_has_ends:
  assumes "Q  𝒫"
      and "X  Q"
      and "finite X"
      and "card X  3"
    shows "aX. bX. a  b  (cX. a  c  b  c  [a;c;b])"
using assms
proof (induct "card X - 3" arbitrary: X)
  case 0
  then have "card X = 3"
    by linarith
  then obtain a b c where X_eq: "X = {a, b, c}"
    by (metis card_Suc_eq numeral_3_eq_3)
  then have abc_neq: "a  b" "a  c" "b  c"
    by (metis card X = 3 empty_iff insert_iff order_refl three_in_set3)+
  then consider "[a;b;c]" | "[b;c;a]" | "[c;a;b]"
    using some_betw [of Q a b c] "0.prems"(1) "0.prems"(2) X_eq by auto
  thus ?case
  proof (cases)
    assume "[a;b;c]"
    thus ?thesis ― ‹All d not equal to a or c is just d = b, so it immediately follows.›
      using X_eq abc_neq(2) by blast
  next
    assume "[b;c;a]"
    thus ?thesis
      by (simp add: X_eq abc_neq(1))
  next
    assume "[c;a;b]"
    thus ?thesis
      using X_eq abc_neq(3) by blast
  qed
next
  case IH: (Suc n)
  obtain Y x where X_eq: "X = insert x Y" and "x  Y"
    by (meson IH.prems(4) Set.set_insert three_in_set3)
  then have "card Y - 3 = n" "card Y  3"
    using IH.hyps(2) IH.prems(3) X_eq x  Y by auto
  then obtain a b where ab_Y: "a  Y" "b  Y" "a  b"
                    and Y_ends: "cY. (a  c  b  c)  [a;c;b]"
    using IH(1) [of Y] IH.prems(1-3) X_eq by auto
  consider "[a;x;b]" | "[x;b;a]" | "[b;a;x]"
    using some_betw [of Q a x b] ab_Y IH.prems(1,2) X_eq x  Y by auto
  thus ?case
  proof (cases)
    assume "[a;x;b]"
    thus ?thesis
      using Y_ends X_eq ab_Y by auto
  next
    assume "[x;b;a]"
    { fix c
      assume "c  X" "x  c" "a  c"
      then have "[x;c;a]"
        by (smt IH.prems(2) X_eq Y_ends [x;b;a] ab_Y(1) abc_abc_neq abc_bcd_abd abc_only_cba(3) abc_sym Q  𝒫 betw_b_in_path insert_iff some_betw subsetD)
    }
    thus ?thesis
      using X_eq [x;b;a] ab_Y(1) abc_abc_neq insert_iff by force
  next
    assume "[b;a;x]"
    { fix c
      assume "c  X" "b  c" "x  c"
      then have "[b;c;x]"
        by (smt IH.prems(2) X_eq Y_ends [b;a;x] ab_Y(1) abc_abc_neq abc_bcd_acd abc_only_cba(1)
            abc_sym Q  𝒫 betw_a_in_path insert_iff some_betw subsetD)
    }
    thus ?thesis
      using X_eq x  Y ab_Y(2) by fastforce
  qed
qed


lemma obtain_fin_path_ends:
  assumes path_X: "X𝒫"
      and fin_Q: "finite Q"
      and card_Q: "card Q  3"
      and events_Q: "QX"
  obtains a b where "ab" and "aQ" and "bQ" and "cQ. (ac  bc)  [a;c;b]"
proof -
  obtain n where "n0" and "card Q = n+3"
    using card_Q nat_le_iff_add
    by auto
  then obtain a b where "ab" and "aQ" and "bQ" and "cQ. (ac  bc)  [a;c;b]"
    using finite_path_has_ends assms n0
    by metis
  thus ?thesis
    using that by auto
qed


lemma path_card_nil:
  assumes "Q𝒫"
  shows "card Q = 0"
proof (rule ccontr)
  assume "card Q  0"
  obtain n where "n = card Q"
    by auto
  hence "n1"
    using card Q  0 by linarith
  then consider (n1) "n=1" | (n2) "n=2" | (n3) "n3"
    by linarith
  thus False
  proof (cases)
    case n1
    thus ?thesis
      using One_nat_def card_Suc_eq ge2_events_lax singletonD assms(1)
      by (metis n = card Q)
  next
    case n2
    then obtain a b where "ab" and "aQ" and "bQ"
      using ge2_events_lax assms(1) by blast
    then obtain c where "cQ" and "ca" and "cb"
      using prolong_betw2 by (metis abc_abc_neq assms(1))
    hence "card Q  2"
      by (metis a  Q a  b b  Q card_2_iff')
    thus False
      using n = card Q n = 2 by blast
  next
    case n3
    have fin_Q: "finite Q"
    proof -
      have "(0::nat)  1"
        by simp
      then show ?thesis
        by (meson card Q  0 card.infinite)
    qed
    have card_Q: "card Q  3"
      using n = card Q n3 by blast
    have "QQ" by simp 
    then obtain a b where "aQ" and "bQ" and "ab"
        and acb: "cQ. (ca  cb)  [a;c;b]"
      using obtain_fin_path_ends card_Q fin_Q assms(1)
      by metis
    then obtain x where "[a;b;x]" and "xQ"
      using prolong_betw2 assms(1) by blast
    thus False
      by (metis acb abc_abc_neq abc_only_cba(2))
  qed
qed


theorem (*6ii*) infinite_paths:
  assumes "P𝒫"
  shows "infinite P"
proof
  assume fin_P: "finite P"
  have "P{}"
    by (simp add: assms)
  hence "card P  0"
    by (simp add: fin_P)
  moreover have "¬(card P  1)"
    using path_card_nil
    by (simp add: assms)
  ultimately show False
    by simp
qed


end (* contex MinkowskiSpacetime *)


section "3.5 Second collinearity theorem"


text ‹We start with a useful betweenness lemma.›
lemma (in MinkowskiBetweenness) some_betw2:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and c_inQ: "c  Q"
  shows "a = b  a = c  b = c  [a;b;c]  [b;c;a]  [c;a;b]"
  using a_inQ b_inQ c_inQ path_Q some_betw by blast

lemma (in MinkowskiPrimitive) paths_tri:
  assumes path_ab: "path ab a b"
      and path_bc: "path bc b c"
      and path_ca: "path ca c a"
      and a_notin_bc: "a  bc"
  shows " a b c"
proof -
  have abc_events: "a    b    c  "
    using path_ab path_bc path_ca in_path_event by auto
  have abc_neq: "a  b  a  c  b  c"
    using path_ab path_bc path_ca by auto
  have paths_neq: "ab  bc  ab  ca  bc  ca"
    using a_notin_bc cross_once_notin path_ab path_bc path_ca by blast
  show ?thesis
    unfolding kinematic_triangle_def
    using abc_events abc_neq paths_neq path_ab path_bc path_ca
    by auto
qed

lemma (in MinkowskiPrimitive) paths_tri2:
  assumes path_ab: "path ab a b"
      and path_bc: "path bc b c"
      and path_ca: "path ca c a"
      and ab_neq_bc: "ab  bc"
  shows " a b c"
by (meson ab_neq_bc cross_once_notin path_ab path_bc path_ca paths_tri)

text ‹
  Schutz states it more like
   ⟦tri_abc; bcd; cea⟧ ⟹ (path de d e ⟶ ∃f∈de. [a;f;b]∧[d;e;f])›.
  Equivalent up to usage of impI›.
›

theorem (*7*) (in MinkowskiChain) collinearity2:
  assumes tri_abc: " a b c"
      and bcd: "[b;c;d]"
      and cea: "[c;e;a]"
      and path_de: "path de d e"
  shows "f. [a;f;b]  [d;e;f]"
proof -
  obtain ab where path_ab: "path ab a b" using tri_abc triangle_paths_unique by blast
  then obtain f where afb: "[a;f;b]"
                  and f_in_de: "f  de" using collinearity tri_abc path_de path_ab bcd cea by blast
  (* af will be used a few times, so obtain it here. *)
  obtain af where path_af: "path af a f" using abc_abc_neq afb betw_b_in_path path_ab by blast
  have "[d;e;f]"
  proof -
    have def_in_de: "d  de  e  de  f  de" using path_de f_in_de by simp
    then have five_poss:"f = d  f = e  [e;f;d]  [f;d;e]  [d;e;f]"
        using path_de some_betw2 by blast
    have "f = d  f = e  (Q𝒫. a  Q  b  Q  c  Q)"
        by (metis abc_abc_neq afb bcd betw_a_in_path betw_b_in_path cea path_ab)
    then have f_neq_d_e: "f  d  f  e" using tri_abc
        using triangle_diff_paths by simp
    then consider "[e;f;d]" | "[f;d;e]" | "[d;e;f]" using five_poss by linarith
    thus ?thesis
    proof (cases)
      assume efd: "[e;f;d]"
      obtain dc where path_dc: "path dc d c" using abc_abc_neq abc_ex_path bcd by blast
      obtain ce where path_ce: "path ce c e" using abc_abc_neq abc_ex_path cea by blast
      have "dcce"
        using bcd betw_a_in_path betw_c_in_path cea path_ce path_dc tri_abc triangle_diff_paths
        by blast
      hence " d c e"
        using paths_tri2 path_ce path_dc path_de by blast
      then obtain x where x_in_af: "x  af"
                      and dxc: "[d;x;c]"
          using collinearity
                [where a = d and b = c and c = e and d = a and e = f and de = af]
                cea efd path_dc path_af by blast
      then have x_in_dc: "x  dc" using betw_b_in_path path_dc by blast
      then have "x = b" using eq_paths by (metis path_af path_dc afb bcd tri_abc x_in_af
                                            betw_a_in_path betw_c_in_path triangle_diff_paths)
      then have "[d;b;c]" using dxc by simp
      then have "False" using bcd abc_only_cba [where a = b and b = c and c = d] by simp
      thus ?thesis by simp
    next
      assume fde: "[f;d;e]"
      obtain bd where path_bd: "path bd b d" using abc_abc_neq abc_ex_path bcd by blast
      obtain ea where path_ea: "path ea e a" using abc_abc_neq abc_ex_path_unique cea by blast
      obtain fe where path_fe: "path fe f e" using f_in_de f_neq_d_e path_de by blast
      have "feea"
        using tri_abc afb cea path_ea path_fe
        by (metis abc_abc_neq betw_a_in_path betw_c_in_path triangle_paths_neq)
      hence " e a f"
        by (metis path_unique path_af path_ea path_fe paths_tri2)
      then obtain y where y_in_bd: "y  bd"
                      and eya: "[e;y;a]" thm collinearity
          using collinearity
                [where a = e and b = a and c = f and d = b and e = d and de = bd]
                afb fde path_bd path_ea by blast
      then have "y = c" by (metis (mono_tags, lifting)
                                  afb bcd cea path_bd tri_abc
                                  abc_ac_neq betw_b_in_path path_unique triangle_paths(2)
                                    triangle_paths_neq)
      then have "[e;c;a]" using eya by simp
      then have "False" using cea abc_only_cba [where a = c and b = e and c = a] by simp
      thus ?thesis by simp
    next
      assume "[d;e;f]"
      thus ?thesis by assumption
    qed
  qed
  thus ?thesis using afb f_in_de by blast
qed



section "3.6 Order on a path - Theorems 8 and 9"
context MinkowskiSpacetime begin

subsection ‹Theorem 8 (as in Veblen (1911) Theorem 6)›
text ‹
  Note a'b'c'› don't necessarily form a triangle, as there still needs to be paths between them.
›

theorem (*8*) (in MinkowskiChain) tri_betw_no_path:
  assumes tri_abc: " a b c"
      and ab'c: "[a; b'; c]"
      and bc'a: "[b; c'; a]"
      and ca'b: "[c; a'; b]"
  shows "¬ (Q𝒫. a'  Q  b'  Q  c'  Q)"
proof -
  have abc_a'b'c'_neq: "a  a'  a  b'  a  c'
                       b  a'  b  b'  b  c'
                       c  a'  c  b'  c  c'"
      using abc_ac_neq
      by (metis ab'c abc_abc_neq bc'a ca'b tri_abc triangle_not_betw_abc triangle_permutes(4))

  have tri_betw_no_path_single_case: False
    if a'b'c': "[a'; b'; c']" and tri_abc: " a b c"
      and ab'c: "[a; b'; c]" and bc'a: "[b; c'; a]" and ca'b: "[c; a'; b]"
    for a b c a' b' c'
  proof -
    have abc_a'b'c'_neq: "a  a'  a  b'  a  c'
                         b  a'  b  b'  b  c'
                         c  a'  c  b'  c  c'"
      using abc_abc_neq that by (metis triangle_not_betw_abc triangle_permutes(4))
    have c'b'a': "[c'; b'; a']" using abc_sym a'b'c' by simp
    have nopath_a'c'b: "¬ (Q𝒫. a'  Q  c'  Q  b  Q)"
    proof (rule notI)
      assume "Q𝒫. a'  Q  c'  Q  b  Q"
      then obtain Q where path_Q: "Q  𝒫"
                      and a'_inQ: "a'  Q"
                      and c'_inQ: "c'  Q"
                      and b_inQ: "b  Q" by blast
      then have ac_inQ: "a  Q  c  Q" using eq_paths
          by (metis abc_a'b'c'_neq ca'b bc'a betw_a_in_path betw_c_in_path)
      thus False using b_inQ path_Q tri_abc triangle_diff_paths by blast
    qed
    then have tri_a'bc': " a' b c'"
        by (smt bc'a ca'b a'b'c' paths_tri abc_ex_path_unique)
    obtain ab' where path_ab': "path ab' a b'" using ab'c abc_a'b'c'_neq abc_ex_path by blast
    obtain a'b where path_a'b: "path a'b a' b" using tri_a'bc' triangle_paths(1) by blast
    then have "xa'b. [a'; x; b]  [a; b'; x]"
        using collinearity2 [where a = a' and b = b and c = c' and e = b' and d = a and de = ab']
              bc'a betw_b_in_path c'b'a' path_ab' tri_a'bc' by blast
    then obtain x where x_in_a'b: "x  a'b"
                    and a'xb: "[a'; x; b]"
                    and ab'x: "[a; b'; x]" by blast
    (* ab' ∩ a'b = {c} doesn't follow as immediately as in Schutz. *)
    have c_in_ab': "c  ab'" using ab'c betw_c_in_path path_ab' by auto
    have c_in_a'b: "c  a'b" using ca'b betw_a_in_path path_a'b by auto 
    have ab'_a'b_distinct: "ab'  a'b"
        using c_in_a'b path_a'b path_ab' tri_abc triangle_diff_paths by blast
    have "ab'  a'b = {c}"
        using paths_cross_at ab'_a'b_distinct c_in_a'b c_in_ab' path_a'b path_ab' by auto
    then have "x = c" using ab'x path_ab' x_in_a'b betw_c_in_path by auto
    then have "[a'; c; b]" using a'xb by auto
    thus ?thesis using ca'b abc_only_cba by blast
  qed

  show ?thesis
  proof (rule notI)
    assume path_a'b'c': "Q𝒫. a'  Q  b'  Q  c'  Q"
    consider "[a'; b'; c']" | "[b'; c'; a']" | "[c'; a'; b']" using some_betw
        by (smt abc_a'b'c'_neq path_a'b'c' bc'a ca'b ab'c tri_abc
                abc_ex_path cross_once_notin triangle_diff_paths)
    thus False
      apply (cases)
      using tri_betw_no_path_single_case[of a' b' c'] ab'c bc'a ca'b tri_abc apply blast
      using tri_betw_no_path_single_case ab'c bc'a ca'b tri_abc triangle_permutes abc_sym by blast+
  qed
qed

subsection ‹Theorem 9›
text ‹
  We now begin working on the transitivity lemmas needed to prove Theorem 9.
  Multiple lemmas below obtain primed variables (e.g. d'›). These are starred in Schutz (e.g. d*›),
  but that notation is already reserved in Isabelle.
›

lemma unreachable_bounded_path_only:
  assumes d'_def: "d' unreach-on ab from e" "d'ab" "d'e"
      and e_event: "e  "
      and path_ab: "ab  𝒫"
      and e_notin_S: "e  ab"
  shows "d'e. path d'e d' e"
proof (rule ccontr)
  assume "¬(d'e. path d'e d' e)"
  hence "¬(R𝒫. d'R  eR  d'e)"
    by blast
  hence "¬(R𝒫.  eR  d'R)"
    using d'_def(3) by blast
  moreover have "ab𝒫  e  eab"
    by (simp add: e_event e_notin_S path_ab)
  ultimately have "d' unreach-on ab from e"
    unfolding unreachable_subset_def using d'_def(2)
    by blast
  thus False
    using d'_def(1) by auto
qed

lemma unreachable_bounded_path:
  assumes S_neq_ab: "S  ab"
      and a_inS: "a  S"
      and e_inS: "e  S"
      and e_neq_a: "e  a"
      and path_S: "S  𝒫"
      and path_ab: "path ab a b"
      and path_be: "path be b e"
      and no_de: "¬(de. path de d e)"
      and abd:"[a;b;d]"
    obtains d' d'e where "d'ab  path d'e d' e  [b; d; d']"
proof -
  have e_event: "e"
    using e_inS path_S by auto
  have "eab"
    using S_neq_ab a_inS e_inS e_neq_a eq_paths path_S path_ab by auto
  have "ab𝒫  eab"
    using S_neq_ab a_inS e_inS e_neq_a eq_paths path_S path_ab
    by auto
  have "b  ab - unreach-on ab from e"
    using cross_in_reachable path_ab path_be
    by blast
  have "d  unreach-on ab from e"
    using no_de abd path_ab e_event eab
      betw_c_in_path unreachable_bounded_path_only
    by blast
  have  "d' d'e. d'ab  path d'e d' e  [b; d; d']"
  proof -
    obtain d' where "[b; d; d']" "d'ab" "d' unreach-on ab from e" "bd'" "ed'"
      using unreachable_set_bounded b  ab - unreach-on ab from e d  unreach-on ab from e e_event eab path_ab
      by (metis DiffE)
    then obtain d'e where "path d'e d' e"
      using unreachable_bounded_path_only e_event eab path_ab
      by blast
    thus ?thesis
      using [b; d; d'] d'  ab
      by blast
  qed
  thus ?thesis
    using that by blast
qed

text ‹
  This lemma collects the first three paragraphs of Schutz' proof of Theorem 9 - Lemma 1.
  Several case splits need to be considered, but have no further importance outside of this lemma:
  thus we parcel them away from the main proof.›
lemma exist_c'd'_alt:
  assumes abc: "[a;b;c]"
      and abd: "[a;b;d]"
      and dbc: "[d;b;c]" (* the assumption that makes this False for ccontr! *)
      and c_neq_d: "c  d"
      and path_ab: "path ab a b"
      and path_S: "S  𝒫"
      and a_inS: "a  S"
      and e_inS: "e  S"
      and e_neq_a: "e  a"
      and S_neq_ab: "S  ab"
      and path_be: "path be b e"
  shows "c' d'. d'e c'e. c'ab  d'ab
                         [a; b; d']  [c'; b; a]  [c'; b; d']
                         path d'e d' e  path c'e c' e"
proof (cases)
  assume "de. path de d e"
  then obtain de where "path de d e"
    by blast
  hence "[a;b;d]  dab"
    using abd betw_c_in_path path_ab by blast
  thus ?thesis
  proof (cases)
    assume "ce. path ce c e"
    then obtain ce where "path ce c e" by blast
    have "c  ab"
      using abc betw_c_in_path path_ab by blast
    thus ?thesis
      using [a;b;d]  d  ab ce. path ce c e c  ab path de d e abc abc_sym dbc
      by blast
  next
    assume "¬(ce. path ce c e)"
    obtain c' c'e where "c'ab  path c'e c' e  [b; c; c']"
      using unreachable_bounded_path [where ab=ab and e=e and b=b and d=c and a=a and S=S and be=be]
        S_neq_ab ¬(ce. path ce c e) a_inS abc e_inS e_neq_a path_S path_ab path_be
    by (metis (mono_tags, lifting))
    hence "[a; b; c']  [d; b; c']"
      using abc dbc by blast
    hence "[c'; b; a]  [c'; b; d]"
      using theorem1 by blast
    thus ?thesis
      using [a;b;d]  d  ab c'  ab  path c'e c' e  [b; c; c'] path de d e
      by blast
  qed
next
  assume "¬ (de. path de d e)"
  obtain d' d'e where d'_in_ab: "d'  ab"
                   and bdd': "[b; d; d']"
                   and "path d'e d' e"
    using unreachable_bounded_path [where ab=ab and e=e and b=b and d=d and a=a and S=S and be=be]
      S_neq_ab de. path de d e a_inS abd e_inS e_neq_a path_S path_ab path_be
    by (metis (mono_tags, lifting))
  hence "[a; b; d']" using abd by blast
  thus ?thesis
  proof (cases)
    assume "ce. path ce c e"
    then obtain ce where "path ce c e" by blast
    have "c  ab"
      using abc betw_c_in_path path_ab by blast
    thus ?thesis
      using [a; b; d'] d'  ab path ce c e c  ab path d'e d' e abc abc_sym dbc
      by (meson abc_bcd_acd bdd')
  next
    assume "¬(ce. path ce c e)"
    obtain c' c'e where "c'ab  path c'e c' e  [b; c; c']"
      using unreachable_bounded_path [where ab=ab and e=e and b=b and d=c and a=a and S=S and be=be]
        S_neq_ab ¬(ce. path ce c e) a_inS abc e_inS e_neq_a path_S path_ab path_be
    by (metis (mono_tags, lifting))
    hence "[a; b; c']  [d; b; c']"
      using abc dbc by blast
    hence "[c'; b; a]  [c'; b; d]"
      using theorem1 by blast
    thus ?thesis
      using [a; b; d'] c'  ab  path c'e c' e  [b; c; c'] path d'e d' e bdd' d'_in_ab
      by blast
  qed
qed

lemma exist_c'd':
  assumes abc: "[a;b;c]"
      and abd: "[a;b;d]"
      and dbc: "[d;b;c]" (* the assumption that makes this False for ccontr! *)
      and path_S: "path S a e"
      and path_be: "path be b e"
      and S_neq_ab: "S  path_of a b"
    shows "c' d'. [a; b; d']  [c'; b; a]  [c'; b; d'] 
                   path_ex d' e  path_ex c' e"
proof (cases "path_ex d e")
  let ?ab = "path_of a b"
  have "path_ex a b"
    using abc abc_abc_neq abc_ex_path by blast
  hence path_ab: "path ?ab a b" using path_of_ex by simp
  have "cd" using abc_ac_neq dbc by blast
  {
    case True
    then obtain de where "path de d e"
      by blast
    hence "[a;b;d]  d?ab"
      using abd betw_c_in_path path_ab by blast
    thus ?thesis
    proof (cases "path_ex c e")
      case True
      then obtain ce where "path ce c e" by blast
      have "c  ?ab"
        using abc betw_c_in_path path_ab by blast
      thus ?thesis
        using [a;b;d]  d  ?ab ce. path ce c e c  ?ab path de d e abc abc_sym dbc
        by blast
    next
      case False
      obtain c' c'e where "c'?ab  path c'e c' e  [b; c; c']"
        using unreachable_bounded_path [where ab="?ab" and e=e and b=b and d=c and a=a and S=S and be=be]
          S_neq_ab ¬(ce. path ce c e) abc path_S path_ab path_be
      by (metis (mono_tags, lifting))
      hence "[a; b; c']  [d; b; c']"
        using abc dbc by blast
      hence "[c'; b; a]  [c'; b; d]"
        using theorem1 by blast
      thus ?thesis
        using [a;b;d]  d  ?ab c'  ?ab  path c'e c' e  [b; c; c'] path de d e
        by blast
    qed
  } {
    case False
    obtain d' d'e where d'_in_ab: "d'  ?ab"
                     and bdd': "[b; d; d']"
                     and "path d'e d' e"
      using unreachable_bounded_path [where ab="?ab" and e=e and b=b and d=d and a=a and S=S and be=be]
        S_neq_ab ¬path_ex d e abd path_S path_ab path_be
      by (metis (mono_tags, lifting))
    hence "[a; b; d']" using abd by blast
    thus ?thesis
    proof (cases "path_ex c e")
      case True
      then obtain ce where "path ce c e" by blast
      have "c  ?ab"
        using abc betw_c_in_path path_ab by blast
      thus ?thesis
        using [a; b; d'] d'  ?ab path ce c e c  ?ab path d'e d' e abc abc_sym dbc
        by (meson abc_bcd_acd bdd')
    next
      case False
      obtain c' c'e where "c'?ab  path c'e c' e  [b; c; c']"
        using unreachable_bounded_path [where ab="?ab" and e=e and b=b and d=c and a=a and S=S and be=be]
          S_neq_ab ¬(path_ex c e) abc path_S path_ab path_be
      by (metis (mono_tags, lifting))
      hence "[a; b; c']  [d; b; c']"
        using abc dbc by blast
      hence "[c'; b; a]  [c'; b; d]"
        using theorem1 by blast
      thus ?thesis
        using [a; b; d'] c'  ?ab  path c'e c' e  [b; c; c'] path d'e d' e bdd' d'_in_ab
        by blast
    qed
  }
qed


lemma exist_f'_alt:
  assumes path_ab: "path ab a b"
      and path_S: "S  𝒫"
      and a_inS: "a  S"
      and e_inS: "e  S"
      and e_neq_a: "e  a"
      and f_def: "[e; c'; f]" "fc'e"
      and S_neq_ab: "S  ab"
      and c'd'_def: "c'ab  d'ab
             [a; b; d']  [c'; b; a]  [c'; b; d']
             path d'e d' e  path c'e c' e"
    shows "f'. f'b. [e; c'; f']  path f'b f' b"
proof (cases)
  assume "bf. path bf b f"
  thus ?thesis
    using [e; c'; f] by blast
next
  assume "¬(bf. path bf b f)"
  hence "f  unreach-on c'e from b"
    using assms(1-5,7-9) abc_abc_neq betw_events eq_paths unreachable_bounded_path_only
    by metis
  moreover have "c'  c'e - unreach-on c'e from b"
    using c'd'_def cross_in_reachable path_ab by blast
  moreover have "b  bc'e"
    using f  unreach-on c'e from b betw_events c'd'_def same_empty_unreach by auto
  ultimately obtain f' where f'_def: "[c'; f; f']" "f'c'e" "f' unreach-on c'e from b" "c'f'" "bf'"
    using unreachable_set_bounded c'd'_def
    by (metis DiffE)
  hence "[e; c'; f']"
    using [e; c'; f] by blast
  moreover obtain f'b where "path f'b f' b"
    using b    b  c'e c'd'_def f'_def(2,3) unreachable_bounded_path_only
    by blast
  ultimately show ?thesis by blast
qed

lemma exist_f':
  assumes path_ab: "path ab a b"
      and path_S: "path S a e"
      and f_def: "[e; c'; f]"
      and S_neq_ab: "S  ab"
      and c'd'_def: "[a; b; d']" "[c'; b; a]" "[c'; b; d']"
            "path d'e d' e" "path c'e c' e"
    shows "f'. [e; c'; f']  path_ex f' b"
proof (cases)
  assume "path_ex b f"
  thus ?thesis
    using f_def by blast
next
  assume no_path: "¬(path_ex b f)"
  have path_S_2: "S  𝒫" "a  S" "e  S" "e  a"
    using path_S by auto
  have "fc'e"
    using betw_c_in_path f_def c'd'_def(5) by blast
  have "c' ab" "d' ab"
    using betw_a_in_path betw_c_in_path c'd'_def(1,2) path_ab by blast+
  have "f  unreach-on c'e from b"
    using no_path assms(1,4-9) path_S_2 fc'e c'ab d'ab
      abc_abc_neq betw_events eq_paths unreachable_bounded_path_only
    by metis
  moreover have "c'  c'e - unreach-on c'e from b"
    using c'd'_def cross_in_reachable path_ab c'  ab by blast
  moreover have "b  bc'e"
    using f  unreach-on c'e from b betw_events c'd'_def same_empty_unreach by auto
  ultimately obtain f' where f'_def: "[c'; f; f']" "f'c'e" "f' unreach-on c'e from b" "c'f'" "bf'"
    using unreachable_set_bounded c'd'_def
    by (metis DiffE)
  hence "[e; c'; f']"
    using [e; c'; f] by blast
  moreover obtain f'b where "path f'b f' b"
    using b    b  c'e c'd'_def f'_def(2,3) unreachable_bounded_path_only
    by blast
  ultimately show ?thesis by blast
qed


lemma abc_abd_bcdbdc:
  assumes abc: "[a;b;c]"
      and abd: "[a;b;d]"
      and c_neq_d: "c  d"
  shows "[b;c;d]  [b;d;c]"
proof -
  have "¬ [d;b;c]"
  proof (rule notI)
    assume dbc: "[d;b;c]"
    obtain ab where path_ab: "path ab a b"
      using abc_abc_neq abc_ex_path_unique abc by blast
    obtain S where path_S: "S  𝒫" 
               and S_neq_ab: "S  ab"
               and a_inS: "a  S"
      using ex_crossing_at path_ab
      by auto
    (* This is not as immediate as Schutz presents it. *)
    have "eS. e  a  (be𝒫. path be b e)"
    proof -
      have b_notinS: "b  S" using S_neq_ab a_inS path_S path_ab path_unique by blast
      then obtain x y z where x_in_unreach: "x  unreach-on S from b"
                        and y_in_unreach: "y  unreach-on S from b"
                        and x_neq_y: "x  y"
                        and z_in_reach: "z  S - unreach-on S from b"
        using two_in_unreach [where Q = S and b = b]
          in_path_event path_S path_ab a_inS cross_in_reachable
        by blast
      then obtain w where w_in_reach: "w  S - unreach-on S from b"
                      and w_neq_z: "w  z"
          using unreachable_set_bounded [where Q = S and b = b and Qx = z and Qy = x]
                b_notinS in_path_event path_S path_ab by blast
      thus ?thesis by (metis DiffD1 b_notinS in_path_event path_S path_ab reachable_path z_in_reach)
    qed
    then obtain e be where e_inS: "e  S"
                       and e_neq_a: "e  a"
                       and path_be: "path be b e"
      by blast
    have path_ae: "path S a e"
      using a_inS e_inS e_neq_a path_S by auto
    have S_neq_ab_2: "S  path_of a b"
      using S_neq_ab cross_once_notin path_ab path_of_ex by blast

    (* Obtain c' and d' as in Schutz (there called c* and d* ) *)
    have "c' d'.
              c'ab  d'ab
             [a; b; d']  [c'; b; a]  [c'; b; d']
             path_ex d' e  path_ex c' e"
      using exist_c'd' [where a=a and b=b and c=c and d=d and e=e and be=be and S=S]
      using assms(1-2) dbc e_neq_a path_ae path_be S_neq_ab_2
      using abc_sym betw_a_in_path path_ab by blast
    then obtain c' d' d'e c'e
      where c'd'_def: "c'ab  d'ab
             [a; b; d']  [c'; b; a]  [c'; b; d']
             path d'e d' e  path c'e c' e"
      by blast

    (* Now obtain f' (called f* in Schutz) *)
    obtain f where f_def: "fc'e" "[e; c'; f]"
      using c'd'_def prolong_betw2 by blast
    then obtain f' f'b where f'_def: "[e; c'; f']  path f'b f' b"
      using exist_f'
        [where e=e and c'=c' and b=b and f=f and S=S and ab=ab and d'=d' and a=a and c'e=c'e]
      using path_ab path_S a_inS e_inS e_neq_a f_def S_neq_ab c'd'_def
      by blast

    (* Now we follow Schutz, who follows Veblen. *)
    obtain ae where path_ae: "path ae a e" using a_inS e_inS e_neq_a path_S by blast
    have tri_aec: " a e c'"
        by (smt cross_once_notin S_neq_ab a_inS abc abc_abc_neq abc_ex_path
                e_inS e_neq_a path_S path_ab c'd'_def paths_tri)
    (* The second collinearity theorem doesn't explicitly capture the fact that it meets at
       ae, so Schutz misspoke, but maybe that's an issue with the statement of the theorem. *)
    then obtain h where h_in_f'b: "h  f'b"
                    and ahe: "[a;h;e]"
                    and f'bh: "[f'; b; h]"
        using collinearity2 [where a = a and b = e and c = c' and d = f' and e = b and de = f'b]
              f'_def c'd'_def f'_def betw_c_in_path by blast
    have tri_dec: " d' e c'"
        using cross_once_notin S_neq_ab a_inS abc abc_abc_neq abc_ex_path
                e_inS e_neq_a path_S path_ab c'd'_def paths_tri by smt
    then obtain g where g_in_f'b: "g  f'b"
                    and d'ge: "[d'; g; e]"
                    and f'bg: "[f'; b; g]"
        using collinearity2 [where a = d' and b = e and c = c' and d = f' and e = b and de = f'b]
              f'_def c'd'_def betw_c_in_path by blast
    have " e a d'" by (smt betw_c_in_path paths_tri2 S_neq_ab a_inS abc_ac_neq
                           abd e_inS e_neq_a c'd'_def path_S path_ab)
    thus False
      using tri_betw_no_path [where a = e and b = a and c = d' and b' = g and a' = b and c' = h]
        f'_def c'd'_def h_in_f'b g_in_f'b abd d'ge ahe abc_sym
      by blast
  qed
  thus ?thesis
    by (smt abc abc_abc_neq abc_ex_path abc_sym abd c_neq_d cross_once_notin some_betw)
qed


(* Lemma 2-3.6. *)
lemma abc_abd_acdadc:
  assumes abc: "[a;b;c]"
      and abd: "[a;b;d]"
      and c_neq_d: "c  d"
  shows "[a;c;d]  [a;d;c]"
proof -
  have cba: "[c;b;a]" using abc_sym abc by simp
  have dba: "[d;b;a]" using abc_sym abd by simp
  have dcb_over_cba: "[d;c;b]  [c;b;a]  [d;c;a]" by auto
  have cdb_over_dba: "[c;d;b]  [d;b;a]  [c;d;a]" by auto

  have bcdbdc: "[b;c;d]  [b;d;c]" using abc abc_abd_bcdbdc abd c_neq_d by auto
  then have dcb_or_cdb: "[d;c;b]  [c;d;b]" using abc_sym by blast
  then have "[d;c;a]  [c;d;a]" using abc_only_cba dcb_over_cba cdb_over_dba cba dba by blast
  thus ?thesis using abc_sym by auto
qed

(* Lemma 3-3.6. *)
lemma abc_acd_bcd:
  assumes abc: "[a;b;c]"
      and acd: "[a;c;d]"
  shows "[b;c;d]"
proof -
  have path_abc: "Q𝒫. a  Q  b  Q  c  Q" using abc by (simp add: abc_ex_path)
  have path_acd: "Q𝒫. a  Q  c  Q  d  Q" using acd by (simp add: abc_ex_path)
  then have "Q𝒫. b  Q  c  Q  d  Q" using path_abc abc_abc_neq acd cross_once_notin by metis
  then have bcd3: "[b;c;d]  [b;d;c]  [c;b;d]" by (metis abc abc_only_cba(1,2) acd some_betw2)
  show ?thesis
  proof (rule ccontr)
    assume "¬ [b;c;d]"
    then have "[b;d;c]  [c;b;d]" using bcd3 by simp
    thus False
    proof (rule disjE)
      assume "[b;d;c]"
      then have "[c;d;b]" using abc_sym by simp
      then have "[a;c;b]" using acd abc_bcd_abd by blast
      thus False using abc abc_only_cba by blast
    next
      assume cbd: "[c;b;d]"
      have cba: "[c;b;a]" using abc abc_sym by blast
      have a_neq_d: "a  d" using abc_ac_neq acd by auto
      then have "[c;a;d]  [c;d;a]" using abc_abd_acdadc cbd cba by simp
      thus False using abc_only_cba acd by blast
    qed
  qed
qed

text ‹
  A few lemmas that don't seem to be proved by Schutz, but can be proven now, after Lemma 3.
  These sometimes avoid us having to construct a chain explicitly.
›
lemma abd_bcd_abc:
  assumes abd: "[a;b;d]"
      and bcd: "[b;c;d]"
  shows "[a;b;c]"
proof -
  have dcb: "[d;c;b]" using abc_sym bcd by simp
  have dba: "[d;b;a]" using abc_sym abd by simp
  have "[c;b;a]" using abc_acd_bcd dcb dba by blast
  thus ?thesis using abc_sym by simp
qed

lemma abc_acd_abd:
  assumes abc: "[a;b;c]"
      and acd: "[a;c;d]"
    shows "[a;b;d]"
  using abc abc_acd_bcd acd by blast

lemma abd_acd_abcacb:
  assumes abd: "[a;b;d]"
      and acd: "[a;c;d]"
      and bc: "bc"
    shows "[a;b;c]  [a;c;b]"
proof -
  obtain P where P_def: "P𝒫" "aP" "bP" "dP"
    using abd abc_ex_path by blast
  hence "cP"
    using acd abc_abc_neq betw_b_in_path by blast
  have "¬[b;a;c]"
    using abc_only_cba abd acd by blast
  thus ?thesis
    by (metis P_def(1-3) c  P abc_abc_neq abc_sym abd acd bc some_betw)
qed

lemma abe_ade_bcd_ace:
  assumes abe: "[a;b;e]"
      and ade: "[a;d;e]"
      and bcd: "[b;c;d]"
    shows "[a;c;e]"
proof -
  have abdadb: "[a;b;d]  [a;d;b]"
    using abc_ac_neq abd_acd_abcacb abe ade bcd by auto
  thus ?thesis
  proof
    assume "[a;b;d]" thus ?thesis
      by (meson abc_acd_abd abc_sym ade bcd)
  next assume "[a;d;b]" thus ?thesis
      by (meson abc_acd_abd abc_sym abe bcd)
  qed
qed

text ‹Now we start on Theorem 9. Based on Veblen (1904) Lemma 2 p357.›

lemma (in MinkowskiBetweenness) chain3:
  assumes path_Q: "Q  𝒫"
      and a_inQ: "a  Q"
      and b_inQ: "b  Q"
      and c_inQ: "c  Q"
      and abc_neq: "a  b  a  c  b  c"
  shows "ch {a,b,c}"
proof -
  have abc_betw: "[a;b;c]  [a;c;b]  [b;a;c]"
    using assms by (meson in_path_event abc_sym some_betw insert_subset)
  have ch1: "[a;b;c]  ch {a,b,c}"
    using abc_abc_neq ch_by_ord_def ch_def ord_ordered_loc between_chain by auto
  have ch2: "[a;c;b]  ch {a,c,b}"
    using abc_abc_neq ch_by_ord_def ch_def ord_ordered_loc between_chain by auto
  have ch3: "[b;a;c]  ch {b,a,c}"
    using abc_abc_neq ch_by_ord_def ch_def ord_ordered_loc between_chain by auto
  show ?thesis
    using abc_betw ch1 ch2 ch3 by (metis insert_commute)
qed

lemma overlap_chain: "[a;b;c]; [b;c;d]  ch {a,b,c,d}"
proof -
  assume "[a;b;c]" and "[b;c;d]"
  have "f. local_ordering f betw {a,b,c,d}"
  proof -
    have f1: "[a;b;d]"
      using [a;b;c] [b;c;d] by blast
    have "[a;c;d]"
      using [a;b;c] [b;c;d] abc_bcd_acd by blast
    then show ?thesis
      using f1 by (metis (no_types) [a;b;c] [b;c;d] abc_abc_neq overlap_ordering_loc)
  qed
  hence "f. local_long_ch_by_ord f {a,b,c,d}"
    apply (simp add: chain_defs eval_nat_numeral)
    using [a;b;c] abc_abc_neq
    by (smt (z3) [b;c;d] card.empty card_insert_disjoint card_insert_le finite.emptyI
      finite.insertI insertE insert_absorb insert_not_empty)
  thus ?thesis
    by (simp add: chain_defs)
qed

text ‹
  The book introduces Theorem 9 before the above three lemmas but can only complete the proof
  once they are proven.
  This doesn't exactly say it the same way as the book, as the book gives the
  termlocal_ordering (abcd) explicitly (for arbitrarly named events), but is equivalent.
›

theorem (*9*) chain4:
  assumes path_Q: "Q  𝒫"
      and inQ: "a  Q" "b  Q" "c  Q" "d  Q"
      and abcd_neq: "a  b  a  c  a  d  b  c  b  d  c  d"
    shows "ch {a,b,c,d}"
proof -
  obtain a' b' c' where a'_pick: "a'  {a,b,c,d}"
                    and b'_pick: "b'  {a,b,c,d}"
                    and c'_pick: "c'  {a,b,c,d}"
                    and a'b'c': "[a'; b'; c']"
      using some_betw by (metis inQ(1,2,4) abcd_neq insert_iff path_Q)
  then obtain d' where d'_neq: "d'  a'  d'  b'  d'  c'"
                   and d'_pick: "d'  {a,b,c,d}"
    using insert_iff abcd_neq by metis
  have all_picked_on_path: "a'Q" "b'Q" "c'Q" "d'Q"
    using a'_pick b'_pick c'_pick d'_pick inQ by blast+
  consider "[d'; a'; b']" | "[a'; d'; b']" | "[a'; b'; d']"
    using some_betw abc_only_cba all_picked_on_path(1,2,4)
    by (metis a'b'c' d'_neq path_Q)
  then have picked_chain: "ch {a',b',c',d'}"
  proof (cases)
    assume "[d'; a'; b']"
    thus ?thesis using a'b'c' overlap_chain by (metis (full_types) insert_commute)
  next
    assume a'd'b': "[a'; d'; b']"
    then have "[d'; b'; c']" using abc_acd_bcd a'b'c' by blast
    thus ?thesis using a'd'b' overlap_chain by (metis (full_types) insert_commute)
  next
    assume a'b'd': "[a'; b'; d']"
    then have two_cases: "[b'; c'; d']  [b'; d'; c']" using abc_abd_bcdbdc a'b'c' d'_neq by blast
    (* Doing it this way avoids SMT. *)
    have case1: "[b'; c'; d']  ?thesis" using a'b'c' overlap_chain by blast
    have case2: "[b'; d'; c']  ?thesis"
        using abc_only_cba abc_acd_bcd a'b'd' overlap_chain
        by (metis (full_types) insert_commute)
    show ?thesis using two_cases case1 case2 by blast
  qed
  have "{a',b',c',d'} = {a,b,c,d}"
  proof (rule Set.set_eqI, rule iffI)
    fix x
    assume "x  {a',b',c',d'}"
    thus "x  {a,b,c,d}" using a'_pick b'_pick c'_pick d'_pick by auto
  next
    fix x
    assume x_pick: "x  {a,b,c,d}"
    have "a'  b'  a'  c'  a'  d'  b'  c'  c'  d'"
        using a'b'c' abc_abc_neq d'_neq by blast
    thus "x  {a',b',c',d'}"
        using a'_pick b'_pick c'_pick d'_pick x_pick d'_neq by auto
  qed
  thus ?thesis using picked_chain by simp
qed

theorem (*9*) chain4_alt:
  assumes path_Q: "Q  𝒫"
      and abcd_inQ: "{a,b,c,d}  Q"
      and abcd_distinct: "card {a,b,c,d} = 4"
    shows "ch {a,b,c,d}"
proof -
  have abcd_neq: "a  b  a  c  a  d  b  c  b  d  c  d"
    using abcd_distinct numeral_3_eq_3
    by (smt (z3) card_1_singleton_iff card_2_iff card_3_dist insert_absorb2 insert_commute numeral_1_eq_Suc_0 numeral_eq_iff semiring_norm(85) semiring_norm(88) verit_eq_simplify(8))
  have inQ: "a  Q" "b  Q" "c  Q" "d  Q"
    using abcd_inQ by auto
  show ?thesis using chain4[OF assms(1) inQ] abcd_neq by simp
qed


end (* context MinkowskiSpacetime *)


section "Interlude - Chains, segments, rays"

context MinkowskiBetweenness begin

subsection "General results for chains"

lemma inf_chain_is_long:
  assumes "[fX|x..]"
  shows "local_long_ch_by_ord f X  f 0 = x  infinite X"
  using chain_defs by (metis assms infinite_chain_alt)

text ‹A reassurance that the starting point $x$ is implied.›
lemma long_inf_chain_is_semifin:
  assumes "local_long_ch_by_ord f X  infinite X"
  shows " x. [fX|x..]"
  using assms infinite_chain_with_def chain_alts by auto

lemma endpoint_in_semifin:
  assumes  "[fX|x..]"
  shows "xX"
  using zero_into_ordering_loc by (metis assms empty_iff inf_chain_is_long local_long_ch_by_ord_alt)

text ‹
  Yet another corollary to Theorem 2, without indices, for arbitrary events on the chain.
›

corollary all_aligned_on_fin_chain:
  assumes "[fX]" "finite X"
  and x: "xX" and y: "yX" and z:"zX" and xy: "xy" and xz: "xz" and yz: "yz" 
  shows "[x;y;z]  [x;z;y]  [y;x;z]"
proof -
  have "card X  3" using assms(2-5) three_subset[OF xy xz yz] by blast
  hence 1: "local_long_ch_by_ord f X"
    using assms(1,3-) chain_defs by (metis short_ch_alt(1) short_ch_card(1) short_ch_card_2)
  obtain i j k where ijk: "x=f i" "i<card X" "y=f j" "j<card X" "z=f k" "k<card X"
    using obtain_index_fin_chain assms(1-5) by metis
  have 2: "[f i;f j;f k]" if "i<j  j<k" "k<card X" for i j k
    using assms order_finite_chain2 that(1,2) by auto
  consider "i<j  j<k"|"i<k  k<j"|"j<i  i<k"|"i>j  j>k"|"i>k  k>j"|"j>i  i>k"
    using xy xz yz ijk(1,3,5) by (metis linorder_neqE_nat)
  thus ?thesis
    apply cases using 2 abc_sym ijk by presburger+
qed

lemma (in MinkowskiPrimitive) card2_either_elt1_or_elt2:
  assumes "card X = 2" and "xX" and "yX" and "xy"
    and "zX" and "zx"
  shows "z=y"
by (metis assms card_2_iff')

(* potential misnomer: Schutz defines bounds only for infinite chains. *)
lemma get_fin_long_ch_bounds:
  assumes "local_long_ch_by_ord f X"
    and "finite X"
  shows "xX. yX. zX. [fX|x..y..z]"
proof (rule bexI)+
  show 1:"[fX|f 0..f 1..f (card X - 1)]"
    using assms unfolding finite_long_chain_with_def using index_injective
    by (auto simp: finite_chain_with_alt local_long_ch_by_ord_def local_ordering_def)
  show "f (card X - 1)  X"
    using 1 points_in_long_chain(3) by auto
  show "f 0  X" "f 1  X"
    using "1" points_in_long_chain by auto
qed

lemma get_fin_long_ch_bounds2:
  assumes "local_long_ch_by_ord f X"
    and "finite X"
  obtains x y z nx ny nz
  where "xX" "yX" "zX" "[fX|x..y..z]" "f nx = x" "f ny = y" "f nz = z"
  using get_fin_long_ch_bounds assms
  by (meson finite_chain_with_def finite_long_chain_with_alt index_middle_element)

lemma long_ch_card_ge3:
  assumes "ch_by_ord f X" "finite X"
  shows "local_long_ch_by_ord f X  card X  3"
  using assms ch_by_ord_def local_long_ch_by_ord_def short_ch_card(1) by auto

lemma fin_ch_betw2:
  assumes "[fX|a..c]" and "bX"
  obtains "b=a"|"b=c"|"[a;b;c]"
  by (metis assms finite_long_chain_with_alt finite_long_chain_with_def)

lemma chain_bounds_unique:
  assumes "[fX|a..c]" "[gX|x..z]"
  shows "(a=x  c=z)  (a=z  c=x)"
  using assms points_in_chain abc_abc_neq abc_bcd_acd abc_sym
  by (metis (full_types) fin_ch_betw2 )


end (* context MinkowskiBetweenness *)

subsection "Results for segments, rays and (sub)chains"


context MinkowskiBetweenness begin

lemma inside_not_bound:
  assumes "[fX|a..c]"
      and "j<card X"
    shows "j>0  f j  a" "j<card X - 1  f j  c"
  using index_injective2 assms finite_chain_def finite_chain_with_def apply metis
  using index_injective2 assms finite_chain_def finite_chain_with_def by auto


text ‹Converse to Theorem 2(i).›
lemma (in MinkowskiBetweenness) order_finite_chain_indices:
  assumes chX: "local_long_ch_by_ord f X" "finite X"
    and abc: "[a;b;c]"
    and ijk: "f i = a" "f j = b" "f k = c" "i<card X" "j<card X" "k<card X"
  shows "i<j  j<k  k<j  j<i"
  by (metis abc_abc_neq abc_only_cba(1,2,3) assms bot_nat_0.extremum linorder_neqE_nat order_finite_chain)


lemma order_finite_chain_indices2:
  assumes "[fX|a..c]"
    and "f j = b" "j<card X"
  obtains "0<j  j<(card X - 1)"|"j=(card X - 1)  b=c"|"j=0  b=a"
proof -
  have finX: "finite X"
    using assms(3) card.infinite gr_implies_not0 by blast
  have "bX"
    using assms unfolding chain_defs local_ordering_def
    by (metis One_nat_def card_2_iff insertI1 insert_commute less_2_cases)
  have a: "f 0 = a" and c: "f (card X - 1) = c"
    using assms(1) finite_chain_with_def by auto

  have "0<j  j<(card X - 1)  j=(card X - 1)  b=c  j=0  b=a"
  proof (cases "short_ch_by_ord f X")
    case True
    hence "X={a,c}"
      using a assms(1) first_neq_last points_in_chain short_ch_by_ord_def by fastforce
    then consider "b=a"|"b=c"
      using bX by fastforce
    thus ?thesis
      apply cases using assms(2,3) a c le_less by fastforce+
  next
    case False
    hence chX: "local_long_ch_by_ord f X"
      using assms(1) unfolding finite_chain_with_alt using chain_defs by meson
    consider "[a;b;c]"|"b=a"|"b=c"
      using bX assms(1) fin_ch_betw2 by blast
    thus ?thesis apply cases
      using f 0 = a chX finX assms(2,3) a c order_finite_chain_indices apply fastforce
      using f 0 = a chX finX assms(2,3) index_injective apply blast
      using a c assms chX finX index_injective linorder_neqE_nat inside_not_bound(2) by metis
  qed
  thus ?thesis using that by blast
qed


lemma index_bij_betw_subset:
  assumes chX: "[fX|a..b..c]" "f i = b" "card X > i"
  shows "bij_betw f {0<..<i} {eX. [a;e;b]}"
proof (unfold bij_betw_def, intro conjI)
  have chX2: "local_long_ch_by_ord f X" "finite X"
    using assms unfolding chain_defs apply (metis chX(1)
      abc_ac_neq fin_ch_betw points_in_long_chain(1,3) short_ch_alt(1) short_ch_path)
    using assms unfolding chain_defs by simp
  from index_bij_betw[OF this] have 1: "bij_betw f {0..<card X} X" .
  have "{0<..<i}  {0..<card X}"
    using assms(1,3) unfolding chain_defs by fastforce
  show "inj_on f {0<..<i}"
    using 1 assms(3) unfolding bij_betw_def
    by (smt (z3) atLeastLessThan_empty_iff2 atLeastLessThan_iff empty_iff greaterThanLessThan_iff
      inj_on_def less_or_eq_imp_le)
  show "f ` {0<..<i} = {e  X. [a;e;b]}"
  proof
    show "f ` {0<..<i}  {e  X. [a;e;b]}"
    proof (auto simp add: image_subset_iff conjI)
      fix j assume asm: "j>0" "j<i"
      hence "j < card X" using chX(3) less_trans by blast
      thus "f j  X" "[a;f j;b]"
        using chX(1) asm(1) unfolding chain_defs local_ordering_def
        apply (metis chX2(1) chX(1) fin_chain_card_geq_2 short_ch_card_2 short_xor_long(2)
          le_antisym set_le_two finite_chain_def finite_chain_with_def finite_long_chain_with_alt)
        using chX asm chX2(1) order_finite_chain unfolding chain_defs local_ordering_def by force
    qed
    show "{e  X. [a;e;b]}  f ` {0<..<i}"
    proof (auto)
      fix e assume e: "e  X" "[a;e;b]"
      obtain j where "f j = e" "j<card X"
        using e chX2 unfolding chain_defs local_ordering_def by blast
      show "e  f ` {0<..<i}"
      proof
        have "0<jj<i  i<jj<0"
          using order_finite_chain_indices chX chain_defs
          by (smt (z3) f j = e j < card X chX2(1) e(2) gr_implies_not_zero linorder_neqE_nat)
        hence "j<i" by simp
        thus "j{0<..<i}" "e = f j"
          using 0 < j  j < i  i < j  j < 0 greaterThanLessThan_iff
          by (blast,(simp add: f j = e))
      qed
    qed
  qed
qed


lemma bij_betw_extend:
  assumes "bij_betw f A B"
    and "f a = b" "aA" "bB"
  shows "bij_betw f (insert a A) (insert b B)"
  by (smt (verit, ccfv_SIG) assms(1) assms(2) assms(4) bij_betwI' bij_betw_iff_bijections insert_iff)


lemma insert_iff2:
  assumes "aX" shows "insert a {xX. P x} = {xX. P x  x=a}"
  using insert_iff assms by blast


lemma index_bij_betw_subset2:
  assumes chX: "[fX|a..b..c]" "f i = b" "card X > i"
  shows "bij_betw f {0..i} {eX. [a;e;b]a=eb=e}"
proof -
  have "bij_betw f {0<..<i} {eX. [a;e;b]}" using index_bij_betw_subset[OF assms] .
  moreover have "0{0<..<i}" "i{0<..<i}" by simp+
  moreover have "a{eX. [a;e;b]}" "b{eX. [a;e;b]}" using abc_abc_neq by auto+
  moreover have "f 0 = a" "f i = b" using assms unfolding chain_defs by simp+
  moreover have "(insert b (insert a {eX. [a;e;b]})) = {eX. [a;e;b]a=eb=e}"
  proof -
    have 1: "(insert a {eX. [a;e;b]}) = {eX. [a;e;b]a=e}"
      using insert_iff2[OF points_in_long_chain(1)[OF chX(1)]] by auto
    have "b{eX. [a;e;b]a=e}"
      using abc_abc_neq chX(1) fin_ch_betw by fastforce
    thus "(insert b (insert a {eX. [a;e;b]})) = {eX. [a;e;b]a=eb=e}"
      using 1 insert_iff2 points_in_long_chain(2)[OF chX(1)] by auto
  qed
  moreover have "(insert i (insert 0 {0<..<i})) = {0..i}" using image_Suc_lessThan by auto
  ultimately show ?thesis using bij_betw_extend[of f]
    by (metis (no_types, lifting) chX(1) finite_long_chain_with_def insert_iff)
qed


lemma chain_shortening:
  assumes "[fX|a..b..c]"
  shows "[f  {eX. [a;e;b]  e=a  e=b} |a..b]"
proof (unfold finite_chain_with_def finite_chain_def, (intro conjI))

  text ‹Different forms of assumptions for compatibility with needed antecedents later.›
  show "f 0 = a" using assms unfolding chain_defs by simp
  have chX: "local_long_ch_by_ord f X"
    using assms first_neq_last points_in_long_chain(1,3) short_ch_card(1) chain_defs
    by (metis card2_either_elt1_or_elt2)
  have finX: "finite X"
    by (meson assms chain_defs)

  text ‹General facts about the shortened set, which we will call Y.›
  let ?Y = "{eX. [a;e;b]  e=a  e=b}"
  show finY: "finite ?Y"
    using assms finite_chain_def finite_chain_with_def finite_long_chain_with_alt by auto
  have "ab" "a?Y" "b?Y" "c?Y"
    using assms finite_long_chain_with_def apply simp
    using assms points_in_long_chain(1,2) apply auto[1]
    using assms points_in_long_chain(2) apply auto[1]
    using abc_ac_neq abc_only_cba(2) assms fin_ch_betw by fastforce
  from this(1-3) finY have cardY: "card ?Y  2"
    by (metis (no_types, lifting) card_le_Suc0_iff_eq not_less_eq_eq numeral_2_eq_2)

  text ‹Obtain index for b› (a› is at index 0›): this index i› is card ?Y - 1›.›
  obtain i where i: "i<card X" "f i=b"
    using assms unfolding chain_defs local_ordering_def using Suc_leI diff_le_self by force
  hence "i<card X - 1"
    using assms unfolding chain_defs
    by (metis Suc_lessI diff_Suc_Suc diff_Suc_eq_diff_pred minus_nat.diff_0 zero_less_diff)
  have card01: "i+1 = card {0..i}" by simp
  have bb: "bij_betw f {0..i} ?Y" using index_bij_betw_subset2[OF assms i(2,1)] Collect_cong by smt
  hence i_eq: "i = card ?Y - 1" using bij_betw_same_card by force
  thus "f (card ?Y - 1) = b" using i(2) by simp

  text ‹The path P› on which X› lies. If ?Y› has two arguments, P› makes it a short chain.›
  obtain P where P_def: "P𝒫" "XP" "Q. Q𝒫  XQ  Q=P"
    using fin_chain_on_path[of f X] assms unfolding chain_defs by force
  have "aP" "bP" using P_def by (meson assms in_mono points_in_long_chain)+

  consider (eq_1)"i=1"|(gt_1)"i>1" using a  b f 0 = a i(2) less_linear by blast
  thus "[f?Y]"
  proof (cases)
    case eq_1
    hence "{0..i}={0,1}" by auto
    hence "bij_betw f {0,1} ?Y" using bb by auto
    from bij_betw_imp_surj_on[OF this] show ?thesis
      unfolding chain_defs using P_def eq_1 a  b f 0 = a i(2) by blast
  next
    case gt_1
    have 1: "3card ?Y" using gt_1 cardY i_eq by linarith
    {
      fix n assume "n < card ?Y"
      hence "n<card X"
        using i<card X - 1 add_diff_inverse_nat i_eq nat_diff_split_asm by linarith
      have "f n  ?Y"
      proof (simp, intro conjI)
        show "f n  X"
          using n<card X assms chX chain_defs local_ordering_def by metis
        consider "0<n  n<card ?Y - 1"|"n=card ?Y - 1"|"n=0"
          using n<card ?Y nat_less_le zero_less_diff by linarith
        thus "[a;f n;b]  f n = a  f n = b"
          using i i_eq f 0 = a chX finX le_numeral_extra(3) order_finite_chain by fastforce
      qed
    } moreover {
      fix x assume "x?Y" hence "xX" by simp
      obtain ix where ix: "ix < card X" "f ix = x"
        using assms obtain_index_fin_chain chain_defs xX by metis
      have "ix < card ?Y"
      proof -
        consider "[a;x;b]"|"x=a"|"x=b" using x?Y by auto
        hence "(ix<i  ix<0)  ix=0  ix=i"
          apply cases
          apply (metis f 0=a chX finX i ix less_nat_zero_code neq0_conv order_finite_chain_indices)
          using f 0 = a chX finX ix index_injective apply blast
          by (metis chX finX i(2) ix index_injective linorder_neqE_nat)
        thus ?thesis using gt_1 i_eq by linarith
      qed
      hence "n. n < card ?Y  f n = x" using ix(2) by blast
    } moreover {
      fix n assume "Suc (Suc n) < card ?Y"
      hence "Suc (Suc n) < card X"
        using i(1) i_eq by linarith
      hence "[f n; f (Suc n); f (Suc (Suc n))]"
        using assms unfolding chain_defs local_ordering_def by auto
    }
    ultimately have 2: "local_ordering f betw ?Y"
      by (simp add: local_ordering_def finY)
    show ?thesis using 1 2 chain_defs by blast
  qed
qed


corollary ord_fin_ch_right:
  assumes "[fX|a..f i..c]" "ji" "j<card X"
  shows "[f i;f j;c]  j = card X - 1  j = i"
proof -
  consider (inter)"j>i  j<card X - 1"|(left)"j=i"|(right)"j=card X - 1"
    using assms(3,2) by linarith
  thus ?thesis
    apply cases
    using assms(1) chain_defs order_finite_chain2 apply force
    by simp+
qed

lemma f_img_is_subset:
  assumes "[fX|(f 0) ..]" "i0" "j>i" "Y=f`{i..j}"
  shows "YX"
proof
  fix x assume "xY"
  then obtain n where "n{i..j}" "f n = x"
    using assms(4) by blast
  hence "f n  X"
    by (metis local_ordering_def assms(1) inf_chain_is_long local_long_ch_by_ord_def)
  thus "xX"
    using f n = x by blast
qed


lemma i_le_j_events_neq:
  assumes "[fX|a..b..c]"
    and "i<j" "j<card X"
  shows "f i  f j"
  using chain_defs by (meson assms index_injective2)

lemma indices_neq_imp_events_neq:
  assumes "[fX|a..b..c]"
      and "ij" "j<card X" "i<card X"
    shows "f i  f j"
  by (metis assms i_le_j_events_neq less_linear)

end (* context MinkowskiChain *)

context MinkowskiSpacetime begin

lemma bound_on_path:
  assumes "Q𝒫" "[fX|(f 0)..]" "XQ" "is_bound_f b X f"
  shows "bQ"
proof -
  obtain a c where "aX" "cX" "[a;c;b]"
    using assms(4)
    by (metis local_ordering_def inf_chain_is_long is_bound_f_def local_long_ch_by_ord_def zero_less_one)
  thus ?thesis
    using abc_abc_neq assms(1) assms(3) betw_c_in_path by blast
qed

lemma pro_basis_change:
  assumes "[a;b;c]"
  shows "prolongation a c = prolongation b c" (is "?ac=?bc")
proof
  show "?ac  ?bc"
  proof
    fix x assume "x?ac"
    hence "[a;c;x]"
      by (simp add: pro_betw)
    hence "[b;c;x]"
      using assms abc_acd_bcd by blast
    thus "x?bc"
      using abc_abc_neq pro_betw by blast
  qed
  show "?bc  ?ac"
  proof
    fix x assume "x?bc"
    hence "[b;c;x]"
      by (simp add: pro_betw)
    hence "[a;c;x]"
      using assms abc_bcd_acd by blast
    thus "x?ac"
      using abc_abc_neq pro_betw by blast
  qed
qed

lemma adjoining_segs_exclusive:
  assumes "[a;b;c]"
  shows "segment a b  segment b c = {}"
proof (cases)
  assume "segment a b = {}" thus ?thesis by blast
next
  assume "segment a b  {}"
  have "xsegment a b  xsegment b c" for x
  proof
    fix x assume "xsegment a b"
    hence "[a;x;b]" by (simp add: seg_betw)
    have "¬[a;b;x]" by (meson [a;x;b] abc_only_cba)
    have "¬[b;x;c]"
      using ¬ [a;b;x] abd_bcd_abc assms by blast
    thus "xsegment b c"
      by (simp add: seg_betw)
  qed
  thus ?thesis by blast
qed

end (* context MinkowskiSpacetime *)

section "3.6 Order on a path - Theorems 10 and 11"

context MinkowskiSpacetime begin

subsection ‹Theorem 10 (based on Veblen (1904) theorem 10).›

lemma (in MinkowskiBetweenness) two_event_chain:
  assumes finiteX: "finite X"
      and path_Q: "Q  𝒫"
      and events_X: "X  Q"
      and card_X: "card X = 2"
    shows "ch X"
proof -
  obtain a b where X_is: "X={a,b}"
    using card_le_Suc_iff numeral_2_eq_2
    by (meson card_2_iff card_X)
  have no_c: "¬(c{a,b}. ca  cb)"
    by blast
  have "ab  aQ & bQ"
    using X_is card_X events_X by force
  hence "short_ch {a,b}"
    using path_Q no_c by (meson short_ch_intros(2))
  thus ?thesis
    by (simp add: X_is chain_defs)
qed

lemma (in MinkowskiBetweenness) three_event_chain:
  assumes finiteX: "finite X"
      and path_Q: "Q  𝒫"
      and events_X: "X  Q"
      and card_X: "card X = 3"
    shows "ch X"
proof -
  obtain a b c where X_is: "X={a,b,c}"
    using numeral_3_eq_3 card_X by (metis card_Suc_eq)
  then have all_neq: "ab  ac  bc"
    using card_X numeral_2_eq_2 numeral_3_eq_3
    by (metis Suc_n_not_le_n insert_absorb2 insert_commute set_le_two)
  have in_path: "aQ  bQ  cQ"
    using X_is events_X by blast
  hence "[a;b;c]  [b;c;a]  [c;a;b]"
    using some_betw all_neq path_Q by auto
  thus "ch X"
    using between_chain X_is all_neq chain3 in_path path_Q by auto
qed

text ‹This is case (i) of the induction in Theorem 10.›

lemma (*for 10*) chain_append_at_left_edge:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and bY: "[b; a1; an]"
    fixes g defines g_def: "g  (λj::nat. if j1 then f (j-1) else b)"
    shows "[g(insert b Y)|b .. a1 .. an]"
proof -
  let ?X = "insert b Y"
  have ord_fY: "local_ordering f betw Y" using long_ch_Y finite_long_chain_with_card chain_defs
    by (meson long_ch_card_ge3)
  have "bY"
    using abc_ac_neq abc_only_cba(1) assms by (metis fin_ch_betw2 finite_long_chain_with_alt)
  have bound_indices: "f 0 = a1  f (card Y - 1) = an"
    using long_ch_Y by (simp add: chain_defs)
  have fin_Y: "card Y  3"
    using finite_long_chain_with_def long_ch_Y numeral_2_eq_2 points_in_long_chain
    by (metis abc_abc_neq bY card2_either_elt1_or_elt2 fin_chain_card_geq_2 leI le_less_Suc_eq numeral_3_eq_3)
  hence num_ord: "0  (0::nat)  0<(1::nat)  1 < card Y - 1  card Y - 1 < card Y"
    by linarith
  hence "[a1; f 1; an]"
    using order_finite_chain chain_defs long_ch_Y
    by auto
  text ‹Schutz has a step here that says $[b a_1 a_2 a_n]$ is a chain (using Theorem 9).
    We have no easy way (yet) of denoting an ordered 4-element chain, so we skip this step
    using a termlocal_ordering lemma from our script for 3.6, which Schutz doesn't list.›
  hence "[b; a1; f 1]"
    using bY abd_bcd_abc by blast
  have "local_ordering g betw ?X"
  proof -
    {
      fix n assume "finite ?X  n<card ?X"
      have "g n  ?X"
        apply (cases "n1")
         prefer 2 apply (simp add: g_def)
      proof
        assume "1n" "g n  Y"
        hence "g n = f(n-1)" unfolding g_def by auto
        hence "g n  Y"
        proof (cases "n = card ?X - 1")
          case True
          thus ?thesis
            using bY card.insert diff_Suc_1 long_ch_Y points_in_long_chain chain_defs
            by (metis g n = f (n - 1))
        next
          case False
          hence "n < card Y"
            using points_in_long_chain finite ?X  n < card ?X g n = f (n - 1) g n  Y bY chain_defs
            by (metis card.insert finite_insert long_ch_Y not_less_simps(1))
          hence "n-1 < card Y - 1"
            using 1  n diff_less_mono by blast
          hence "f(n-1)Y"
            using long_ch_Y fin_Y unfolding chain_defs local_ordering_def
            by (metis Suc_le_D card_3_dist diff_Suc_1 insert_absorb2 le_antisym less_SucI numeral_3_eq_3 set_le_three)
          thus ?thesis
            using g n = f (n - 1) by presburger
        qed
        hence False using g n  Y by auto
        thus "g n = b" by simp
      qed
    } moreover {
      fix n assume "(finite ?X  Suc(Suc n) < card ?X)"
      hence "[g n; g (Suc n); g (Suc(Suc n))]"
        apply (cases "n1")
        using bY [b; a1; f 1] g_def ordering_ord_ijk_loc[OF ord_fY] fin_Y
        apply (metis Suc_diff_le card_insert_disjoint diff_Suc_1 finite_insert le_Suc_eq not_less_eq)
        by (metis One_nat_def Suc_leI [b;a1;f 1] bound_indices diff_Suc_1 g_def
          not_less_less_Suc_eq zero_less_Suc)
    } moreover {
      fix x assume "x?X" "x=b"
      have "(finite ?X  0 < card ?X)  g 0 = x"
        by (simp add: bY x = b g_def)
    } moreover {
      fix x assume "x?X" "xb"
      hence "n. (finite ?X  n < card ?X)  g n = x"
      proof -
        obtain n where "f n = x" "n < card Y"
          using x?X xb local_ordering_def insert_iff long_ch_Y chain_defs by (metis ord_fY)
        have "(finite ?X  n+1 < card ?X)" "g(n+1) = x"
          apply (simp add: bY n < card Y)
          by (simp add: f n = x g_def)
        thus ?thesis by auto
      qed
    }
    ultimately show ?thesis
      unfolding local_ordering_def
      by smt
  qed
  hence "local_long_ch_by_ord g ?X"
    unfolding local_long_ch_by_ord_def
    using fin_Y bY
    by (meson card_insert_le finite_insert le_trans)
  show ?thesis
  proof (intro finite_long_chain_with_alt2)
    show "local_long_ch_by_ord g ?X" using local_long_ch_by_ord g ?X by simp
    show "[b;a1;an]  a1  ?X" using bY long_ch_Y points_in_long_chain(1) by auto
    show "g 0 = b" using g_def by simp
    show "finite ?X" 
      using fin_Y bY eval_nat_numeral by (metis card.infinite finite.insertI not_numeral_le_zero)
    show "g (card ?X - 1) = an"
      using g_def bY bound_indices eval_nat_numeral
      by (metis One_nat_def card.infinite card_insert_disjoint diff_Suc_Suc
        diff_is_0_eq' less_nat_zero_code minus_nat.diff_0 nat_le_linear num_ord)
  qed
qed

text ‹
  This is case (iii) of the induction in Theorem 10.
  Schutz says merely ``The proof for this case is similar to that for Case (i).''
  Thus I feel free to use a result on symmetry, rather than going through
  the pain of Case (i) (chain_append_at_left_edge›) again.
›
lemma (*for 10*) chain_append_at_right_edge:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and Yb: "[a1; an; b]"
    fixes g defines g_def: "g  (λj::nat. if j  (card Y - 1) then f j else b)"
    shows "[g(insert b Y)|a1 .. an .. b]"
proof -
  let ?X = "insert b Y"
  have "bY"
    using Yb abc_abc_neq abc_only_cba(2) long_ch_Y
    by (metis fin_ch_betw2 finite_long_chain_with_def)
  have fin_Y: "card Y  3"
    using finite_long_chain_with_card long_ch_Y by auto
  hence fin_X: "finite ?X"
    by (metis card.infinite finite.insertI not_numeral_le_zero)
  have "a1Y  anY  aY"
    using long_ch_Y points_in_long_chain by meson
  have "a1a  a an  a1an"
    using Yb abc_abc_neq finite_long_chain_with_def long_ch_Y by auto
  have "Suc (card Y) = card ?X"
    using bY fin_X finite_long_chain_with_def long_ch_Y by auto
  obtain f2 where f2_def: "[f2Y|an..a..a1]" "f2=(λn. f (card Y - 1 - n))"
    using chain_sym long_ch_Y by blast
  obtain g2 where g2_def: "g2 = (λj::nat. if j1 then f2 (j-1) else b)"
    by simp
  have "[b; an; a1]"
    using abc_sym Yb by blast
  hence g2_ord_X: "[g2?X|b .. an .. a1]"
    using chain_append_at_left_edge [where a1="an" and an="a1" and f="f2"]
      fin_X bY f2_def g2_def
    by blast
  then obtain g1 where g1_def: "[g1?X|a1..an..b]" "g1=(λn. g2 (card ?X - 1 - n))"
    using chain_sym by blast
  have sYX: "(card Y) = (card ?X) - 1"
    using assms(2,3) finite_long_chain_with_def long_ch_Y Suc (card Y) = card ?X by linarith
  have "g1=g"
    unfolding g1_def g2_def f2_def g_def
  proof
    fix n
    show "(
            if 1  card ?X - 1 - n then
              f (card Y - 1 - (card ?X - 1 - n - 1))
            else b
          ) = (
            if n  card Y - 1 then
              f n
            else b
          )" (is "?lhs=?rhs")
    proof (cases)
      assume "n  card ?X - 2"
      show "?lhs=?rhs"
        using n  card ?X - 2 finite_long_chain_with_def long_ch_Y sYX Suc (card Y) = card ?X
        by (metis (mono_tags, opaque_lifting) Suc_1 Suc_leD diff_Suc_Suc diff_commute diff_diff_cancel
          diff_le_mono2 fin_chain_card_geq_2)
    next
      assume "¬ n  card ?X - 2"
      thus "?lhs=?rhs"
        by (metis Suc (card Y) = card ?X Suc_1 diff_Suc_1 diff_Suc_eq_diff_pred diff_diff_cancel
            diff_is_0_eq' nat_le_linear not_less_eq_eq)
    qed
  qed
  thus ?thesis
    using g1_def(1) by blast
qed


lemma S_is_dense:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and S_def: "S = {k::nat. [a1; f k; b]  k < card Y}"
      and k_def: "S{}" "k = Max S"
      and k'_def: "k'>0" "k'<k"
  shows "k'  S"
proof -
  text ‹
    We will prove this by contradiction. We can obtain the path that Y› lies on, and show b› is
    on it too. Then since f`S› must be on this path, there must be an ordering involving b›, f k›
    and f k'› that leads to contradiction with the definition of S› and k∉S›.
    Notice we need no knowledge about b› except how it relates to S›.
  ›
  have "[fY]" using long_ch_Y chain_defs by meson
  have "card Y  3" using finite_long_chain_with_card long_ch_Y by blast
  hence "finite Y" by (metis card.infinite not_numeral_le_zero)
  have "kS" using k_def Max_in S_def by (metis finite_Collect_conjI finite_Collect_less_nat)
  hence "k<card Y" using S_def by auto
  have "k'<card Y" using S_def k'_def kS by auto
  show "k'  S"
  proof (rule ccontr)
    assume asm: "¬k'S"
    have 1: "[f 0;f k;f k']"
    proof -
      have "[a1; b; f k']"
        using order_finite_chain2 long_ch_Y k  S k' < card Y chain_defs
        by (smt (z3) abc_acd_abd asm le_numeral_extra(3) assms mem_Collect_eq)
      have "[a1; f k; b]"
        using S_def k  S by blast
      have "[f k; b; f k']"
        using abc_acd_bcd [a1; b; f k'] [a1; f k; b] by blast
      thus ?thesis
        using [a1;f k;b] long_ch_Y unfolding finite_long_chain_with_def finite_chain_with_def
        by blast
    qed
    have 2: "[f 0;f k';f k]"
      apply (intro order_finite_chain2[OF [fY] finite Y]) by (simp add: k < card Y k'_def)
    show False using 1 2 abc_only_cba(2) by blast
  qed
qed


lemma (*for 10*) smallest_k_ex:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and Y_def: "bY"
      and Yb: "[a1; b; an]"
    shows "k>0. [a1; b; f k]  k < card Y  ¬(k'<k. [a1; b; f k'])"
proof -
(* the usual suspects first, they'll come in useful I'm sure *)
  have bound_indices: "f 0 = a1  f (card Y - 1) = an"
    using chain_defs long_ch_Y by auto
  have fin_Y: "finite Y"
    using chain_defs long_ch_Y by presburger
  have card_Y: "card Y  3"
    using long_ch_Y points_in_long_chain finite_long_chain_with_card by blast

  text ‹We consider all indices of chain elements between a1 and b›, and find the maximal one.›
  let ?S = "{k::nat. [a1; f k; b]  k < card Y}"
  obtain S where S_def: "S=?S"
    by simp
  have "S{0..card Y}"
    using S_def by auto
  hence "finite S"
    using finite_subset by blast

  show ?thesis
  proof (cases)
    assume "S={}"
    show ?thesis
    proof
      show "(0::nat)<1  [a1; b; f 1]  1 < card Y  ¬ (k'::nat. k' < 1  [a1; b; f k'])"
      proof (intro conjI)
        show "(0::nat)<1" by simp
        show "1 < card Y"
          using Yb abc_ac_neq bound_indices not_le by fastforce
        show "¬ (k'::nat. k' < 1  [a1; b; f k'])"
          using abc_abc_neq bound_indices
          by blast
        show "[a1; b; f 1]"
        proof -
          have "f 1  Y"
            using long_ch_Y chain_defs local_ordering_def by (metis 1 < card Y short_ch_ord_in(2))
          hence "[a1; f 1; an]"
            using bound_indices long_ch_Y chain_defs local_ordering_def card_Y
            by (smt (z3) Nat.lessE One_nat_def Suc_le_lessD Suc_lessD diff_Suc_1 diff_Suc_less
              fin_ch_betw2 i_le_j_events_neq less_numeral_extra(1) numeral_3_eq_3)
          hence "[a1; b; f 1]  [a1; f 1; b]  [b; a1; f 1]"
            using abc_ex_path_unique some_betw abc_sym
            by (smt Y_def Yb f 1  Y abc_abc_neq cross_once_notin)
          thus "[a1; b; f 1]"
          proof -
            have "n. ¬ ([a1; f n; b]  n < card Y)"
              using S_def S = {}
              by blast
            then have "[a1; b; f 1]  ¬ [an; f 1; b]  ¬ [a1; f 1; b]"
              using bound_indices abc_sym abd_bcd_abc Yb
              by (metis (no_types) diff_is_0_eq' nat_le_linear nat_less_le)
            then show ?thesis
              using abc_bcd_abd abc_sym
              by (meson [a1; b; f 1]  [a1; f 1; b]  [b; a1; f 1] [a1; f 1; an])
          qed
        qed
      qed
    qed
  next assume "¬S={}"
    obtain k where "k = Max S"
      by simp
    hence  "k  S" using Max_in
      by (simp add: S  {} finite S)
    have "k1"
    proof (rule ccontr)
      assume "¬ 1  k"
      hence "k=0" by simp
      have "[a1; f k; b]"
        using kS S_def
        by blast
      thus False
        using bound_indices k = 0 abc_abc_neq
        by blast
    qed

    show ?thesis
    proof
      let ?k = "k+1"
      show "0<?k  [a1; b; f ?k]  ?k < card Y  ¬ (k'::nat. k' < ?k  [a1; b; f k'])"
      proof (intro conjI)
        show "(0::nat)<?k" by simp
        show "?k < card Y"
          by (metis (no_types, lifting) S_def Yb k  S abc_only_cba(2) add.commute
              add_diff_cancel_right' bound_indices less_SucE mem_Collect_eq nat_add_left_cancel_less
              plus_1_eq_Suc)
        show "[a1; b; f ?k]"
        proof -
          have "f ?k  Y"
            using k + 1 < card Y long_ch_Y card_Y unfolding local_ordering_def chain_defs
            by (metis One_nat_def Suc_numeral not_less_eq_eq numeral_3_eq_3 numerals(1) semiring_norm(2) set_le_two)
          have "[a1; f ?k; an]  f ?k = an"
            using fin_ch_betw2 inside_not_bound(1) long_ch_Y chain_defs
            by (metis 0 < k + 1 k + 1 < card Y f (k + 1)  Y)
          thus  "[a1; b; f ?k]"
          proof (rule disjE)
            assume "[a1; f ?k; an]"
            hence "f ?k  an"
              by (simp add: abc_abc_neq)
            hence "[a1; b; f ?k]  [a1; f ?k; b]  [b; a1; f ?k]"
              using abc_ex_path_unique some_betw abc_sym [a1; f ?k; an]
                f ?k  Y Yb abc_abc_neq assms(3) cross_once_notin
              by (smt Y_def)
            moreover have "¬ [a1; f ?k; b]"
            proof
              assume "[a1; f ?k; b]"
              hence "?k  S"
                using S_def [a1; f ?k; b] k + 1 < card Y by blast
              hence "?k  k"
                by (simp add: finite S k = Max S)
              thus False
                by linarith
            qed
            moreover have "¬ [b; a1; f ?k]"
              using Yb [a1; f ?k; an] abc_only_cba
              by blast
            ultimately show "[a1; b; f ?k]"
              by blast
          next assume "f ?k = an"
            show ?thesis
              using Yb f (k + 1) = an by blast
          qed
        qed
        show "¬(k'::nat. k' < k + 1  [a1; b; f k'])"
        proof
          assume "k'::nat. k' < k + 1  [a1; b; f k']"
          then obtain k' where k'_def: "k'>0" "k' < k + 1" "[a1; b; f k']"
            using abc_ac_neq bound_indices neq0_conv
            by blast
          hence "k'<k"
            using S_def k  S abc_only_cba(2) less_SucE by fastforce
          hence "k'S"
            using S_is_dense long_ch_Y S_def ¬S={} k = Max S k'>0
            by blast
          thus False
            using S_def abc_only_cba(2) k'_def(3) by blast
        qed
      qed
    qed
  qed
qed


(* TODO: there's definitely a way of doing this using chain_sym and smallest_k_ex. *)
lemma greatest_k_ex:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and Y_def: "bY"
      and Yb: "[a1; b; an]"
    shows "k. [f k; b; an]  k < card Y - 1  ¬(k'<card Y. k'>k  [f k'; b; an])"
proof -
  have bound_indices: "f 0 = a1  f (card Y - 1) = an"
    using chain_defs long_ch_Y by simp
  have fin_Y: "finite Y"
    using chain_defs long_ch_Y by presburger
  have card_Y: "card Y  3"
    using long_ch_Y points_in_long_chain finite_long_chain_with_card by blast
  have chY2: "local_long_ch_by_ord f Y"
    using long_ch_Y chain_defs by (meson card_Y long_ch_card_ge3)

  text ‹Again we consider all indices of chain elements between a1 and b›.›
  let ?S = "{k::nat. [an; f k; b]  k < card Y}"
  obtain S where S_def: "S=?S"
    by simp
  have "S{0..card Y}"
    using S_def by auto
  hence "finite S"
    using finite_subset by blast

  show ?thesis
  proof (cases)
    assume "S={}"
    show ?thesis
    proof
      let ?n = "card Y - 2"
      show "[f ?n; b; an]  ?n < card Y - 1  ¬(k'<card Y. k'>?n  [f k'; b; an])"
      proof (intro conjI)
        show "?n < card Y - 1"
          using Yb abc_ac_neq bound_indices not_le by fastforce
      next show "¬(k'<card Y. k'>?n  [f k'; b; an])"
          using abc_abc_neq bound_indices
          by (metis One_nat_def Suc_diff_le Suc_leD Suc_lessI card_Y diff_Suc_1 diff_Suc_Suc
              not_less_eq numeral_2_eq_2 numeral_3_eq_3)
      next show "[f ?n; b; an]"
        proof -
          have "[f 0;f ?n; f (card Y - 1)]"
            apply (intro order_finite_chain[of f Y], (simp_all add: chY2 fin_Y))
            using card_Y by linarith
          hence "[a1; f ?n; an]"
            using long_ch_Y unfolding chain_defs by simp
          have "f ?n  Y"
            using long_ch_Y eval_nat_numeral unfolding local_ordering_def chain_defs
            by (metis card_1_singleton_iff card_Suc_eq card_gt_0_iff diff_Suc_less diff_self_eq_0 insert_iff numeral_2_eq_2)
          hence "[an; b; f ?n]  [an; f ?n; b]  [b; an; f ?n]"
            using abc_ex_path_unique some_betw abc_sym [a1; f ?n; an]
            by (smt Y_def Yb f ?n  Y abc_abc_neq cross_once_notin)
          thus "[f ?n; b; an]"
          proof -
            have "n. ¬ ([an; f n; b]  n < card Y)"
              using S_def S = {}
              by blast
            then have "[an; b; f ?n]  ¬ [a1; f ?n; b]  ¬ [an; f ?n; b]"
              using bound_indices abc_sym abd_bcd_abc Yb
              by (metis (no_types, lifting) f (card Y - 2)  Y card_gt_0_iff diff_less empty_iff fin_Y zero_less_numeral)
            then show ?thesis
              using abc_bcd_abd abc_sym
              by (meson [an; b; f ?n]  [an; f ?n; b]  [b; an; f ?n] [a1; f ?n; an])
          qed
        qed
      qed
    qed
  next assume "¬S={}"
    obtain k where "k = Min S"
      by simp
    hence  "k  S"
      by (simp add: S  {} finite S)

    show ?thesis
    proof
      let ?k = "k-1"
      show "[f ?k; b; an]  ?k < card Y - 1  ¬ (k'<card Y. ?k < k'  [f k'; b; an])"
      proof (intro conjI)
        show "?k < card Y - 1"
          using S_def k  S less_imp_diff_less card_Y
          by (metis (no_types, lifting) One_nat_def diff_is_0_eq' diff_less_mono lessI less_le_trans
              mem_Collect_eq nat_le_linear numeral_3_eq_3 zero_less_diff)
        show "[f ?k; b; an]"
        proof -
          have "f ?k  Y"
            using k - 1 < card Y - 1 long_ch_Y card_Y eval_nat_numeral unfolding local_ordering_def chain_defs
            by (metis Suc_pred' less_Suc_eq less_nat_zero_code not_less_eq not_less_eq_eq set_le_two)
          have "[a1; f ?k; an]  f ?k = a1"
            using bound_indices long_ch_Y k - 1 < card Y - 1 chain_defs
            unfolding finite_long_chain_with_alt
            by (metis f (k - 1)  Y card_Diff1_less card_Diff_singleton_if chY2 index_injective)
          thus  "[f ?k; b; an]"
          proof (rule disjE)
            assume "[a1; f ?k; an]"
            hence "f ?k  a1"
              using abc_abc_neq by blast
            hence "[an; b; f ?k]  [an; f ?k; b]  [b; an; f ?k]"
              using abc_ex_path_unique some_betw abc_sym [a1; f ?k; an]
                f ?k  Y Yb abc_abc_neq assms(3) cross_once_notin
              by (smt Y_def)
            moreover have "¬ [an; f ?k; b]"
            proof
              assume "[an; f ?k; b]"
              hence "?k  S"
                using S_def [an; f ?k; b] k - 1 < card Y - 1
                by simp
              hence "?k  k"
                by (simp add: finite S k = Min S)
              thus False
                using f (k - 1)  a1 chain_defs long_ch_Y
                by auto
            qed
            moreover have "¬ [b; an; f ?k]"
              using Yb [a1; f ?k; an] abc_only_cba(2) abc_bcd_acd
              by blast
            ultimately show "[f ?k; b; an]"
              using abc_sym by auto
          next assume "f ?k = a1"
            show ?thesis
              using Yb f (k - 1) = a1 by blast
          qed
        qed
        show "¬(k'<card Y. k-1 < k'  [f k'; b; an])"
        proof
          assume "k'<card Y. k-1 < k'  [f k'; b; an]"
          then obtain k' where k'_def: "k'<card Y -1" "k' > k - 1" "[an; b; f k']"
            using abc_ac_neq bound_indices neq0_conv
            by (metis Suc_diff_1 abc_sym gr_implies_not0 less_SucE)
          hence "k'>k"
            using S_def k  S abc_only_cba(2) less_SucE
            by (metis (no_types, lifting) add_diff_inverse_nat less_one mem_Collect_eq
                not_less_eq plus_1_eq_Suc)thm S_is_dense
          hence "k'S"         
            apply (intro S_is_dense[of f Y a1 a an _ b "Max S"])
            apply (simp add: long_ch_Y)
            apply (smt (verit, ccfv_SIG) S_def k  S abc_acd_abd abc_only_cba(4)
              add_diff_inverse_nat bound_indices chY2 diff_add_zero diff_is_0_eq fin_Y k'_def(1,3)
              less_add_one less_diff_conv2 less_nat_zero_code mem_Collect_eq nat_diff_split order_finite_chain)
            apply (simp add: S  {}, simp, simp)
            using k'_def S_def
            by (smt (verit, ccfv_SIG) k  S abc_acd_abd abc_only_cba(4) add_diff_cancel_right'
              add_diff_inverse_nat bound_indices chY2 fin_Y le_eq_less_or_eq less_nat_zero_code
              mem_Collect_eq nat_diff_split nat_neq_iff order_finite_chain zero_less_diff zero_less_one)
          thus False
            using S_def abc_only_cba(2) k'_def(3)
            by blast
        qed
      qed
    qed
  qed
qed


lemma get_closest_chain_events:
  assumes long_ch_Y: "[fY|a0..a..an]"
      and x_def: "xY" "[a0; x; an]"
    obtains nb nc b c
      where "b=f nb" "c=f nc" "[b;x;c]" "bY" "cY" "nb = nc - 1" "nc<card Y" "nc>0"
            "¬(k < card Y. [f k; x; an]  k>nb)" "¬(k<nc. [a0; x; f k])"
proof -
  have " nb nc b c. b=f nb  c=f nc  [b;x;c]  bY  cY  nb = nc - 1  nc<card Y  nc>0
     ¬(k < card Y. [f k; x; an]  k>nb)  ¬(k < nc. [a0; x; f k])"
  proof -
    have bound_indices: "f 0 = a0  f (card Y - 1) = an"
      using chain_defs long_ch_Y by simp
    have fin_Y: "finite Y"
      using chain_defs long_ch_Y by presburger
    have card_Y: "card Y  3"
      using long_ch_Y points_in_long_chain finite_long_chain_with_card by blast
    have chY2: "local_long_ch_by_ord f Y"
      using long_ch_Y chain_defs by (meson card_Y long_ch_card_ge3)
    obtain P where P_def: "P𝒫" "YP"
      using fin_chain_on_path long_ch_Y fin_Y chain_defs by meson
    hence "xP"
      using betw_b_in_path x_def(2) long_ch_Y points_in_long_chain
      by (metis abc_abc_neq in_mono)
    obtain nc where nc_def: "¬(k. [a0; x; f k]  k<nc)" "[a0; x; f nc]" "nc<card Y" "nc>0"
      using smallest_k_ex [where a1=a0 and a=a and an=an and b=x and f=f and Y=Y]
        long_ch_Y x_def
      by blast
    then obtain c where c_def: "c=f nc" "cY"
      using chain_defs local_ordering_def by (metis chY2)
    have c_goal: "c=f nc  cY  nc<card Y  nc>0  ¬(k < card Y. [a0; x; f k]  k<nc)"
      using c_def nc_def(1,3,4) by blast
    obtain nb where nb_def: "¬(k < card Y. [f k; x; an]  k>nb)" "[f nb; x; an]" "nb<card Y-1"
      using greatest_k_ex [where a1=a0 and a=a and an=an and b=x and f=f and Y=Y]
        long_ch_Y x_def
      by blast
    hence "nb<card Y"
      by linarith
    then obtain b where b_def: "b=f nb" "bY"
      using nb_def chY2 local_ordering_def by (metis local_long_ch_by_ord_alt)
    have "[b;x;c]"
    proof -
      have "[b; x; an]"
        using b_def(1) nb_def(2) by blast
      have "[a0; x; c]"
        using c_def(1) nc_def(2) by blast
      moreover have "a. [a;x;b]  ¬ [a; an; x]"
        using [b; x; an] abc_bcd_acd
        by (metis (full_types) abc_sym)
      moreover have "a. [a;x;b]  ¬ [an; a; x]"
        using [b; x; an] by (meson abc_acd_bcd abc_sym)
      moreover have "an = c  [b;x;c]"
        using [b; x; an] by meson
      ultimately show ?thesis
        using abc_abd_bcdbdc abc_sym x_def(2)
        by meson
    qed
    have "nb<nc"
      using [b;x;c] nc<card Y nb<card Y c = f nc b = f nb
      by (smt (z3) abc_abd_bcdbdc abc_ac_neq abc_acd_abd abc_only_cba(4) abc_sym bot_nat_0.extremum
      bound_indices chY2 fin_Y nat_neq_iff nc_def(2) nc_def(4) order_finite_chain)
    have "nb = nc - 1"
    proof (rule ccontr)
      assume "nb  nc - 1"
      have "nb<nc-1"
        using nb  nc - 1 nb<nc by linarith
      hence "[f nb; (f(nc-1)); f nc]"
        using nb  nc - 1 long_ch_Y nc_def(3) order_finite_chain chain_defs
        by auto
      have "¬[a0; x; (f(nc-1))]"
        using nc_def(1,4) diff_less less_numeral_extra(1)
        by blast
      have "nc-10"
        using nb < nc nb  nc - 1 by linarith
      hence "f(nc-1)a0  a0x"
        using bound_indices nb < nc - 1 abc_abc_neq less_imp_diff_less nb_def(1) nc_def(3) x_def(2)
        by blast
      have "xf(nc-1)"
        using x_def(1) nc_def(3) chY2 unfolding chain_defs local_ordering_def
        by (metis One_nat_def Suc_pred less_Suc_eq nc_def(4) not_less_eq)
      hence "[a0; f (nc-1); x]"
        using long_ch_Y nc_def c_def chain_defs
        by (metis [f nb;f (nc - 1);f nc] ¬ [a0;x;f (nc - 1)] abc_ac_neq abc_acd_abd abc_bcd_acd
          abd_acd_abcacb abd_bcd_abc b_def(1) b_def(2) fin_ch_betw2 nb_def(2))
      hence "[(f(nc-1)); x; an]"
        using abc_acd_bcd x_def(2) by blast
      thus False using nb_def(1)
        using nb < nc - 1 less_imp_diff_less nc_def(3)
        by blast
    qed
    have b_goal: "b=f nb  bY  nb=nc-1  ¬(k < card Y. [f k; x; an]  k>nb)"
      using b_def nb_def(1) nb_def(3) nb=nc-1 by blast
    thus ?thesis
      using [b;x;c] c_goal
      using nb < card Y nc_def(1) by auto
  qed
  thus ?thesis
    using that by auto
qed


text ‹This is case (ii) of the induction in Theorem 10.›
lemma (*for 10*) chain_append_inside:
  assumes long_ch_Y: "[fY|a1..a..an]"
      and Y_def: "bY"
      and Yb: "[a1; b; an]"
      and k_def: "[a1; b; f k]" "k < card Y" "¬(k'. (0::nat)<k'  k'<k  [a1; b; f k'])"
    fixes g
  defines g_def: "g  (λj::nat. if (jk-1) then f j else (if (j=k) then b else f (j-1)))"
    shows "[ginsert b Y|a1 .. b .. an]"
proof -
  let ?X = "insert b Y"
  have fin_X: "finite ?X"
    by (meson chain_defs finite.insertI long_ch_Y)
  have bound_indices: "f 0 = a1  f (card Y - 1) = an"
    using chain_defs long_ch_Y
    by auto
  have fin_Y: "finite Y"
    using chain_defs long_ch_Y by presburger
  have f_def: "local_long_ch_by_ord f Y"
    using chain_defs long_ch_Y by (meson finite_long_chain_with_card long_ch_card_ge3)
  have  a1  an  a1  b  b  an
    using Yb abc_abc_neq by blast
  have "k  0"
    using abc_abc_neq bound_indices k_def
    by metis

  have b_middle: "[f(k-1); b; f k]"
  proof (cases)
    assume "k=1" show "[f(k-1); b; f k]"
      using [a1; b; f k] k = 1 bound_indices by auto
  next assume "k1" show "[f(k-1); b; f k]"
    proof -
      have a1k: "[a1; f (k-1); f k]" using bound_indices
        using k < card Y k  0 k  1 long_ch_Y fin_Y order_finite_chain
        unfolding chain_defs by auto
      text ‹In fact, the comprehension below gives the order of elements too.
         Our notation and Theorem 9 are too weak to say that just now.›
      have ch_with_b: "ch {a1, (f (k-1)), b, (f k)}" using chain4
        using k_def(1) abc_ex_path_unique between_chain cross_once_notin
        by (smt [a1; f (k-1); f k] abc_abc_neq insert_absorb2)
      have "f (k-1)  b  (f k)  (f (k-1))  b  (f k)"
        using abc_abc_neq f_def k_def(2) Y_def
        by (metis local_ordering_def [a1; f (k-1); f k] less_imp_diff_less local_long_ch_by_ord_def)
      hence some_ord_bk: "[f(k-1); b; f k]  [b; f (k-1); f k]  [f (k-1); f k; b]"
        using fin_chain_on_path ch_with_b some_betw Y_def chain_defs
        by (metis a1k abc_acd_bcd abd_acd_abcacb k_def(1))
      thus "[f(k-1); b; f k]"
      proof -
        have "¬ [a1; f k; b]"
          by (simp add: [a1; b; f k] abc_only_cba(2))
        thus ?thesis
          using some_ord_bk k_def abc_bcd_acd abd_bcd_abc bound_indices
          by (metis diff_is_0_eq' diff_less less_imp_diff_less less_irrefl_nat not_less
              zero_less_diff zero_less_one [a1; b; f k] a1k)
      qed
    qed
  qed
  (* not xor but it works anyway *)
  let "?case1  ?case2" = "k-2  0  k+1  card Y -1"

  have b_right: "[f (k-2); f (k-1); b]" if "k  2"
  proof -
    have "k-1 < (k::nat)"
      using k  0 diff_less zero_less_one by blast
    hence "k-2 < k-1"
      using 2  k by linarith
    have "[f (k-2); f (k-1); b]"
      using abd_bcd_abc b_middle f_def k_def(2) fin_Y k-2 < k-1 k-1 < k thm2_ind2 chain_defs
      by (metis Suc_1 Suc_le_lessD diff_Suc_eq_diff_pred that zero_less_diff)
    thus "[f (k-2); f (k-1); b]"
      using [f(k - 1); b; f k] abd_bcd_abc
      by blast
  qed

  have b_left: "[b; f k; f (k+1)]" if "k+1  card Y -1"
  proof -
    have "[f (k-1); f k; f (k+1)]"
      using k  0 f_def fin_Y order_finite_chain that
      by auto
    thus "[b; f k; f (k+1)]"
      using [f (k - 1); b; f k] abc_acd_bcd
      by blast
  qed

  have "local_ordering g betw ?X"
  proof -
    have "n. (finite ?X  n < card ?X)  g n  ?X"
    proof (clarify)
      fix n assume "finite ?X  n < card ?X" "g n  Y"
      consider "nk-1" | "nk+1" | "n=k"
        by linarith
      thus "g n = b"
      proof (cases)
        assume "n  k - 1"
        thus "g n = b"
          using f_def k_def(2) Y_def(1) chain_defs local_ordering_def g_def
          by (metis g n  Y k  0 diff_less le_less less_one less_trans not_le)
      next
        assume "k + 1  n"
        show "g n = b"
        proof -
          have "f n  Y  ¬(n < card Y)" for n
            using chain_defs by (metis local_ordering_def f_def)
          then show "g n = b"
            using finite ?X  n < card ?X fin_Y g_def Y_def g n  Y k + 1  n
              not_less not_less_simps(1) not_one_le_zero
            by fastforce
        qed
      next
        assume "n=k"
        thus "g n = b"
          using Y_def k  0 g_def
          by auto
      qed
    qed
    moreover have "x?X. n. (finite ?X  n < card ?X)  g n = x"
    proof
      fix x assume "x?X"
      show "n. (finite ?X  n < card ?X)  g n = x"
      proof (cases)
        assume "xY"
        show ?thesis
        proof -
          obtain ix where "f ix = x" "ix < card Y"
            using  x  Y f_def fin_Y
            unfolding chain_defs local_ordering_def
            by auto
          have "ixk-1  ixk"
            by linarith
          thus ?thesis
          proof
            assume "ixk-1"
            hence "g ix = x"
              using f ix = x g_def by auto
            moreover have "finite ?X  ix < card ?X"
              using Y_def ix < card Y by auto
            ultimately show ?thesis by metis
          next assume "ixk"
            hence "g (ix+1) = x"
              using f ix = x g_def by auto
            moreover have "finite ?X  ix+1 < card ?X"
              using Y_def ix < card Y by auto
            ultimately show ?thesis by metis
          qed
        qed
      next assume "xY"
        hence "x=b"
          using Y_def x  ?X by blast
        thus ?thesis
          using Y_def k  0 k_def(2) ordered_cancel_comm_monoid_diff_class.le_diff_conv2 g_def
          by auto
      qed
    qed
    moreover have "n n' n''. (finite ?X  n'' < card ?X)  Suc n = n'  Suc n' = n''
           [g n; g (Suc n); g (Suc (Suc n))]"
    proof (clarify)
      fix n n' n'' assume  a: "(finite ?X  (Suc (Suc n)) < card ?X)"
      
      text ‹Introduce the two-case splits used later.›
      have  cases_sn: "Suc nk-1  Suc n=k" if "nk-1"
        using k  0 that by linarith
      have cases_ssn: "Suc(Suc n)k-1  Suc(Suc n)=k" if "nk-1" "Suc nk-1"
        using that(2) by linarith

      consider "nk-1" | "nk+1" | "n=k"
        by linarith
      then show "[g n; g (Suc n); g (Suc (Suc n))]"
      proof (cases)
        assume "nk-1" show ?thesis
          using cases_sn
        proof (rule disjE)
          assume "Suc n  k - 1"
          show ?thesis using cases_ssn
          proof (rule disjE)
            show  "n  k - 1" using n  k - 1 by blast
            show Suc n  k - 1 using Suc n  k - 1 by blast
          next
            assume "Suc (Suc n)  k - 1"
            thus ?thesis
              using  Suc n  k - 1 k  0 n  k - 1 ordering_ord_ijk_loc f_def g_def k_def(2)
              by (metis (no_types, lifting) add_diff_inverse_nat less_Suc_eq_le
                less_imp_le_nat less_le_trans less_one local_long_ch_by_ord_def plus_1_eq_Suc)
          next
            assume "Suc (Suc n) = k"
            thus ?thesis
              using b_right g_def by force
          qed
        next
          assume "Suc n = k"
          show ?thesis
            using b_middle Suc n = k n  k - 1 g_def
            by auto
        next show "n  k-1" using n  k - 1 by blast
        qed
      next assume "nk+1" show ?thesis
        proof -
          have "g n = f (n-1)"
            using k + 1  n less_imp_diff_less g_def
            by auto
          moreover have "g (Suc n) = f (n)"
            using k + 1  n g_def by auto
          moreover have "g (Suc (Suc n)) = f (Suc n)"
            using k + 1  n g_def by auto
          moreover have "n-1<n  n<Suc n"
            using k + 1  n by auto
          moreover have "finite Y  Suc n < card Y"
            using Y_def a by auto
          ultimately show ?thesis
            using f_def unfolding chain_defs local_ordering_def
            by (metis k + 1  n add_leD2 le_add_diff_inverse plus_1_eq_Suc)
        qed
      next assume "n=k"
        show ?thesis
          using k  0 n = k b_left g_def Y_def(1) a assms(3) fin_Y
          by auto
      qed
    qed
    ultimately show "local_ordering g betw ?X"
      unfolding local_ordering_def
      by presburger
  qed
  hence "local_long_ch_by_ord g ?X"
    using Y_def f_def local_long_ch_by_ord_def local_long_ch_by_ord_def
    by auto
  thus "[g?X|a1..b..an]"
      using fin_X a1  an  a1  b  b  an bound_indices k_def(2) Y_def g_def chain_defs
      by simp
qed


lemma card4_eq:
  assumes "card X = 4"
  shows "a b c d. a  b  a  c  a  d  b  c  b  d  c  d  X = {a, b, c, d}"
proof -
  obtain a X' where "X = insert a X'" and "a  X'"
    by (metis Suc_eq_numeral assms card_Suc_eq)
  then have "card X' = 3"
    by (metis add_2_eq_Suc' assms card_eq_0_iff card_insert_if diff_Suc_1 finite_insert numeral_3_eq_3 numeral_Bit0 plus_nat.add_0 zero_neq_numeral)
  then obtain b X'' where "X' = insert b X''" and "b  X''"
    by (metis card_Suc_eq numeral_3_eq_3)
  then have "card X'' = 2"
    by (metis Suc_eq_numeral card X' = 3 card.infinite card_insert_if finite_insert pred_numeral_simps(3) zero_neq_numeral)
  then have "c d. c  d  X'' = {c, d}"
    by (meson card_2_iff)
  thus ?thesis
    using X = insert a X' X' = insert b X'' a  X' b  X'' by blast
qed


theorem (*10*) path_finsubset_chain:
  assumes "Q  𝒫"
      and "X  Q"
      and "card X  2"
  shows "ch X"
proof -
  have "finite X"
    using assms(3) not_numeral_le_zero by fastforce
  consider "card X = 2" | "card X = 3" | "card X  4"
    using card X  2 by linarith
  thus ?thesis
  proof (cases)
    assume "card X = 2"
    thus ?thesis
      using finite X assms two_event_chain by blast
  next
    assume "card X = 3"
    thus ?thesis
      using finite X assms three_event_chain by blast
  next
    assume "card X  4"
    thus ?thesis
      using assms(1,2) finite X
    proof (induct "card X - 4" arbitrary: X)
      case 0
      then have "card X = 4"
        by auto
      then have "a b c d. a  b  a  c  a  d  b  c  b  d  c  d  X = {a, b, c, d}"
        using card4_eq by fastforce
      thus ?case
        using "0.prems"(3) assms(1) chain4 by auto
    next
      case IH: (Suc n)

      then obtain Y b where X_eq: "X = insert b Y" and "b  Y"
        by (metis Diff_iff card_eq_0_iff finite.cases insertI1 insert_Diff_single not_numeral_le_zero)
      have "card Y  4" "n = card Y - 4"
        using IH.hyps(2) IH.prems(4) X_eq b  Y by auto
      then have "ch Y"
        using IH(1) [of Y] IH.prems(3,4) X_eq assms(1) by auto

      then obtain f where f_ords: "local_long_ch_by_ord f Y"
        using 4  card Y ch_alt short_ch_card(2) by auto
      then obtain a1 a an where long_ch_Y: "[fY|a1..a..an]"
        using 4  card Y get_fin_long_ch_bounds by fastforce
      hence bound_indices: "f 0 = a1  f (card Y - 1) = an"
        by (simp add: chain_defs)
      have "a1  an  a1  b  b  an"
        using b  Y abc_abc_neq fin_ch_betw long_ch_Y points_in_long_chain by metis
      moreover have "a1  Q  an  Q  b  Q"
        using IH.prems(3) X_eq long_ch_Y points_in_long_chain by auto
      ultimately consider "[b; a1; an]" | "[a1; an; b]" | "[an; b; a1]"
        using some_betw [of Q b a1 an] Q  𝒫 by blast
      thus "ch X"
      proof (cases)
        (* case (i) *)
        assume "[b; a1; an]"
        have X_eq': "X = Y  {b}"
          using X_eq by auto
        let ?g = "λj. if j  1 then f (j - 1) else b"
        have "[?gX|b..a1..an]"
          using chain_append_at_left_edge IH.prems(4) X_eq' [b; a1; an] b  Y long_ch_Y X_eq
          by presburger
        thus "ch X"
          using chain_defs by auto
      next
        (* case (ii) *)
        assume "[a1; an; b]"
        let ?g = "λj. if j  (card X - 2) then f j else b"
        have "[?gX|a1..an..b]"
          using chain_append_at_right_edge IH.prems(4) X_eq [a1; an; b] b  Y long_ch_Y
          by auto
        thus "ch X" using chain_defs by (meson ch_def)
      next
        (* case (iii) *)
        assume "[an; b; a1]"
        then have "[a1; b; an]"
          by (simp add: abc_sym)
        obtain k where
            k_def: "[a1; b; f k]" "k < card Y" "¬ (k'. 0 < k'  k' < k  [a1; b; f k'])"
          using [a1; b; an] b  Y long_ch_Y smallest_k_ex by blast
        obtain g where "g = (λj::nat. if j  k - 1
                                        then f j
                                        else if j = k
                                          then b else f (j - 1))"
          by simp
        hence "[gX|a1..b..an]"
          using chain_append_inside [of f Y a1 a an b k] IH.prems(4) X_eq
            [a1; b; an] b  Y k_def long_ch_Y
          by auto
        thus "ch X"
          using chain_defs ch_def by auto
      qed
    qed
  qed
qed


lemma path_finsubset_chain2:
  assumes "Q  𝒫" and "X  Q" and "card X  2"
  obtains f a b where "[fX|a..b]"
proof -
  have finX: "finite X"
    by (metis assms(3) card.infinite rel_simps(28))
  have ch_X: "ch X"
    using path_finsubset_chain[OF assms] by blast
  obtain f a b where f_def: "[fX|a..b]" "aX  bX"
    using assms finX ch_X get_fin_long_ch_bounds chain_defs
    by (metis ch_def points_in_chain)
  thus ?thesis
    using that by auto
qed


subsection ‹Theorem 11›


text ‹
  Notice this case is so simple, it doesn't even require the path density larger sets of segments
  rely on for fixing their cardinality.
›

lemma (*for 11*) segmentation_ex_N2:
  assumes path_P: "P𝒫"
      and Q_def: "finite (Q::'a set)" "card Q = N" "QP" "N=2"
      and f_def: "[fQ|a..b]"
      and S_def: "S = {segment a b}"
      and P1_def: "P1 = prolongation b a"
      and P2_def: "P2 = prolongation a b"
    shows "P = ((S)  P1  P2  Q) 
           card S = (N-1)  (xS. is_segment x) 
           P1P2={}  (xS. (xP1={}  xP2={}  (yS. xy  xy={})))"
proof -
  have "aQ  bQ  ab"
    using chain_defs f_def points_in_chain first_neq_last
    by (metis)
  hence "Q={a,b}"
    using assms(3,5)
    by (smt card_2_iff insert_absorb insert_commute insert_iff singleton_insert_inj_eq)
  have "aP  bP"
    using Q={a,b} assms(4) by auto
  have "ab" using Q={a,b}
    using N = 2 assms(3) by force
  obtain s where s_def: "s = segment a b" by simp
  let ?S = "{s}"
  have "P = (({s})  P1  P2  Q) 
          card {s} = (N-1)  (x{s}. is_segment x) 
          P1P2={}  (x{s}. (xP1={}  xP2={}  (y{s}. xy  xy={})))"
  proof (rule conjI)
    { fix x assume "xP"
      have "[a;x;b]  [b;a;x]  [a;b;x]  x=a  x=b"
        using aP  bP some_betw path_P ab
        by (meson x  P abc_sym)
      then have "xs  xP1  xP2  x=a  x=b"
        using pro_betw seg_betw P1_def P2_def s_def Q = {a, b}
        by auto
      hence "x  ({s})  P1  P2  Q"
        using Q = {a, b} by auto
    } moreover {
      fix x assume "x  ({s})  P1  P2  Q"
      hence "xs  xP1  xP2  x=a  x=b"
        using Q = {a, b} by blast
      hence "[a;x;b]  [b;a;x]  [a;b;x]  x=a  x=b"
        using s_def P1_def P2_def
        unfolding segment_def prolongation_def
        by auto
      hence "xP"
        using a  P  b  P a  b betw_b_in_path betw_c_in_path path_P
        by blast
    }
    ultimately show union_P: "P = (({s})  P1  P2  Q)"
      by blast
    show "card {s} = (N-1)  (x{s}. is_segment x)  P1P2={} 
          (x{s}. (xP1={}  xP2={}  (y{s}. xy  xy={})))"
    proof (safe)
      show "card {s} = N - 1"
        using Q = {a, b} a  b assms(3) by auto
      show "is_segment s"
        using s_def by blast
      show "x. x  P1  x  P2  x  {}"
      proof -
        fix x assume "xP1" "xP2"
        show "x{}"
          using P1_def P2_def x  P1 x  P2 abc_only_cba pro_betw
          by metis
      qed
      show "x xa. xa  s  xa  P1  xa  {}"
      proof -
        fix x xa assume "xas" "xaP1"
        show "xa{}"
          using abc_only_cba seg_betw pro_betw  P1_def xa  P1 xa  s s_def
          by (metis)
      qed
      show "x xa. xa  s  xa  P2  xa  {}"
      proof -
        fix x xa assume "xas" "xaP2"
        show "xa{}"
          using abc_only_cba seg_betw pro_betw
          by (metis P2_def xa  P2 xa  s s_def)
      qed
    qed
  qed
  thus ?thesis
    by (simp add: S_def s_def)
qed



lemma int_split_to_segs:
  assumes f_def: "[fQ|a..b..c]"
  fixes S defines S_def: "S  {segment (f i) (f(i+1)) | i. i<card Q-1}"
  shows "interval a c = (S)  Q"
proof
  let ?N = "card Q"
  have f_def_2: "aQ  bQ  cQ"
    using f_def points_in_long_chain by blast
  hence "?N  3"
    using f_def long_ch_card_ge3 chain_defs
    by (meson finite_long_chain_with_card)
  have bound_indices: "f 0 = a  f (card Q - 1) = c"
    using f_def chain_defs by auto
  let "?i = ?u" = "interval a c = (S)  Q"
  show "?i?u"
  proof
    fix p assume "p  ?i"
    show "p?u"
    proof (cases)
      assume "pQ" thus ?thesis by blast
    next assume "pQ"
      hence "pa  pc"
        using f_def f_def_2 by blast
      hence "[a;p;c]"
        using seg_betw p  interval a c interval_def
        by auto
      then obtain ny nz y z
        where yz_def: "y=f ny" "z=f nz" "[y;p;z]" "yQ" "zQ" "ny=nz-1" "nz<card Q"
          "¬(k < card Q. [f k; p; c]  k>ny)" "¬(k<nz. [a; p; f k])"
        using get_closest_chain_events [where f=f and x=p and Y=Q and an=c and a0=a and a=b]
          f_def pQ
        by metis
      have "ny<card Q-1"
        using yz_def(6,7) f_def index_middle_element
        by fastforce
      let ?s = "segment (f ny) (f nz)"
      have "p?s"
        using [y;p;z] abc_abc_neq seg_betw yz_def(1,2)
        by blast
      have "nz = ny + 1"
        using yz_def(6)
        by (metis abc_abc_neq add.commute add_diff_inverse_nat less_one yz_def(1,2,3) zero_diff)
      hence "?sS"
        using S_def ny<card Q-1 assms(2)
        by blast
      hence "pS"
        using p  ?s by blast
      thus ?thesis by blast
    qed
  qed
  show "?u?i"
  proof
    fix p assume "p  ?u"
    hence "pS  pQ" by blast
    thus "p?i"
    proof
      assume "pQ"
      then consider "p=a"|"p=c"|"[a;p;c]"
        using f_def by (meson fin_ch_betw2 finite_long_chain_with_alt)
      thus ?thesis
      proof (cases)
        assume "p=a"
        thus ?thesis by (simp add: interval_def)
      next assume "p=c"
        thus ?thesis by (simp add: interval_def)
      next assume "[a;p;c]"
        thus ?thesis using interval_def seg_betw by auto
      qed
    next assume "pS"
      then obtain s where "ps" "sS"
        by blast
      then obtain y where "s = segment (f y) (f (y+1))" "y<?N-1"
        using S_def by blast
      hence "y+1<?N" by (simp add: assms(2))
      hence fy_in_Q: "(f y)Q  f (y+1)  Q"
        using f_def add_lessD1 unfolding chain_defs local_ordering_def
        by (metis One_nat_def Suc_eq_plus1 Zero_not_Suc 3card Q card_1_singleton_iff card_gt_0_iff
          card_insert_if diff_add_inverse2 diff_is_0_eq' less_numeral_extra(1) numeral_3_eq_3 plus_1_eq_Suc)
      have "[a; f y; c]  y=0"
        using y <?N - 1 assms(2) f_def chain_defs order_finite_chain by auto
      moreover have "[a; f (y+1); c]  y = ?N-2"
        using y+1 < card Q assms(2) f_def chain_defs order_finite_chain i_le_j_events_neq
        using indices_neq_imp_events_neq fin_ch_betw2 fy_in_Q
        by (smt (z3) Nat.add_0_right Nat.add_diff_assoc add_gr_0 card_Diff1_less card_Diff_singleton_if
          diff_diff_left diff_is_0_eq' le_numeral_extra(4) less_numeral_extra(1) nat_1_add_1)
      ultimately consider "y=0"|"y=?N-2"|"([a; f y; c]  [a; f (y+1); c])"
        by linarith
      hence "[a;p;c]"
      proof (cases)
        assume "y=0"
        hence "f y = a"
          by (simp add: bound_indices)
        hence "[a; p; (f(y+1))]"
          using p  s s = segment (f y) (f (y + 1)) seg_betw
          by auto
        moreover have "[a; (f(y+1)); c]"
          using [a; (f(y+1)); c]  y = ?N - 2 y = 0 ?N3
          by linarith
        ultimately show "[a;p;c]"
          using abc_acd_abd by blast
      next
        assume "y=?N-2"
        hence "f (y+1) = c"
          using bound_indices ?N3 numeral_2_eq_2 numeral_3_eq_3
          by (metis One_nat_def Suc_diff_le add.commute add_leD2 diff_Suc_Suc plus_1_eq_Suc)
        hence "[f y; p; c]"
          using p  s s = segment (f y) (f (y + 1)) seg_betw
          by auto
        moreover have "[a; f y; c]"
          using [a; f y; c]  y = 0 y = ?N - 2 ?N3
          by linarith
        ultimately show "[a;p;c]"
          by (meson abc_acd_abd abc_sym)
      next
        assume "[a; f y; c]  [a; (f(y+1)); c]"
        thus "[a;p;c]"
          using abe_ade_bcd_ace [where a=a and b="f y" and d="f (y+1)" and e=c and c=p]
          using p  s s = segment (f y) (f(y+1)) seg_betw
          by auto
      qed
      thus ?thesis
        using interval_def seg_betw by auto
    qed
  qed
qed


lemma (*for 11*) path_is_union:
  assumes path_P: "P𝒫"
      and Q_def: "finite (Q::'a set)" "card Q = N" "QP" "N3"
      and f_def: "aQ  bQ  cQ" "[fQ|a..b..c]"
      and S_def: "S = {s. i<(N-1). s = segment (f i) (f (i+1))}"
      and P1_def: "P1 = prolongation b a"
      and P2_def: "P2 = prolongation b c"
    shows "P = ((S)  P1  P2  Q)"
proof -
  (* For future use, as always *)
  have in_P: "aP  bP  cP"
    using assms(4) f_def by blast
  have bound_indices: "f 0 = a  f (card Q - 1) = c"
    using f_def chain_defs by auto
  have points_neq: "ab  bc  ac"
    using f_def chain_defs by (metis first_neq_last)

  text ‹The proof in two parts: subset inclusion one way, then the other.›
  { fix x assume "xP"
    have "[a;x;c]  [b;a;x]  [b;c;x]  x=a  x=c"
      using in_P some_betw path_P points_neq x  P abc_sym
      by (metis (full_types) abc_acd_bcd fin_ch_betw f_def(2))
    then have "(sS. xs)  xP1  xP2  xQ"
    proof (cases)
      assume "[a;x;c]"
      hence only_axc: "¬([b;a;x]  [b;c;x]  x=a  x=c)"
        using abc_only_cba
        by (meson abc_bcd_abd abc_sym f_def fin_ch_betw)
      have "x  interval a c"
        using [a;x;c] interval_def seg_betw by auto
      hence "xQ  xS"
        using int_split_to_segs S_def assms(2,3,5) f_def
        by blast
      thus ?thesis by blast
    next assume "¬[a;x;c]"
      hence "[b;a;x]  [b;c;x]  x=a  x=c"
        using [a;x;c]  [b;a;x]  [b;c;x]  x = a  x = c by blast
      hence " xP1  xP2  xQ"
        using P1_def P2_def f_def pro_betw by auto
      thus ?thesis by blast
    qed
    hence "x  (S)  P1  P2  Q" by blast
  } moreover {
    fix x assume "x  (S)  P1  P2  Q"
    hence "(sS. xs)  xP1  xP2  xQ"
      by blast
    hence "xS  [b;a;x]  [b;c;x]  xQ"
      using S_def P1_def P2_def
      unfolding segment_def prolongation_def
      by auto
    hence "xP"
    proof (cases)
      assume "xS"
      have "S = {segment (f i) (f(i+1)) | i. i<N-1}"
        using S_def by blast
      hence "xinterval a c"
        using int_split_to_segs [OF f_def(2)] assms xS
        by (simp add: UnCI)
      hence "[a;x;c]  x=a  x=c"
        using interval_def seg_betw by auto
      thus ?thesis
      proof (rule disjE)
        assume "x=a  x=c"
        thus ?thesis
          using in_P by blast
      next
        assume "[a;x;c]"
        thus ?thesis
          using betw_b_in_path in_P path_P points_neq by blast
      qed
     next assume "xS"
      hence "[b;a;x]  [b;c;x]  xQ"
        using x   S  [b;a;x]  [b;c;x]  x  Q
        by blast
      thus ?thesis
        using assms(4) betw_c_in_path in_P path_P points_neq
        by blast
    qed
  }
  ultimately show "P = ((S)  P1  P2  Q)"
    by blast
qed


lemma (*for 11*) inseg_axc:
  assumes path_P: "P𝒫"
      and Q_def: "finite (Q::'a set)" "card Q = N" "QP" "N3"
      and f_def: "aQ  bQ  cQ" "[fQ|a..b..c]"
      and S_def: "S = {s. i<(N-1). s = segment (f i) (f (i+1))}"
      and x_def: "xs" "sS"
    shows "[a;x;c]"
proof -
  have fQ: "local_long_ch_by_ord f Q"
    using f_def Q_def chain_defs by (metis ch_long_if_card_geq3 path_P short_ch_card(1) short_xor_long(2))
  have inseg_neq_ac: "xa  xc" if "xs" "sS" for x s
  proof
    show "xa"
    proof (rule notI)
      assume "x=a"
      obtain n where s_def: "s = segment (f n) (f (n+1))" "n<N-1"
        using S_def s  S by blast
      hence "n<card Q" using assms(3) by linarith
      hence "f n  Q"
        using fQ unfolding chain_defs local_ordering_def by blast 
      hence "[a; f n; c]"
        using f_def finite_long_chain_with_def assms(3) order_finite_chain seg_betw that(1)
        using n < N - 1 s = segment (f n) (f (n + 1)) x = a
        by (metis abc_abc_neq add_lessD1 fin_ch_betw inside_not_bound(2) less_diff_conv)
      moreover have "[(f(n)); x; (f(n+1))]"
        using xs seg_betw s_def(1) by simp
      ultimately show False
        using x=a abc_only_cba(1) assms(3) fQ chain_defs s_def(2)
        by (smt (z3) n < card Q f_def(2) order_finite_chain_indices2 thm2_ind1)
    qed

    show "xc"
    proof (rule notI)
      assume "x=c"
      obtain n where s_def: "s = segment (f n) (f (n+1))" "n<N-1"
        using S_def s  S by blast
      hence "n+1<N" by simp
      have "[(f(n)); x; (f(n+1))]"
        using xs seg_betw s_def(1) by simp
      have "f (n)  Q"
        using fQ n+1 < N chain_defs local_ordering_def
        by (metis add_lessD1 assms(3))
      have "f (n+1)  Q"
        using n+1 < N fQ chain_defs local_ordering_def
        by (metis assms(3))
      have "f(n+1)  c"
        using x=c [(f(n)); x; (f(n+1))] abc_abc_neq
        by blast
      hence "[a; (f(n+1)); c]"
        using f_def finite_long_chain_with_def assms(3) order_finite_chain seg_betw that(1)
          abc_abc_neq abc_only_cba fin_ch_betw
        by (metis [f n; x; f (n + 1)] f (n + 1)  Q f n  Q x = c)
      thus False
        using x=c [(f(n)); x; (f(n+1))] assms(3) f_def s_def(2)
          abc_only_cba(1) finite_long_chain_with_def order_finite_chain
        by (metis f n  Q abc_bcd_acd abc_only_cba(1,2) fin_ch_betw)
    qed
  qed

  show "[a;x;c]"
  proof -
    have "xinterval a c"
      using int_split_to_segs [OF f_def(2)] S_def assms(2,3,5) x_def
      by blast
    have "xa  xc" using inseg_neq_ac
      using x_def by auto
    thus ?thesis
      using seg_betw x  interval a c interval_def
      by auto
  qed
qed


lemma disjoint_segmentation:
  assumes path_P: "P𝒫"
      and Q_def: "finite (Q::'a set)" "card Q = N" "QP" "N3"
      and f_def: "aQ  bQ  cQ" "[fQ|a..b..c]"
      and S_def: "S = {s. i<(N-1). s = segment (f i) (f (i+1))}"
      and P1_def: "P1 = prolongation b a"
      and P2_def: "P2 = prolongation b c"
    shows "P1P2={}  (xS. (xP1={}  xP2={}  (yS. xy  xy={})))"
proof (rule conjI)
  have fQ: "local_long_ch_by_ord f Q"
    using f_def Q_def chain_defs by (metis ch_long_if_card_geq3 path_P short_ch_card(1) short_xor_long(2))
  show "P1  P2 = {}"
  proof (safe)
    fix x assume "xP1" "xP2"
    show "x{}"
      using abc_only_cba pro_betw P1_def P2_def
      by (metis x  P1 x  P2 abc_bcd_abd f_def(2) fin_ch_betw)
  qed

  show "xS. (xP1={}  xP2={}  (yS. xy  xy={}))"
  proof (rule ballI)
    fix s assume "sS"
    show "s  P1 = {}  s  P2 = {}  (yS. s  y  s  y = {})"
    proof (intro conjI ballI impI)
      show "sP1={}"
      proof (safe)
        fix x assume "xs" "xP1"
        hence "[a;x;c]"
          using inseg_axc s  S assms by blast
        thus "x{}"
          by (metis P1_def x  P1 abc_bcd_abd abc_only_cba(1) f_def(2) fin_ch_betw pro_betw)
      qed
      show "sP2={}"
      proof (safe)
        fix x assume "xs" "xP2"
        hence "[a;x;c]"
          using inseg_axc s  S assms by blast
        thus "x{}"
          by (metis P2_def x  P2 abc_bcd_acd abc_only_cba(2) f_def(2) fin_ch_betw pro_betw)
      qed
      fix r assume "rS" "sr"
      show "sr={}"
      proof (safe)
        fix y assume "y  r" "y  s"
        obtain n m where rs_def: "r = segment (f n) (f(n+1))" "s = segment (f m) (f(m+1))"
                                 "nm" "n<N-1" "m<N-1"
          using S_def r  S s  r s  S by blast
        have y_betw: "[f n; y; (f(n+1))]  [f m; y; (f(m+1))]"
          using seg_betw yr ys rs_def(1,2) by simp
        have False
        proof (cases)
          assume "n<m"
          have "[f n; f m; (f(m+1))]"
            using n < m assms(3) fQ chain_defs order_finite_chain rs_def(5) by (metis assms(2) thm2_ind1)
          have "n+1<m"
            using [f n; f m; f(m + 1)] n < m abc_only_cba(2) abd_bcd_abc y_betw
            by (metis Suc_eq_plus1 Suc_leI le_eq_less_or_eq)
          hence "[f n; (f(n+1)); f m]"
            using fQ assms(3) rs_def(5) unfolding chain_defs local_ordering_def
            by (metis (full_types) [f n;f m;f (m + 1)] abc_only_cba(1) abc_sym abd_bcd_abc assms(2) fQ thm2_ind1 y_betw)
          hence "[f n; (f(n+1)); y]"
            using [f n; f m; f(m + 1)] abc_acd_abd abd_bcd_abc y_betw
            by blast
          thus ?thesis
            using abc_only_cba y_betw by blast
        next
          assume "¬n<m"
          hence "n>m" using nat_neq_iff rs_def(3) by blast
          have "[f m; f n; (f(n+1))]"
            using n > m assms(3) fQ chain_defs rs_def(4) by (metis assms(2) thm2_ind1)
          hence "m+1<n"
            using  n > m abc_only_cba(2) abd_bcd_abc y_betw
            by (metis Suc_eq_plus1 Suc_leI le_eq_less_or_eq)
          hence "[f m; (f(m+1)); f n]"
            using fQ assms(2,3) rs_def(4) unfolding chain_defs local_ordering_def
            by (metis (no_types, lifting) [f m;f n;f (n + 1)] abc_only_cba(1) abc_sym abd_bcd_abc fQ thm2_ind1 y_betw)
          hence "[f m; (f(m+1)); y]"
            using [f m; f n; f(n + 1)] abc_acd_abd abd_bcd_abc y_betw
            by blast
          thus ?thesis
            using abc_only_cba y_betw by blast
        qed
        thus "y{}" by blast
      qed
    qed
  qed
qed


lemma (*for 11*) segmentation_ex_Nge3:
  assumes path_P: "P𝒫"
      and Q_def: "finite (Q::'a set)" "card Q = N" "QP" "N3"
      and f_def: "aQ  bQ  cQ" "[fQ|a..b..c]"
      and S_def: "S = {s. i<(N-1). s = segment (f i) (f (i+1))}"
      and P1_def: "P1 = prolongation b a"
      and P2_def: "P2 = prolongation b c"
    shows "P = ((S)  P1  P2  Q) 
           (xS. is_segment x) 
           P1P2={}  (xS. (xP1={}  xP2={}  (yS. xy  xy={})))"
proof (intro disjoint_segmentation conjI)
  show "P = ((S)  P1  P2  Q)"
    using path_is_union assms
    by blast
  show "xS. is_segment x"
  proof
    fix s assume "sS"
    thus "is_segment s" using S_def by auto
  qed
qed (use assms disjoint_segmentation in auto)


text ‹Some unfolding of the definition for a finite chain that happens to be short.›
lemma finite_chain_with_card_2:
  assumes f_def: "[fQ|a..b]"
    and card_Q: "card Q = 2"
  shows "finite Q" "f 0 = a" "f (card Q - 1) = b" "Q = {f 0, f 1}" "Q. path Q (f 0) (f 1)"
    using assms unfolding chain_defs by auto


text ‹
  Schutz says "As in the proof of the previous theorem [...]" - does he mean to imply that this
  should really be proved as induction? I can see that quite easily, induct on $N$, and add a segment
  by either splitting up a segment or taking a piece out of a prolongation.
  But I think that might be too much trouble.
›

theorem (*11*) show_segmentation:  
  assumes path_P: "P𝒫"
      and Q_def: "QP"
      and f_def: "[fQ|a..b]"
    fixes P1 defines P1_def: "P1  prolongation b a"
    fixes P2 defines P2_def: "P2  prolongation a b"
    fixes S  defines S_def: "S  {segment (f i) (f (i+1)) | i. i<card Q-1}"
    shows "P = ((S)  P1  P2  Q)" "(xS. is_segment x)"
          "disjoint (S{P1,P2})" "P1P2" "P1S" "P2S"
proof -
  have card_Q: "card Q  2"
    using fin_chain_card_geq_2 f_def by blast
  have "finite Q"
    by (metis card.infinite card_Q rel_simps(28))
  have f_def_2: "aQ  bQ"
    using f_def points_in_chain finite_chain_with_def by auto
  have "ab"
    using f_def chain_defs by (metis first_neq_last)
  {
    assume "card Q = 2"
    hence "card Q - 1 = Suc 0" by simp
    have "Q = {f 0, f 1}" "Q. path Q (f 0) (f 1)" "f 0 = a" "f (card Q - 1) = b"
      using card Q = 2 finite_chain_with_card_2 f_def by auto
    hence "S={segment a b}"
      unfolding S_def using card Q - 1 = Suc 0 by (simp add: eval_nat_numeral)
    hence "P = ((S)  P1  P2  Q)" "(xS. is_segment x)" "P1P2={}"
         "(xS. (xP1={}  xP2={}  (yS. xy  xy={})))"
      using assms f_def finite Q segmentation_ex_N2
        [where P=P and Q=Q and N="card Q"]
      by (metis (no_types, lifting) card Q = 2)+
  } moreover {
    assume "card Q  2"
    hence "card Q  3"
      using card_Q by auto
    then obtain c where c_def: "[fQ|a..c..b]"
      using assms(3,5) ab chain_defs
      by (metis f_def three_in_set3)
    have pro_equiv: "P1 = prolongation c a  P2 = prolongation c b"
      using pro_basis_change
      using P1_def P2_def abc_sym c_def fin_ch_betw by auto
    have S_def2: "S = {s. i<(card Q-1). s = segment (f i) (f (i+1))}"
      using S_def card Q  3 by auto
    have "P = ((S)  P1  P2  Q)" "(xS. is_segment x)" "P1P2={}"
         "(xS. (xP1={}  xP2={}  (yS. xy  xy={})))"
      using f_def_2 assms f_def card Q  3 c_def pro_equiv
        segmentation_ex_Nge3 [where P=P and Q=Q and N="card Q" and S=S and a=a and b=c and c=b and f=f]
      using points_in_long_chain finite Q S_def2 by metis+
  }
  ultimately have old_thesis: "P = ((S)  P1  P2  Q)" "(xS. is_segment x)" "P1P2={}"
                  "(xS. (xP1={}  xP2={}  (yS. xy  xy={})))" by meson+
  thus "disjoint (S{P1,P2})" "P1P2" "P1S" "P2S"
       "P = ((S)  P1  P2  Q)" "(xS. is_segment x)"
         unfolding disjoint_def apply (simp add: Int_commute)
        apply (metis P2_def Un_iff old_thesis(1,3) a  b disjoint_iff f_def_2 path_P pro_betw prolong_betw2)
       apply (metis P1_def Un_iff old_thesis(1,4) a  b disjoint_iff f_def_2 path_P pro_betw prolong_betw3)
      apply (metis P2_def Un_iff old_thesis(1,4) a  b disjoint_iff f_def_2 path_P pro_betw prolong_betw)
    using old_thesis(1,2) by linarith+
qed


theorem (*11*) segmentation:
  assumes path_P: "P𝒫"
      and Q_def: "card Q2" "QP"
    shows "S P1 P2. P = ((S)  P1  P2  Q) 
                     disjoint (S{P1,P2})  P1P2  P1S  P2S 
                     (xS. is_segment x)  is_prolongation P1  is_prolongation P2"
proof -
  let ?N = "card Q"
(* Hooray for theorem 10! Without it, we couldn't so brazenly go from a set of events
   to an ordered chain of events. *)
  obtain f a b where f_def: "[fQ|a..b]"
    using path_finsubset_chain2[OF path_P Q_def(2,1)]
    by metis
  let ?S = "{segment (f i) (f (i+1)) | i. i<card Q-1}"
  let ?P1 = "prolongation b a"
  let ?P2 = "prolongation a b"
  have from_seg: "P = ((?S)  ?P1  ?P2  Q)" "(x?S. is_segment x)"
          "disjoint (?S{?P1,?P2})" "?P1?P2" "?P1?S" "?P2?S"
    using show_segmentation[OF path_P Q_def(2) [fQ|a..b]]
    by force+
  thus ?thesis
    by blast
qed



end (* context MinkowskiSpacetime *)


section "Chains are unique up to reversal"
context MinkowskiSpacetime begin

lemma chain_remove_at_right_edge:
  assumes "[fX|a..c]" "f (card X - 2) = p" "3  card X" "X = insert c Y" "cY"
  shows "[fY|a..p]"
proof -
  have lch_X: "local_long_ch_by_ord f X"
    using assms(1,3) chain_defs short_ch_card_2
    by fastforce
  have "pX"
    by (metis local_ordering_def assms(2) card.empty card_gt_0_iff diff_less lch_X
        local_long_ch_by_ord_def not_numeral_le_zero zero_less_numeral)
  have bound_ind: "f 0 = a  f (card X - 1) = c"
    using lch_X assms(1,3) unfolding finite_chain_with_def finite_long_chain_with_def
    by metis

  have "[a;p;c]"
  proof -
    have "card X - 2 < card X - 1"
      using 3  card X by auto
    moreover have "card X - 2 > 0"
      using 3  card X by linarith
    ultimately show ?thesis
      using order_finite_chain[OF lch_X] 3  card X assms(2) bound_ind
      by (metis card.infinite diff_less le_numeral_extra(3) less_numeral_extra(1) not_gr_zero not_numeral_le_zero)
  qed

  have "[fX|a..p..c]"
    unfolding finite_long_chain_with_alt by (simp add: assms(1) [a;p;c] pX)

  have 1: "x = a" if "x  Y" "¬ [a;x;p]" "x  p" for x
  proof -
    have "xX"
      using that(1) assms(4) by simp
    hence 01: "x=a  [a;p;x]"
      by (metis that(2,3) [a;p;c] abd_acd_abcacb assms(1) fin_ch_betw2)
    have 02: "x=c" if "[a;p;x]"
    proof -
      obtain i where i_def: "f i = x" "i<card X"
        using xX chain_defs by (meson assms(1) obtain_index_fin_chain)
      have "f 0 = a"
        by (simp add: bound_ind)
      have "card X - 2 < i"
        using order_finite_chain_indices[OF lch_X _ that f 0 = a assms(2) i_def(1) _ _ i_def(2)]
        by (metis card_eq_0_iff card_gt_0_iff diff_less i_def(2) less_nat_zero_code zero_less_numeral)
      hence "i = card X - 1" using i_def(2) by linarith
      thus ?thesis using bound_ind i_def(1) by blast
    qed
    show ?thesis using 01 02 assms(5) that(1) by auto
  qed

  have "Y = {e  X. [a;e;p]  e = a  e = p}"
    apply (safe, simp_all add: assms(4) 1)
    using [a;p;c] abc_only_cba(2) abc_abc_neq assms(4) by blast+

  thus ?thesis using chain_shortening[OF [fX|a..p..c]] by simp
qed


lemma (in MinkowskiChain) fin_long_ch_imp_fin_ch:
  assumes "[fX|a..b..c]"
  shows "[fX|a..c]"
  using assms by (simp add: finite_long_chain_with_alt)


text ‹
  If we ever want to have chains less strongly identified by endpoints,
  this result should generalise - $a,c,x,z$ are only used to identify reversal/no-reversal cases.
›
lemma chain_unique_induction_ax:
  assumes "card X  3"
      and "i < card X"
      and "[fX|a..c]"
      and "[gX|x..z]"
      and "a = x  c = z"
    shows "f i = g i"
using assms
proof (induct "card X - 3" arbitrary: X a c x z)
  case Nil: 0
  have "card X = 3"
    using Nil.hyps Nil.prems(1) by auto

  obtain b where f_ch: "[fX|a..b..c]"
    using chain_defs by (metis Nil.prems(1,3) three_in_set3)
  obtain y where g_ch: "[gX|x..y..z]"
    using Nil.prems chain_defs by (metis three_in_set3)

  have "i=1  i=0  i=2"
    using card X = 3 Nil.prems(2) by linarith
  thus ?case
  proof (rule disjE)
    assume "i=1"
    hence "f i = b  g i = y"
      using index_middle_element f_ch g_ch card X = 3 numeral_3_eq_3
      by (metis One_nat_def add_diff_cancel_left' less_SucE not_less_eq plus_1_eq_Suc)
    have "f i = g i"
    proof (rule ccontr)
      assume "f i  g i"
      hence "g i  b"
        by (simp add: f i = b  g i = y)
      have "g i  X"
        using f i = b  g i = y  g_ch points_in_long_chain by blast
      have "X = {a,b,c}"
        using f_ch unfolding finite_long_chain_with_alt
        using card X = 3 points_in_long_chain[OF f_ch] abc_abc_neq[of a b c]
        by (simp add: card_3_eq'(2))
      hence "(g i = a  g i = c)"
        using g i  b g i  X by blast
      hence "¬ [a; g i; c]"
        using abc_abc_neq by blast
      hence "g i  X"
        using f i=b  g i=y g i=a  g i=c  f_ch g_ch chain_bounds_unique finite_long_chain_with_def
        by blast
      thus False
        by (simp add: g i  X)
    qed
    thus ?thesis
      by (simp add: card X = 3 i = 1)
  next
    assume "i = 0  i = 2"
    show ?thesis
      using Nil.prems(5) card X = 3 i = 0  i = 2 chain_bounds_unique f_ch g_ch chain_defs
      by (metis diff_Suc_1 numeral_2_eq_2 numeral_3_eq_3)
  qed
next
  case IH: (Suc n)
  have lch_fX: "local_long_ch_by_ord f X"
    using chain_defs long_ch_card_ge3 IH(3,5)
    by fastforce
  have lch_gX: "local_long_ch_by_ord g X"
    using IH(3,6) chain_defs long_ch_card_ge3
    by fastforce
  have fin_X: "finite X"
    using IH(4) le_0_eq by fastforce

  have "ch_by_ord f X"
    using lch_fX unfolding ch_by_ord_def by blast
  have "card X  4"
    using IH.hyps(2) by linarith

  obtain b where f_ch: "[fX|a..b..c]"
    using IH(3,5) chain_defs by (metis three_in_set3)
  obtain y where g_ch: "[gX|x..y..z]"
    using IH.prems(1,4) chain_defs by (metis three_in_set3)

  obtain p where p_def: "p = f (card X - 2)" by simp
  have "[a;p;c]"
  proof -
    have "card X - 2 < card X - 1"
      using 4  card X by auto
    moreover have "card X - 2 > 0"
      using 3  card X by linarith
    ultimately show ?thesis
      using f_ch p_def chain_defs [fX] order_finite_chain2 by force
  qed
  hence "pc  pa"
    using abc_abc_neq by blast

  obtain Y where Y_def: "X = insert c Y" "cY"
    using f_ch points_in_long_chain
    by (meson mk_disjoint_insert)
  hence fin_Y: "finite Y"
    using f_ch chain_defs by auto
  hence "n = card Y - 3"
    using Suc n = card X - 3 X = insert c Y cY card_insert_if
    by auto
  hence card_Y: "card Y = n + 3"
    using Y_def(1) Y_def(2) fin_Y IH.hyps(2) by fastforce 
  have "card Y = card X - 1"
    using Y_def(1,2) fin_X by auto
  have "pY"
    using X = insert c Y [a;p;c] abc_abc_neq lch_fX p_def IH.prems(1,3) Y_def(2)
    by (metis chain_remove_at_right_edge points_in_chain)
  have "[fY|a..p]"
    using chain_remove_at_right_edge [where f=f and a=a and c=c and X=X and p=p and Y=Y]
    using fin_long_ch_imp_fin_ch  [where f=f and a=a and c=c and b=b and X=X]
    using f_ch p_def card X  3 Y_def
    by blast
  hence ch_fY: "local_long_ch_by_ord f Y"
    using card_Y fin_Y chain_defs long_ch_card_ge3
    by force

  have p_closest: "¬ (qX. [p;q;c])"
  proof
    assume "(qX. [p;q;c])"
    then obtain q where "qX" "[p;q;c]" by blast
    then obtain j where "j < card X" "f j = q"
      using lch_fX lch_gX fin_X points_in_chain pc  pa chain_defs
      by (metis local_ordering_def)
    have "j > card X - 2  j < card X - 1"
    proof -
      have "j > card X - 2  j < card X - 1  j > card X - 1  j < card X - 2"
        apply (intro order_finite_chain_indices[OF lch_fX finite X [p;q;c]])
        using p_def f j = q IH.prems(3) finite_chain_with_def j < card X by auto
      thus ?thesis by linarith
    qed
    thus False by linarith
  qed

  have "g (card X - 2) = p"
  proof (rule ccontr)
    assume asm_false: "g (card X - 2)  p"
    obtain j where "g j = p" "j < card X - 1" "j>0"
      using X = insert c Y pY  points_in_chain pc  pa
      by (metis (no_types) chain_bounds_unique f_ch
          finite_long_chain_with_def g_ch index_middle_element insert_iff)
    hence "j < card X - 2"
      using asm_false le_eq_less_or_eq by fastforce
    hence "j < card Y - 1"
      by (simp add: Y_def(1,2) fin_Y)
    obtain d where "d = g (card X - 2)" by simp
    have "[p;d;z]"
    proof -
      have "card X - 1 > card X - 2"
        using j < card X - 1 by linarith
      thus ?thesis
        using lch_gX j < card Y - 1 card Y = card X - 1 d = g (card X - 2) g j = p
          order_finite_chain[OF lch_gX] chain_defs local_ordering_def
          by (smt (z3) IH.prems(3-5) asm_false chain_bounds_unique chain_remove_at_right_edge p_def
            thesis. (Y. X = insert c Y; c  Y  thesis)  thesis)
    qed
    moreover have "dX"
      using lch_gX d = g (card X - 2) unfolding local_long_ch_by_ord_def local_ordering_def
      by auto
    ultimately show False
      using p_closest abc_sym IH.prems(3-5) chain_bounds_unique f_ch g_ch
      by blast
  qed

  hence ch_gY: "local_long_ch_by_ord g Y"
    using IH.prems(1,4,5) g_ch f_ch ch_fY card_Y chain_remove_at_right_edge fin_Y chain_defs
    by (metis Y_def chain_bounds_unique long_ch_card_ge3)
  
  have "f i  Y  f i = c"
    by (metis local_ordering_def X = insert c Y i < card X lch_fX insert_iff local_long_ch_by_ord_def)
  thus "f i = g i"
  proof (rule disjE)
    assume "f i  Y"
    hence "f i  c"
      using c  Y by blast
    hence "i < card Y"
      using X = insert c Y cY IH(3,4) f_ch fin_Y chain_defs not_less_less_Suc_eq
      by (metis card Y = card X - 1 card_insert_disjoint)
    hence "3  card Y"
      using card_Y le_add2 by presburger
    show "f i = g i"
      using IH(1) [of Y]
      using n = card Y - 3 3  card Y i < card Y
      using Y_def card_Y chain_remove_at_right_edge le_add2
      by (metis IH.prems(1,3,4,5) chain_bounds_unique)
  next
    assume "f i = c"
    show ?thesis
      using IH.prems(2,5) f i = c chain_bounds_unique f_ch g_ch indices_neq_imp_events_neq chain_defs
      by (metis card Y = card X - 1 Y_def card_insert_disjoint fin_Y lessI)
  qed
qed

text ‹I'm really impressed sledgehammer›/smt› can solve this if I just tell them "Use symmetry!".›

lemma chain_unique_induction_cx:
  assumes "card X  3"
      and "i < card X"
      and "[fX|a..c]"
      and "[gX|x..z]"
      and "c = x  a = z"
    shows "f i = g (card X - i - 1)"
  using chain_sym_obtain2 chain_unique_induction_ax assms diff_right_commute by smt

text ‹
  This lemma has to exclude two-element chains again, because no order exists within them.
  Alternatively, the result is trivial: any function that assigns one element to index 0 and
  the other to 1 can be replaced with the (unique) other assignment, without destroying any
  (trivial, since ternary) termlocal_ordering of the chain.
  This could be made generic over the termlocal_ordering
  similar to @{thm chain_sym} relying on @{thm ordering_sym_loc}.
›

lemma chain_unique_upto_rev_cases:
  assumes ch_f: "[fX|a..c]"
      and ch_g: "[gX|x..z]"
      and card_X: "card X  3"
      and valid_index: "i < card X"
  shows "((a=x  c=z)  (f i = g i))" "((a=z  c=x)  (f i = g (card X - i - 1)))"
proof -
  obtain n where n_def: "n = card X - 3"
    by blast
  hence valid_index_2: "i < n + 3"
    by (simp add: card_X valid_index)

  show "((a=x  c=z)  (f i = g i))"
    using card_X ch_f ch_g chain_unique_induction_ax valid_index by blast
  show "((a=z  c=x)  (f i = g (card X - i - 1)))"
    using assms(3) ch_f ch_g chain_unique_induction_cx valid_index by blast
qed

lemma chain_unique_upto_rev:
  assumes "[fX|a..c]" "[gX|x..z]" "card X  3" "i < card X"
  shows "f i = g i  f i = g (card X - i - 1)" "a=xc=z  c=xa=z"
proof -
  have "(a=x  c=z)  (a=z  c=x)"
    using chain_bounds_unique by (metis assms(1,2))
  thus "f i = g i  f i = g (card X - i - 1)"
    using assms(3) i < card X assms chain_unique_upto_rev_cases by blast
  thus "(a=xc=z)  (c=xa=z)"
    by (meson assms(1-3) chain_bounds_unique)
    qed


end (* context MinkowskiSpacetime *)


section "Interlude: betw4 and WLOG"

subsection "betw4 - strict and non-strict, basic lemmas"
context MinkowskiBetweenness begin

text ‹Define additional notation for non-strict termlocal_ordering -
  cf Schutz' monograph cite‹ p.~27› in "schutz1997".›

abbreviation nonstrict_betw_right :: "'a  'a  'a  bool" ("[_;_;_") where
  "nonstrict_betw_right a b c  [a;b;c]  b = c"

abbreviation nonstrict_betw_left :: "'a  'a  'a  bool" ("_;_;_]") where
  "nonstrict_betw_left a b c  [a;b;c]  b = a"

abbreviation nonstrict_betw_both :: "'a  'a  'a  bool" (* ("[(_ _ _)]") *) where
  "nonstrict_betw_both a b c  nonstrict_betw_left a b c  nonstrict_betw_right a b c"

abbreviation betw4 :: "'a  'a  'a  'a  bool" ("[_;_;_;_]") where
  "betw4 a b c d  [a;b;c]  [b;c;d]"

abbreviation nonstrict_betw_right4 :: "'a  'a  'a  'a  bool" ("[_;_;_;_") where
  "nonstrict_betw_right4 a b c d  betw4 a b c d  c = d"

abbreviation nonstrict_betw_left4 :: "'a  'a  'a  'a  bool" ("_;_;_;_]") where
  "nonstrict_betw_left4 a b c d  betw4 a b c d  a = b"

abbreviation nonstrict_betw_both4 :: "'a  'a  'a  'a  bool" (* ("[(_ _ _ _)]") *) where
  "nonstrict_betw_both4 a b c d  nonstrict_betw_left4 a b c d  nonstrict_betw_right4 a b c d"

lemma betw4_strong:
  assumes "betw4 a b c d"
  shows "[a;b;d]  [a;c;d]"
  using abc_bcd_acd assms by blast

lemma betw4_imp_neq:
  assumes "betw4 a b c d"
  shows "ab  ac  ad  bc  bd  cd"
  using abc_only_cba assms by blast


end (* context MinkowskiBetweenness *)
context MinkowskiSpacetime begin


lemma betw4_weak:
  fixes a b c d :: 'a
  assumes "[a;b;c]  [a;c;d]
           [a;b;c]  [b;c;d]
           [a;b;d]  [b;c;d]
           [a;b;d]  [b;c;d]"
  shows "betw4 a b c d"
  using abc_acd_bcd abd_bcd_abc assms by blast

lemma betw4_sym:
  fixes a::'a and b::'a and c::'a and d::'a
  shows "betw4 a b c d  betw4 d c b a"
  using abc_sym by blast

lemma abcd_dcba_only:
  fixes a::'a and b::'a and c::'a and d::'a
  assumes "[a;b;c;d]"
  shows "¬[a;b;d;c]" "¬[a;c;b;d]" "¬[a;c;d;b]" "¬[a;d;b;c]" "¬[a;d;c;b]"
        "¬[b;a;c;d]" "¬[b;a;d;c]" "¬[b;c;a;d]" "¬[b;c;d;a]" "¬[b;d;c;a]" "¬[b;d;a;c]"
        "¬[c;a;b;d]" "¬[c;a;d;b]" "¬[c;b;a;d]" "¬[c;b;d;a]" "¬[c;d;a;b]" "¬[c;d;b;a]"
        "¬[d;a;b;c]" "¬[d;a;c;b]" "¬[d;b;a;c]" "¬[d;b;c;a]" "¬[d;c;a;b]"
  using abc_only_cba assms by blast+ 

lemma some_betw4a:
  fixes a::'a and b::'a and c::'a and d::'a and P
  assumes "P𝒫" "aP" "bP" "cP" "dP" "ab  ac  ad  bc  bd  cd"
      and "¬([a;b;c;d]  [a;b;d;c]  [a;c;b;d]  [a;c;d;b]  [a;d;b;c]  [a;d;c;b])"
    shows "[b;a;c;d]  [b;a;d;c]  [b;c;a;d]  [b;d;a;c]  [c;a;b;d]  [c;b;a;d]"
  by (smt abc_bcd_acd abc_sym abd_bcd_abc assms some_betw_xor)

lemma some_betw4b:
  fixes a::'a and b::'a and c::'a and d::'a and P
  assumes "P𝒫" "aP" "bP" "cP" "dP" "ab  ac  ad  bc  bd  cd"
      and "¬([b;a;c;d]  [b;a;d;c]  [b;c;a;d]  [b;d;a;c]  [c;a;b;d]  [c;b;a;d])"
    shows "[a;b;c;d]  [a;b;d;c]  [a;c;b;d]  [a;c;d;b]  [a;d;b;c]  [a;d;c;b]"
  by (smt abc_bcd_acd abc_sym abd_bcd_abc assms some_betw_xor)


lemma abd_acd_abcdacbd:
  fixes a::'a and b::'a and c::'a and d::'a
  assumes abd: "[a;b;d]" and acd: "[a;c;d]" and "bc"
  shows "[a;b;c;d]  [a;c;b;d]"
proof -
  obtain P where "P𝒫" "aP" "bP" "dP"
    using abc_ex_path abd by blast
  have "cP"
    using P  𝒫 a  P d  P abc_abc_neq acd betw_b_in_path by blast
  have "¬[b;d;c]"
    using abc_sym abcd_dcba_only(5) abd acd by blast
  hence "[b;c;d]  [c;b;d]"
    using abc_abc_neq abc_sym abd acd assms(3) some_betw
    by (metis P  𝒫 b  P c  P d  P)
  thus ?thesis
    using abd acd betw4_weak by blast
qed

end (*context MinkowskiSpacetime*)

subsection "WLOG for two general symmetric relations of two elements on a single path"
context MinkowskiBetweenness begin

text ‹
  This first one is really just trying to get a hang of how to write these things.
  If you have a relation that does not care which way round the ``endpoints'' (if $Q$ is the
  interval-relation) go, then anything you want to prove about both undistinguished endpoints,
  follows from a proof involving a single endpoint.
›

lemma wlog_sym_element:
  assumes symmetric_rel: "a b I. Q I a b  Q I b a"
      and one_endpoint: "a b x I. Q I a b; x=a  P x I"
    shows other_endpoint: "a b x I. Q I a b; x=b  P x I"
  using assms by fastforce

text ‹
  This one gives the most pertinent case split: a proof involving e.g. an element of an interval
  must consider the edge case and the inside case.
›
lemma wlog_element:
  assumes symmetric_rel: "a b I. Q I a b  Q I b a"
      and one_endpoint: "a b x I. Q I a b; x=a  P x I"
      and neither_endpoint: "a b x I. Q I a b; xI; (xa  xb)  P x I"
    shows any_element: "x I. xI; (a b. Q I a b)  P x I"
  by (metis assms)

text ‹
  Summary of the two above. Use for early case splitting in proofs.
  Doesn't need $P$ to be symmetric - the context in the conclusion is explicitly symmetric.
›

lemma wlog_two_sets_element:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and case_split: "a b c d x I J. Q I a b; Q J c d 
              (x=a  x=c  P x I J)  (¬(x=a  x=b  x=c  x=d)  P x I J)"
    shows "x I J. a b. Q I a b; a b. Q J a b  P x I J"
  by (smt case_split symmetric_Q)

text ‹
  Now we start on the actual result of interest. First we assume the events are all distinct,
  and we deal with the degenerate possibilities after.
›

lemma wlog_endpoints_distinct1:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and "I J a b c d. Q I a b; Q J c d; [a;b;c;d]  P I J"
    shows "I J a b c d. Q I a b; Q J c d;
              [b;a;c;d]  [a;b;d;c]  [b;a;d;c]  [d;c;b;a]  P I J"
  by (meson abc_sym assms(2) symmetric_Q)

lemma wlog_endpoints_distinct2:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and "I J a b c d. Q I a b; Q J c d; [a;c;b;d]  P I J"
    shows "I J a b c d. Q I a b; Q J c d;
              [b;c;a;d]  [a;d;b;c]  [b;d;a;c]  [d;b;c;a]  P I J"
  by (meson abc_sym assms(2) symmetric_Q)

lemma wlog_endpoints_distinct3:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and symmetric_P: "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d. Q I a b; Q J c d; [a;c;d;b]  P I J"
    shows "I J a b c d. Q I a b; Q J c d;
              [a;d;c;b]  [b;c;d;a]  [b;d;c;a]  [c;a;b;d]  P I J"
  by (meson assms)

lemma (in MinkowskiSpacetime) wlog_endpoints_distinct4:
    fixes Q:: "('a set)  'a  'a  bool" (* cf ‹I = interval a b› *)
      and P:: "('a set)  ('a set)  bool"
      and A:: "('a set)" (* the path that takes the role of the real line *)
  assumes path_A: "A𝒫"
      and symmetric_Q: "a b I. Q I a b  Q I b a"
      and Q_implies_path: "a b I. IA; Q I a b  bA  aA"
      and symmetric_P: "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d.
          Q I a b; Q J c d; IA; JA; [a;b;c;d]  [a;c;b;d]  [a;c;d;b]  P I J"
    shows "I J a b c d. Q I a b; Q J c d; IA; JA;
                ab  ac  ad  bc  bd  cd  P I J"
proof -
  fix I J a b c d
  assume asm: "Q I a b" "Q J c d" "I  A" "J  A"
              "ab  ac  ad  bc  bd  cd"
  have endpoints_on_path: "aA" "bA" "cA" "dA"
    using Q_implies_path asm by blast+
  show "P I J"
  proof (cases) (* have to split like this, because the full ‹some_betw› is too large for Isabelle *)
    assume "[b;a;c;d]  [b;a;d;c]  [b;c;a;d] 
            [b;d;a;c]  [c;a;b;d]  [c;b;a;d]"
    then consider "[b;a;c;d]"|"[b;a;d;c]"|"[b;c;a;d]"|
                  "[b;d;a;c]"|"[c;a;b;d]"|"[c;b;a;d]"
      by linarith
    thus "P I J"
      apply (cases)
           apply (metis(mono_tags) asm(1-4) assms(5) symmetric_Q)+
       apply (metis asm(1-4) assms(4,5))
      by (metis asm(1-4) assms(2,4,5) symmetric_Q)
  next
    assume "¬([b;a;c;d]  [b;a;d;c]  [b;c;a;d] 
              [b;d;a;c]  [c;a;b;d]  [c;b;a;d])"
    hence "[a;b;c;d]  [a;b;d;c]  [a;c;b;d] 
           [a;c;d;b]  [a;d;b;c]  [a;d;c;b]"
      using some_betw4b [where P=A and a=a and b=b and c=c and d=d]
      using endpoints_on_path asm path_A by simp
    then consider "[a;b;c;d]"|"[a;b;d;c]"|"[a;c;b;d]"|
                  "[a;c;d;b]"|"[a;d;b;c]"|"[a;d;c;b]"
      by linarith
    thus "P I J"
      apply (cases)
      by (metis asm(1-4) assms(5) symmetric_Q)+
  qed
qed


lemma (in MinkowskiSpacetime) wlog_endpoints_distinct':
  assumes "A  𝒫"
      and "a b I. Q I a b  Q I b a"
      and "a b I. I  A; Q I a b  a  A"
      and "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d.
          Q I a b; Q J c d; IA; JA; betw4 a b c d  betw4 a c b d  betw4 a c d b  P I J"
      and "Q I a b"
      and "Q J c d"
      and "I  A"
      and "J  A"
      and "a  b" "a  c" "a  d" "b  c" "b  d" "c  d"
  shows "P I J"
proof -
  { 
    let ?R = "(λI. (a b. Q I a b))"
    have "I J. ?R I; ?R J; P I J  P J I"
      using assms(4) by blast
  }
  thus ?thesis
    using wlog_endpoints_distinct4
      [where P=P and Q=Q and A=A and I=I and J=J and a=a and b=b and c=c and d=d]
    by (smt assms(1-3,5-))
qed

lemma (in MinkowskiSpacetime) wlog_endpoints_distinct:
  assumes path_A: "A𝒫"
      and symmetric_Q: "a b I. Q I a b  Q I b a"
      and Q_implies_path: "a b I. IA; Q I a b  bA  aA"
      and symmetric_P: "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d.
          Q I a b; Q J c d; IA; JA; [a;b;c;d]  [a;c;b;d]  [a;c;d;b]  P I J"
  shows "I J a b c d. Q I a b; Q J c d; IA; JA;
              ab  ac  ad  bc  bd  cd  P I J"
  by (smt (verit, ccfv_SIG) assms some_betw4b)


lemma wlog_endpoints_degenerate1:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and symmetric_P: "I J. a b. Q I a b; a b. Q I a b; P I J  P J I"
          (* two singleton intervals *)
      and two: "I J a b c d. Q I a b; Q J c d;
                  (a=b  b=c  c=d)  (a=b  bc  c=d)  P I J"
          (* one singleton interval *)
      and one: "I J a b c d. Q I a b; Q J c d;
                  (a=b  b=c  cd)  (a=b  bc  cd  ad)  P I J"
          (* no singleton interval - the all-distinct case is a separate theorem *)
      and no: "I J a b c d. Q I a b; Q J c d;
                  (ab  bc  cd  a=d)  (ab  b=c  cd  a=d)  P I J"
    shows "I J a b c d. Q I a b; Q J c d; ¬(ab  bc  cd  ad  ac  bd)  P I J"
  by (metis assms)

lemma wlog_endpoints_degenerate2:
  assumes symmetric_Q: "a b I. Q I a b  Q I b a"
      and Q_implies_path: "a b I A. IA; A𝒫; Q I a b  bA  aA"
      and symmetric_P: "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d A. Q I a b; Q J c d; IA; JA; A𝒫;
              [a;b;c]  a=d  P I J"
      and "I J a b c d A. Q I a b; Q J c d; IA; JA; A𝒫;
              [b;a;c]  a=d  P I J"
    shows "I J a b c d A. Q I a b; Q J c d; IA; JA; A𝒫;
              ab  bc  cd  a=d  P I J"
proof -
  have last_case: "I J a b c d A. Q I a b; Q J c d; IA; JA; A𝒫;
              [b;c;a]  a=d  P I J"
    using assms(1,3-5) by (metis abc_sym)
  thus "I J a b c d A. Q I a b; Q J c d; IA; JA; A𝒫;
              ab  bc  cd  a=d  P I J"
    by (smt (z3) abc_sym assms(2,4,5) some_betw)
qed


lemma wlog_endpoints_degenerate:
  assumes path_A: "A𝒫"
      and symmetric_Q: "a b I. Q I a b  Q I b a"
      and Q_implies_path: "a b I. IA; Q I a b  bA  aA"
      and symmetric_P: "I J. a b. Q I a b; a b. Q J a b; P I J  P J I"
      and "I J a b c d. Q I a b; Q J c d; IA; JA
             ((a=b  b=c  c=d)  P I J)  ((a=b  bc  c=d)  P I J)
               ((a=b  b=c  cd)  P I J)  ((a=b  bc  cd  ad)  P I J)
               ((ab  b=c  cd  a=d)  P I J)
               (([a;b;c]  a=d)  P I J)  (([b;a;c]  a=d)  P I J)"
    shows "I J a b c d. Q I a b; Q J c d; IA; JA;
            ¬(ab  bc  cd  ad  ac  bd)  P I J"
proof -

  text ‹We first extract some of the assumptions of this lemma into the form
  of other WLOG lemmas' assumptions.›
  have ord1: "I J a b c d. Q I a b; Q J c d; IA; JA;
              [a;b;c]  a=d  P I J"
    using assms(5) by auto
  have ord2: "I J a b c d. Q I a b; Q J c d; IA; JA;
              [b;a;c]  a=d  P I J"
    using assms(5) by auto
  have last_case: "I J a b c d. Q I a b; Q J c d; IA; JA;
              ab  bc  cd  a=d  P I J"
    using ord1 ord2 wlog_endpoints_degenerate2 symmetric_P symmetric_Q Q_implies_path path_A
    by (metis abc_sym some_betw)
  show "I J a b c d. Q I a b; Q J c d; IA; JA;
            ¬(ab  bc  cd  ad  ac  bd)  P I J"
  proof -

    text ‹Fix the sets on the path, and obtain the assumptions of wlog_endpoints_degenerate1›.›
    fix I J
    assume asm1: "IA" "JA"
    have two: "a b c d. Q I a b; Q J c d; a=b  b=c  c=d  P I J"
            "a b c d. Q I a b; Q J c d; a=b  bc  c=d  P I J"
      using J  A I  A path_A assms(5) by blast+ 
    have one: " a b c d. Q I a b; Q J c d; a=b  b=c  cd  P I J"
          " a b c d. Q I a b; Q J c d; a=b  bc  cd  ad  P I J"
      using I  A J  A path_A assms(5) by blast+ 
    have no: " a b c d. Q I a b; Q J c d; ab  bc  cd  a=d  P I J"
          " a b c d. Q I a b; Q J c d; ab  b=c  cd  a=d  P I J"
      using I  A J  A path_A last_case apply blast
      using I  A J  A path_A assms(5) by auto

    text ‹Now unwrap the remaining object logic and finish the proof.›
    fix a b c d
    assume asm2: "Q I a b" "Q J c d" "¬(ab  bc  cd  ad  ac  bd)"
    show "P I J"
      using two [where a=a and b=b and c=c and d=d]
      using one [where a=a and b=b and c=c and d=d]
      using no [where a=a and b=b and c=c and d=d]
      using wlog_endpoints_degenerate1
        [where I=I and J=J and a=a and b=b and c=c and d=d and P=P and Q=Q]
      using asm1 asm2 symmetric_P last_case assms(5) symmetric_Q
      by smt
  qed
qed


lemma (in MinkowskiSpacetime) wlog_intro:
  assumes path_A: "A𝒫"
      and symmetric_Q: "a b I. Q I a b  Q I b a"
      and Q_implies_path: "a b I. IA; Q I a b  bA  aA"
      and symmetric_P: "I J. a b. Q I a b; c d. Q J c d; P I J  P J I"
      and essential_cases: "I J a b c d. Q I a b; Q J c d; IA; JA
             ((a=b  b=c  c=d)  P I J)
               ((a=b  bc  c=d)  P I J)
               ((a=b  b=c  cd)  P I J)
               ((a=b  bc  cd  ad)  P I J)
               ((ab  b=c  cd  a=d)  P I J)
               (([a;b;c]  a=d)  P I J)
               (([b;a;c]  a=d)  P I J)
               ([a;b;c;d]  P I J)
               ([a;c;b;d]  P I J)
               ([a;c;d;b]  P I J)"
      and antecedants: "Q I a b" "Q J c d" "IA" "JA"
  shows "P I J"
    using essential_cases antecedants
    and wlog_endpoints_degenerate[OF path_A symmetric_Q Q_implies_path symmetric_P]
    and wlog_endpoints_distinct[OF path_A symmetric_Q Q_implies_path symmetric_P]
    by (smt (z3) Q_implies_path path_A symmetric_P symmetric_Q some_betw2 some_betw4b abc_only_cba(1))


end (*context MinkowskiSpacetime*)


subsection "WLOG for two intervals"
context MinkowskiBetweenness begin

text ‹
  This section just specifies the results for a generic relation $Q$
  in the previous section to the interval relation.
›

lemma wlog_two_interval_element:
  assumes "x I J. is_interval I; is_interval J; P x J I  P x I J"
      and "a b c d x I J. I = interval a b; J = interval c d 
              (x=a  x=c  P x I J)  (¬(x=a  x=b  x=c  x=d)  P x I J)"
    shows "x I J. is_interval I; is_interval J  P x I J"
  by (metis assms(2) int_sym)


lemma (in MinkowskiSpacetime) wlog_interval_endpoints_distinct:
  assumes "I J. is_interval I; is_interval J; P I J  P J I" (*P does not distinguish between intervals*)
          "I J a b c d. I = interval a b; J = interval c d
           ([a;b;c;d]  P I J)  ([a;c;b;d]  P I J)  ([a;c;d;b]  P I J)"
  shows "I J Q a b c d. I = interval a b; J = interval c d; IQ; JQ; Q𝒫;
              ab  ac  ad  bc  bd  cd  P I J"
proof -
  let ?Q = "λ I a b. I = interval a b"

  fix I J A a b c d
  assume asm: "?Q I a b" "?Q J c d" "IA" "JA" "A𝒫" "ab  ac  ad  bc  bd  cd"
  show "P I J"
  proof (rule wlog_endpoints_distinct)
    show "a b I. ?Q I a b  ?Q I b a"
      by (simp add: int_sym)
    show "a b I. I  A  ?Q I a b  b  A  a  A"
      by (simp add: ends_in_int subset_iff)
    show "I J. is_interval I  is_interval J  P I J  P J I"
      using assms(1) by blast
    show "I J a b c d. ?Q I a b; ?Q J c d; [a;b;c;d]  [a;c;b;d]  [a;c;d;b]
         P I J"
      by (meson assms(2))
    show "I = interval a b" "J = interval c d" "IA" "JA" "A𝒫"
        "ab  ac  ad  bc  bd  cd"
      using asm by simp+
  qed
qed


lemma wlog_interval_endpoints_degenerate:
  assumes symmetry: "I J. is_interval I; is_interval J; P I J  P J I"
      and "I J a b c d Q. I = interval a b; J = interval c d; IQ; JQ; Q𝒫
             ((a=b  b=c  c=d)  P I J)  ((a=b  bc  c=d)  P I J)
               ((a=b  b=c  cd)  P I J)  ((a=b  bc  cd  ad)  P I J)
               ((ab  b=c  cd  a=d)  P I J)
               (([a;b;c]  a=d)  P I J)  (([b;a;c]  a=d)  P I J)"
    shows "I J a b c d Q. I = interval a b; J = interval c d; IQ; JQ; Q𝒫;
            ¬(ab  bc  cd  ad  ac  bd)  P I J"
proof -
  let ?Q = "λ I a b. I = interval a b"

  fix I J a b c d A
  assume asm: "?Q I a b" "?Q J c d" "IA" "JA" "A𝒫" "¬(ab  bc  cd  ad  ac  bd)"
  show "P I J"
  proof (rule wlog_endpoints_degenerate)
    show "a b I. ?Q I a b  ?Q I b a"
      by (simp add: int_sym)
    show "a b I. I  A  ?Q I a b  b  A  a  A"
      by (simp add: ends_in_int subset_iff)
    show "I J. is_interval I  is_interval J  P I J  P J I"
      using symmetry by blast
    show "I = interval a b" "J = interval c d" "IA" "JA" "A𝒫"
      "¬ (ab  bc  cd  ad  ac  bd)"
      using asm by auto+ 
    show "I J a b c d. ?Q I a b; ?Q J c d; I  A; J  A 
        (a = b  b = c  c = d  P I J) 
        (a = b  b  c  c = d  P I J) 
        (a = b  b = c  c  d  P I J) 
        (a = b  b  c  c  d  a  d  P I J) 
        (a  b  b = c  c  d  a = d  P I J) 
        ([a;b;c]  a = d  P I J)  ([b;a;c]  a = d  P I J)"
      using assms(2) A𝒫 by auto
  qed
qed

end (* context MinkowskiBetweenness *)


section "Interlude: Intervals, Segments, Connectedness"
context MinkowskiSpacetime begin

text ‹
  In this secion, we apply the WLOG lemmas from the previous section in order to reduce the
  number of cases we need to consider when thinking about two arbitrary intervals on a path.
  This is used to prove that the (countable) intersection of intervals is an interval.
  These results cannot be found in Schutz, but he does use them (without justification)
  in his proof of Theorem 12 (even for uncountable intersections).
›

lemma int_of_ints_is_interval_neq: (* Proof using WLOG *)
  assumes  "I1 = interval a b" "I2 = interval c d" "I1P" "I2P" "P𝒫" "I1I2  {}"
      and events_neq: "ab" "ac" "ad" "bc" "bd" "cd"
    shows "is_interval (I1  I2)"
proof -
  have on_path: "aP  bP  cP  dP"
    using assms(1-4) interval_def by auto

  let ?prop = "λ I J. is_interval (IJ)  (IJ) = {}" (* The empty intersection is excluded in assms. *)

  have symmetry: "(I J. is_interval I  is_interval J  ?prop I J  ?prop J I)"
    by (simp add: Int_commute)

  {
    fix I J a b c d
    assume "I = interval a b" "J = interval c d"
    have "([a;b;c;d]  ?prop I J)"
         "([a;c;b;d]  ?prop I J)"
         "([a;c;d;b]  ?prop I J)"
    proof (rule_tac [!] impI)
      assume "betw4 a b c d"
      have "IJ = {}"
      proof (rule ccontr)
        assume "IJ{}"
        then obtain x where "xIJ"
          by blast
        show False
        proof (cases)
          assume "xa  xb  xc  xd"
          hence "[a;x;b]" "[c;x;d]"
            using I=interval a b xIJ J=interval c d xIJ
            by (simp add: interval_def seg_betw)+
          thus False
            by (meson betw4 a b c d abc_only_cba(3) abc_sym abd_bcd_abc)
        next
          assume "¬(xa  xb  xc  xd)"
          thus False
            using interval_def seg_betw I = interval a b J = interval c d abcd_dcba_only(21)
                 x  I  J betw4 a b c d abc_bcd_abd abc_bcd_acd abc_only_cba(1,2)
            by (metis (full_types) insert_iff Int_iff)
        qed
      qed 
      thus "?prop I J" by simp
    next
      assume "[a;c;b;d]"
      then have "a  b  a  c  a  d  b  c  b  d  c  d"
        using betw4_imp_neq by blast
      have "IJ = interval c b"
      proof (safe)
        fix x
        assume "x  interval c b"
        {
          assume "x=b  x=c"
          hence "xI"
            using [a;c;b;d] I = interval a b interval_def seg_betw by auto
          have "xJ"
            using x=b  x=c
            using [a;c;b;d] J = interval c d interval_def seg_betw by auto 
          hence "xI  xJ" using x  I by blast
        } moreover {
          assume  "¬(x=b  x=c)"
          hence "[c;x;b]"
            using xinterval c b unfolding interval_def segment_def by simp
          hence "[a;x;b]"
            by (meson [a;c;b;d] abc_acd_abd abc_sym)
          have "[c;x;d]"
            using [a;c;b;d] [c;x;b] abc_acd_abd by blast
          have "xI" "xJ"
            using I = interval a b [a;x;b] J = interval c d [c;x;d] 
                  interval_def seg_betw by auto
        }
        ultimately show "xI" "xJ" by blast+
      next
        fix x
        assume "xI" "xJ"
        show "x  interval c b"
        proof (cases)
          assume not_eq: "xa  xb  xc  xd"
          have "[a;x;b]" "[c;x;d]"
            using xI I = interval a b  xJ J = interval c d 
                  not_eq unfolding interval_def segment_def by blast+
          hence "[c;x;b]"
            by (meson [a;c;b;d] abc_bcd_acd betw4_weak)
          thus ?thesis
            unfolding interval_def segment_def using seg_betw segment_def by auto
        next
          assume not_not_eq: "¬(xa  xb  xc  xd)"
          {
            assume "x=a"
            have "¬[d;a;c]"
              using [a;c;b;d] abcd_dcba_only(9) by blast
            hence "a  interval c d" unfolding interval_def segment_def
              using abc_sym a  b  a  c  a  d  b  c  b  d  c  d by blast
            hence "False" using xJ J = interval c d x=a by blast
          } moreover {
            assume "x=d"
            have "¬[a;d;b]" using betw4 a c b d abc_sym abcd_dcba_only(9) by blast
            hence "dinterval a b" unfolding interval_def segment_def
              using a  b  a  c  a  d  b  c  b  d  c  d by blast
            hence "False" using xI x=d I = interval a b by blast
          }
          ultimately show ?thesis
            using interval_def not_not_eq by auto
        qed
      qed
      thus "?prop I J" by auto
    next
      assume "[a;c;d;b]"
      have "IJ = interval c d"
      proof (safe)
        fix x
        assume "x  interval c d"
        {
          assume "xc  xd"
          have "x  J"
            by (simp add: J = interval c d x  interval c d)
          have "[c;x;d]"
            using x  interval c d x  c  x  d interval_def seg_betw by auto
          have "[a;x;b]"
            by (meson betw4 a c d b [c;x;d] abc_bcd_abd abc_sym abe_ade_bcd_ace)
          have "x  I"
            using I = interval a b [a;x;b] interval_def seg_betw by auto
          hence "xI  xJ" by (simp add: x  J)
        } moreover {
          assume "¬ (xc  xd)"
          hence "xI  xJ"
            by (metis I = interval a b J = interval c d [a;c;d;b] x  interval c d
                abc_bcd_abd abc_bcd_acd insertI2 interval_def seg_betw)
        }
        ultimately show "xI" "xJ" by blast+
      next
        fix x
        assume "xI" "xJ"
        show "x  interval c d"
          using J = interval c d x  J by auto
      qed
      thus "?prop I J" by auto
    qed
  }
    
  then show "is_interval (I1I2)"
    using wlog_interval_endpoints_distinct
      [where P="?prop" and I=I1 and J=I2 and Q=P and a=a and b=b and c=c and d=d]
    using symmetry assms by simp
qed


lemma int_of_ints_is_interval_deg: (* Proof using WLOG *)
  assumes  "I = interval a b" "J = interval c d" "IJ  {}" "IP" "JP" "P𝒫"
      and events_deg: "¬(ab  bc  cd  ad  ac  bd)"
    shows "is_interval (I  J)"
proof -

  let ?p = "λ I J. (is_interval (I  J)  IJ = {})"

  have symmetry: "I J. is_interval I; is_interval J; ?p I J  ?p J I"
    by (simp add: inf_commute)

  have degen_cases: "I J a b c d Q. I = interval a b; J = interval c d; IQ; JQ; Q𝒫
             ((a=b  b=c  c=d)  ?p I J)  ((a=b  bc  c=d)  ?p I J)
               ((a=b  b=c  cd)  ?p I J)  ((a=b  bc  cd  ad)  ?p I J)
               ((ab  b=c  cd  a=d)  ?p I J)
               (([a;b;c]  a=d)  ?p I J)  (([b;a;c]  a=d)  ?p I J)"
  proof -
    fix I J a b c d Q
    assume "I = interval a b" "J = interval c d" "IQ" "JQ" "Q𝒫"
    show "((a=b  b=c  c=d)  ?p I J)  ((a=b  bc  c=d)  ?p I J)
               ((a=b  b=c  cd)  ?p I J)  ((a=b  bc  cd  ad)  ?p I J)
               ((ab  b=c  cd  a=d)  ?p I J)
               (([a;b;c]  a=d)  ?p I J)  (([b;a;c]  a=d)  ?p I J)"
    proof (intro conjI impI)
      assume "a = b  b = c  c = d" thus "?p I J"
        using I = interval a b J = interval c d by auto
    next
      assume "a = b  b  c  c = d" thus "?p I J"
        using J = interval c d empty_segment interval_def by auto
    next
      assume "a = b  b = c  c  d" thus "?p I J"
        using I = interval a b empty_segment interval_def by auto
    next
      assume "a = b  b  c  c  d  a  d" thus "?p I J"
        using I = interval a b empty_segment interval_def by auto
    next
      assume "a  b  b = c  c  d  a = d" thus "?p I J"
        using I = interval a b J = interval c d int_sym by auto
    next
      assume "[a;b;c]  a = d" show "?p I J"
      proof (cases)
        assume "IJ = {}" thus ?thesis by simp
      next
        assume "IJ  {}"
        have "IJ = interval a b"
        proof (safe)
          fix x assume "xI" "xJ"
          thus "xinterval a b"
            using I = interval a b by blast
        next
          fix x assume "xinterval a b"
          show "xI"
            by (simp add: I = interval a b x  interval a b)
          have "[d;b;c]"
            using [a;b;c]  a = d by blast
          have "[a;x;b]  x=a  x=b"
            using I = interval a b x  I interval_def seg_betw by auto
          consider "[d;x;c]"|"x=a  x=b"
            using [a;b;c]  a = d [a;x;b]  x = a  x = b abc_acd_abd by blast
          thus "xJ" 
          proof (cases)
            case 1
            then show ?thesis 
              by (simp add: J = interval c d abc_abc_neq abc_sym interval_def seg_betw)
          next
            case 2
            then have "x  interval c d"
              using  [a;b;c]  a = d int_sym interval_def seg_betw 
              by force 
            then show ?thesis
              using J = interval c d by blast  
          qed
        qed
        thus "?p I J" by blast
      qed
    next
      assume "[b;a;c]  a = d" show "?p I J"
      proof (cases)
        assume "IJ = {}" thus ?thesis by simp
      next
        assume "IJ  {}"
        have "IJ = {a}"
        proof (safe)
          fix x assume "xI" "xJ" "x{}"
          have cxd: "[c;x;d]  x=c  x=d"
            using J = interval c d x  J interval_def seg_betw by auto
          consider "[a;x;b]"|"x=a"|"x=b"
            using I = interval a b x  I interval_def seg_betw by auto
          then show "x=a"
          proof (cases)
            assume "[a;x;b]"
            hence "[b;x;d;c]"
              using [b;a;c]  a = d abc_acd_bcd abc_sym by meson
            hence False
              using cxd abc_abc_neq by blast
            thus ?thesis by simp
          next
            assume "x=b"
            hence "[b;d;c]"
              using [b;a;c]  a = d by blast
            hence False
              using cxd x = b abc_abc_neq by blast
            thus ?thesis
              by simp
          next
            assume "x=a" thus "x=a" by simp
          qed
        next
          show "aI"
            by (simp add: I = interval a b ends_in_int)
          show "aJ"
            by (simp add: J = interval c d [b;a;c]  a = d ends_in_int)
        qed
        thus "?p I J"
          by (simp add: empty_segment interval_def)
      qed
    qed
  qed

  have "?p I J"
    using wlog_interval_endpoints_degenerate
      [where P="?p" and I=I and J=J and a=a and b=b and c=c and d=d and Q=P]
    using degen_cases
    using symmetry assms
    by smt

  thus ?thesis
    using assms(3) by blast
qed


lemma int_of_ints_is_interval:
  assumes "is_interval I" "is_interval J" "IP" "JP" "P𝒫" "IJ  {}"
  shows "is_interval (I  J)"
  using int_of_ints_is_interval_neq int_of_ints_is_interval_deg
  by (meson assms)


lemma int_of_ints_is_interval2:
  assumes "xS. (is_interval x  xP)" "P𝒫" "S  {}" "finite S" "S{}"
  shows "is_interval (S)"
proof -
  obtain n where "n = card S"
    by simp
  consider "n=0"|"n=1"|"n2"
    by linarith
  thus ?thesis
  proof (cases)
    assume "n=0"
    then have False
      using n = card S assms(4,5) by simp
    thus ?thesis
      by simp
  next
    assume "n=1"
    then obtain I where "S = {I}"
      using n = card S card_1_singletonE by auto
    then have "S = I"
      by simp
    moreover have "is_interval I"
      by (simp add: S = {I} assms(1))
    ultimately show ?thesis
      by blast
  next
    assume "2n"
    obtain m where "m+2=n"
      using 2  n le_add_diff_inverse2 by blast
    have ind: "S. xS. (is_interval x  xP); P𝒫; S  {}; finite S; S{}; m+2=card S
        is_interval (S)"
    proof (induct m)
      case 0
      then have "card S = 2"
        by auto
      then obtain I J where "S={I,J}" "IJ"
        by (meson card_2_iff)
      then have "IS" "JS"
        by blast+
      then have "is_interval I" "is_interval J" "IP" "JP"
           by (simp add: "0.prems"(1))+ 
      also have "IJ  {}"
        using S={I,J} "0.prems"(3) by force
      then have "is_interval(IJ)"
        using assms(2) calculation int_of_ints_is_interval[where I=I and J=J and P=P]
        by fastforce
      then show ?case
        by (simp add: S = {I, J})
    next
      case (Suc m)
      obtain S' I where "IS" "S = insert I S'" "IS'"
        using Suc.prems(4,5) by (metis Set.set_insert finite.simps insertI1)
      then have "is_interval (S')"
      proof -
        have "m+2 = card S'"
          using Suc.prems(4,6) S = insert I S' IS' by auto
        moreover have "xS'. is_interval x  x  P"
          by (simp add: Suc.prems(1) S = insert I S')
        moreover have " S'  {}"
          using Suc.prems(3) S = insert I S' by auto
        moreover have "finite S'"
          using Suc.prems(4) S = insert I S' by auto
        ultimately show ?thesis
          using assms(2) Suc(1) [where S=S'] by fastforce
      qed
      then have "is_interval ((S')I)"
      proof (rule int_of_ints_is_interval)
        show "is_interval I"
          by (simp add: Suc.prems(1) I  S)
        show "S'  P"
          using I  S' S = insert I S' Suc.prems(1,4,6) Inter_subset
          by (metis Suc_n_not_le_n card.empty card_insert_disjoint finite_insert
              le_add2 numeral_2_eq_2 subset_eq subset_insertI)
        show "I  P"
          by (simp add: Suc.prems(1) I  S)
        show "P  𝒫"
          using assms(2) by auto
        show "S'  I  {}"
          using Suc.prems(3) S = insert I S' by auto
      qed
      thus ?case
        using S = insert I S' by (simp add: inf.commute)
    qed
    then show ?thesis
      using m + 2 = n n = card S assms by blast
  qed
qed


end (*context MinkowskiSpacetime*)



section "3.7 Continuity and the monotonic sequence property"
context MinkowskiSpacetime begin

text ‹
  This section only includes a proof of the first part of Theorem 12, as well as some results
  that would be useful in proving part (ii).
›

theorem (*12(i)*) two_rays:
  assumes path_Q: "Q𝒫"
      and event_a: "aQ"
    shows "R L. (is_ray_on R Q  is_ray_on L Q
                 Q-{a}  (R  L)               ⌦‹events of Q excl. a belong to two rays›
                 (rR. lL. [l;a;r])       ⌦‹a is betw any 'a of one ray and any 'a of the other›
                 (xR. yR. ¬ [x;a;y])     ⌦‹but a is not betw any two events …›
                 (xL. yL. ¬ [x;a;y]))"   ⌦‹… of the same ray›
proof -
  text ‹Schutz here uses Theorem 6, but we don't need it.›
  obtain b where "b" and "bQ" and "ba"
    using event_a ge2_events in_path_event path_Q by blast
  let ?L = "{x. [x;a;b]}"
  let ?R = "{y. [a;y;b]  [a;b;y}"
  have "Q = ?L  {a}  ?R"
  proof -
    have inQ: "xQ. [x;a;b]  x=a  [a;x;b]  [a;b;x"
      by (meson b  Q b  a abc_sym event_a path_Q some_betw)
    show ?thesis
    proof (safe)
      fix x
      assume "x  Q" "x  a" "¬ [x;a;b]" "¬ [a;x;b]" "b  x" 
      then show "[a;b;x]"
        using inQ by blast
    next
      fix x  
      assume "[x;a;b]" 
      then show "x  Q"
        by (simp add: b  Q abc_abc_neq betw_a_in_path event_a path_Q)
    next
      show "a  Q"
        by (simp add: event_a)
    next
      fix x
      assume "[a;x;b]"
      then show  "x  Q"
        by (simp add: b  Q abc_abc_neq betw_b_in_path event_a path_Q)
    next
      fix x
      assume "[a;b;x]"
      then show  "x  Q"
        by (simp add: b  Q abc_abc_neq betw_c_in_path event_a path_Q)
    next
      show "b  Q" using b  Q .
    qed
  qed
  have disjointLR: "?L  ?R = {}"
    using abc_abc_neq abc_only_cba by blast

  have wxyz_ord: "[x;a;y;b  [x;a;b;y
       (([w;x;a]  [x;a;y])  ([x;w;a]  [w;a;y]))
       (([x;a;y]  [a;y;z])  ([x;a;z]  [a;z;y]))"
    if "x?L" "w?L" "y?R" "z?R" "wx" "yz" for x w y z
    using path_finsubset_chain order_finite_chain (* Schutz says: implied by thm 10 & 2 *)
    by (smt abc_abd_bcdbdc abc_bcd_abd abc_sym abd_bcd_abc mem_Collect_eq that) (* impressive, sledgehammer! *)

  obtain x y where "x?L" "y?R"
    by (metis (mono_tags) b  Q b  a abc_sym event_a mem_Collect_eq path_Q prolong_betw2)
  obtain w where "w?L" "wx"
    by (metis b  Q b  a abc_sym event_a mem_Collect_eq path_Q prolong_betw3) 
  obtain z where "z?R" "yz"
    by (metis (mono_tags) b  Q b  a event_a mem_Collect_eq path_Q prolong_betw3)

  have "is_ray_on ?R Q 
          is_ray_on ?L Q 
          Q - {a}  ?R  ?L 
          (r?R. l?L. [l;a;r]) 
          (x?R. y?R. ¬ [x;a;y]) 
          (x?L. y?L. ¬ [x;a;y])"
  proof (intro conjI)
    show "is_ray_on ?L Q"
    proof (unfold is_ray_on_def, safe)
      show "Q  𝒫" 
        by (simp add: path_Q)
    next
      fix x 
      assume "[x;a;b]" 
      then show "x  Q"
        using b  Q b  a betw_a_in_path event_a path_Q by blast
    next
      show "is_ray {x. [x;a;b]}"
    proof -
      have "[x;a;b]"
        using x?L by simp
      have "?L = ray a x"
      proof
        show "ray a x  ?L"
        proof
          fix e assume "eray a x"
          show "e?L"
            using wxyz_ord ray_cases abc_bcd_abd abd_bcd_abc abc_sym
            by (metis [x;a;b] e  ray a x mem_Collect_eq)
        qed
        show "?L  ray a x"
        proof
          fix e assume "e?L"
          hence "[e;a;b]"
            by simp
          show "eray a x"
          proof (cases)
            assume "e=x"
            thus ?thesis
              by (simp add: ray_def)
          next
            assume "ex"
            hence "[e;x;a]  [x;e;a]" using wxyz_ord
              by (meson [e;a;b] [x;a;b] abc_abd_bcdbdc abc_sym)
            thus "eray a x"
              by (metis Un_iff abc_sym insertCI pro_betw ray_def seg_betw)
          qed
        qed
      qed
      thus "is_ray ?L" by auto
    qed
  qed
  show "is_ray_on ?R Q"
  proof (unfold is_ray_on_def, safe)
    show "Q  𝒫" 
      by (simp add: path_Q)
  next 
    fix x 
    assume "[a;x;b]" 
    then show "x  Q"
      by (simp add: b  Q abc_abc_neq betw_b_in_path event_a path_Q)
  next
    fix x
    assume "[a;b;x]" 
    then show "x  Q"
      by (simp add: b  Q abc_abc_neq betw_c_in_path event_a path_Q)
  next
    show "b  Q" using b  Q .
  next
    show "is_ray {y. [a;y;b]  [a;b;y}"
    proof -
      have "[a;y;b]  [a;b;y]  y=b"
        using y?R by blast
      have "?R = ray a y"
      proof
        show "ray a y  ?R"
        proof
          fix e assume "eray a y"
          hence "[a;e;y]  [a;y;e]  y=e"
            using ray_cases by auto
          show "e?R"
          proof -
            { assume "e  b"
              have "(e  y  e  b)  [w;a;y]  [a;e;b]  [a;b;e"
                using [a;y;b]  [a;b;y]  y = b w  {x. [x;a;b]} abd_bcd_abc by blast
              hence "[a;e;b]  [a;b;e"
                using abc_abd_bcdbdc abc_bcd_abd abd_bcd_abc
                by (metis [a;e;y]  [a;y;e w  ?L mem_Collect_eq)
            }
            thus ?thesis
              by blast
          qed
        qed
        show "?R  ray a y"
        proof
          fix e assume "e?R"
          hence aeb_cases: "[a;e;b]  [a;b;e]  e=b"
            by blast
          hence aey_cases: "[a;e;y]  [a;y;e]  e=y"
            using abc_abd_bcdbdc abc_bcd_abd abd_bcd_abc
            by (metis [a;y;b]  [a;b;y]  y = b x  {x. [x;a;b]} mem_Collect_eq)
          show "eray a y"
          proof -
            {
              assume "e=b"
              hence ?thesis
                using [a;y;b]  [a;b;y]  y = b b  a pro_betw ray_def seg_betw by auto
            } moreover {
              assume "[a;e;b]  [a;b;e]"
              assume "ye"
              hence "[a;e;y]  [a;y;e]"
                using aey_cases by auto
              hence "eray a y"
                unfolding ray_def using abc_abc_neq pro_betw seg_betw by auto
            } moreover {
              assume "[a;e;b]  [a;b;e]"
              assume "y=e"
              have "eray a y"
                unfolding ray_def by (simp add: y = e)
            }
            ultimately show ?thesis
              using aeb_cases by blast 
          qed
        qed
      qed
      thus "is_ray ?R" by auto
    qed
  qed
    show "(r?R. l?L. [l;a;r])"
      using abd_bcd_abc by blast
    show "x?R. y?R. ¬ [x;a;y]"
      by (smt abc_ac_neq abc_bcd_abd abd_bcd_abc mem_Collect_eq)
    show "x?L. y?L. ¬ [x;a;y]"
      using abc_abc_neq abc_abd_bcdbdc abc_only_cba by blast
    show "Q-{a}  ?R  ?L"
      using Q = {x. [x;a;b]}  {a}  {y. [a;y;b]  [a;b;y} by blast
  qed
  thus ?thesis
    by (metis (mono_tags, lifting))
qed

text ‹
  The definition closest_to› in prose:
  Pick any $r \in R$. The closest event $c$ is such that there is no closer event in $L$,
  i.e. all other events of $L$ are further away from $r$.
  Thus in $L$, $c$ is the element closest to $R$.
›
definition closest_to :: "('a set)  'a  ('a set)  bool"
  where "closest_to L c R  cL  (rR. lL-{c}. [l;c;r])"


lemma int_on_path:
  assumes "lL" "rR" "Q𝒫"
      and partition: "LQ" "L{}" "RQ" "R{}" "LR=Q"
    shows "interval l r  Q"
proof
  fix x assume "xinterval l r"
  thus "xQ"
    unfolding interval_def segment_def
    using betw_b_in_path partition(5) Q𝒫 seg_betw l  L r  R
    by blast
qed


lemma ray_of_bounds1:
  assumes "Q𝒫" "[fX|(f 0)..]" "XQ" "closest_bound c X" "is_bound_f b X f" "bc"
  assumes "is_bound_f x X f"
  shows "x=b  x=c  [c;x;b]  [c;b;x]"
proof -
  have "xQ"
    using bound_on_path assms(1,3,7) unfolding all_bounds_def is_bound_def is_bound_f_def
    by auto
  {
    assume "x=b"
    hence ?thesis by blast
  } moreover {
    assume "x=c"
    hence ?thesis by blast
  } moreover {
    assume "xb" "xc"
    hence ?thesis
      by (meson abc_abd_bcdbdc assms(4,5,6,7) closest_bound_def is_bound_def)
  }
  ultimately show ?thesis by blast
qed


lemma ray_of_bounds2:
  assumes "Q𝒫" "[fX|(f 0)..]" "XQ" "closest_bound_f c X f" "is_bound_f b X f" "bc"
  assumes "x=b  x=c  [c;x;b]  [c;b;x]"
  shows "is_bound_f x X f"
proof -
  have "xQ"
    using assms(1,3,4,5,6,7) betw_b_in_path betw_c_in_path bound_on_path
    using closest_bound_f_def is_bound_f_def by metis
  {
    assume "x=b"
    hence ?thesis
      by (simp add: assms(5))
  } moreover {
    assume "x=c"
    hence ?thesis using assms(4)
      by (simp add: closest_bound_f_def)
  } moreover {
    assume "[c;x;b]"
    hence ?thesis unfolding is_bound_f_def
    proof (safe)
      fix i j::nat
      show "[fX|f 0..]"
        by (simp add: assms(2))
      assume "i<j"
      hence "[f i; f j; b]"
        using assms(5) is_bound_f_def by blast
      hence "[f j; b; c]  [f j; c; b]"
        using i < j abc_abd_bcdbdc assms(4,6) closest_bound_f_def is_bound_f_def by auto
      thus "[f i; f j; x]"
        by (meson [c;x;b] [f i; f j; b] abc_bcd_acd abc_sym abd_bcd_abc)
    qed
  } moreover {
    assume "[c;b;x]"
    hence ?thesis unfolding is_bound_f_def
    proof (safe)
      fix i j::nat
      show "[fX|f 0..]"
        by (simp add: assms(2))
      assume "i<j"
      hence "[f i; f j; b]"
        using assms(5) is_bound_f_def by blast
      hence "[f j; b; c]  [f j; c; b]"
        using i < j abc_abd_bcdbdc assms(4,6) closest_bound_f_def is_bound_f_def by auto
      thus "[f i; f j; x]"
      proof -
        have "(c = b)  [f 0; c; b]"
          using assms(4,5) closest_bound_f_def is_bound_def by auto
        hence "[f j; b; c]  [x; f j; f i]"
          by (metis abc_bcd_acd abc_only_cba(2) assms(5) is_bound_f_def neq0_conv)
        thus ?thesis
          using [c;b;x] [f i; f j; b] [f j; b; c]  [f j; c; b] abc_bcd_acd abc_sym
          by blast
      qed
    qed
  }
  ultimately show ?thesis using assms(7) by blast
qed


lemma ray_of_bounds3:
  assumes "Q𝒫" "[fX|(f 0)..]" "XQ" "closest_bound_f c X f" "is_bound_f b X f" "bc"
  shows "all_bounds X = insert c (ray c b)"
proof
  let ?B = "all_bounds X"
  let ?C = "insert c (ray c b)"
  show "?B  ?C"
  proof
    fix x assume "x?B"
    hence "is_bound x X"
      by (simp add: all_bounds_def)
    hence "x=b  x=c  [c;x;b]  [c;b;x]"
      using ray_of_bounds1 abc_abd_bcdbdc assms(4,5,6)
      by (meson closest_bound_f_def is_bound_def)
    thus "x?C"
      using pro_betw ray_def seg_betw by auto
  qed
  show "?C  ?B"
  proof
    fix x assume "x?C"
    hence "x=b  x=c  [c;x;b]  [c;b;x]"
      using pro_betw ray_def seg_betw by auto
    hence "is_bound x X"
      unfolding is_bound_def using ray_of_bounds2 assms
      by blast
    thus "x?B"
      by (simp add: all_bounds_def)
  qed
qed


lemma int_in_closed_ray:
  assumes "path ab a b"
  shows "interval a b  insert a (ray a b)"
proof
  let ?i = "interval a b"
  show "interval a b  insert a (ray a b)"
  proof -
    obtain c where "[a;b;c]" using prolong_betw2
      using assms by blast
    hence "cray a b"
      using abc_abc_neq pro_betw ray_def by auto
    have "cinterval a b"
      using [a;b;c] abc_abc_neq abc_only_cba(2) interval_def seg_betw by auto
    thus ?thesis
      using c  ray a b by blast
  qed
  show "interval a b  insert a (ray a b)"
    using interval_def ray_def by auto
qed


end (* context MinkowskiSpacetime *)


section "3.8 Connectedness of the unreachable set"
context MinkowskiSpacetime begin

subsection ‹Theorem 13 (Connectedness of the Unreachable Set)›

theorem (*13*) unreach_connected:
  assumes path_Q: "Q𝒫"
      and event_b: "bQ" "b"
      and unreach: "Qx  unreach-on Q from b" "Qz  unreach-on Q from b"
      and xyz: "[Qx; Qy; Qz]"
    shows "Qy  unreach-on Q from b"
proof -
  have xz: "Qx  Qz" using abc_ac_neq xyz by blast

  text ‹First we obtain the chain from @{thm I6}.›
  have in_Q: "QxQ  QyQ  QzQ"
    using betw_b_in_path path_Q unreach(1,2) xz unreach_on_path xyz by blast
  hence event_y: "Qy"
    using in_path_event path_Q by blast
  text‹legacy: @{thm I6_old} instead of @{thm I6}
  obtain X f where X_def: "ch_by_ord f X" "f 0 = Qx" "f (card X - 1) = Qz"
      "(i{1 .. card X - 1}. (f i)  unreach-on Q from b  (Qy. [f (i - 1); Qy; f i]  Qy  unreach-on Q from b))"
      "short_ch X  Qx  X  Qz  X  (Qy. [Qx; Qy; Qz]  Qy  unreach-on Q from b)"
    using I6_old [OF assms(1-5) xz] by blast
  hence fin_X: "finite X"
    using xz not_less by fastforce
  obtain N where "N=card X" "N2"
    using X_def(2,3) xz by fastforce

  text ‹
  Then we have to manually show the bounds, defined via indices only, are in the obtained chain.
›
  let ?a = "f 0"
  let ?d = "f (card X - 1)"
  {
    assume "card X = 2"
    hence "short_ch X" "?a  X  ?d  X" "?a  ?d"
      using X_def card X = 2 short_ch_card_2 xz by blast+
  }
  hence "[fX|Qx..Qz]"
    using chain_defs by (metis X_def(1-3) fin_X)

  text ‹
  Further on, we split the proof into two cases, namely the split Schutz absorbs into his
  non-strict termlocal_ordering. Just below is the statement we use @{thm disjE} with.›
  have y_cases: "QyX  QyX" by blast
  have y_int: "Qyinterval Qx Qz"
    using interval_def seg_betw xyz by auto
  have X_in_Q: "XQ"
    using chain_on_path_I6 [where Q=Q and X=X] X_def event_b path_Q unreach xz [fX|Qx .. Qz] by blast

  show ?thesis
  proof (cases)
    text ‹We treat short chains separately.
      (Legacy: they used to have a separate clause in @{thm I6}, now @{thm I6_old})›
    assume "N=2"
    thus ?thesis
      using X_def(1,5) xyz N = card X event_y short_ch_card_2 by auto
  next
    text ‹
    This is where Schutz obtains the chain from Theorem 11. We instead use the chain we already have
    with only a part of Theorem 11, namely @{thm int_split_to_segs}.
    ?S› is defined like in @{thm segmentation}.›
    assume "N2"
    hence "N3" using 2  N by auto
    have "2card X" using 2  N N = card X by blast
    show ?thesis using y_cases
    proof (rule disjE)
      assume "QyX"
      then obtain i where i_def: "i<card X" "Qy = f i"
        using X_def(1) by (metis fin_X obtain_index_fin_chain)
      have "i0  icard X - 1"
        using X_def(2,3)
        by (metis abc_abc_neq i_def(2) xyz)
      hence "i{1..card X -1}"
        using i_def(1) by fastforce
      thus ?thesis using X_def(4) i_def(2) by metis
    next
      assume "QyX"

      let ?S = "if card X = 2 then {segment ?a ?d} else {segment (f i) (f(i+1)) | i. i<card X - 1}"

      have "Qy?S"
      proof -
        obtain c where "[fX|Qx..c..Qz]"
          using X_def(1) N = card X N2 [fX|Qx..Qz] short_ch_card_2
          by (metis 2  N le_neq_implies_less long_chain_2_imp_3)
        have "interval Qx Qz = ?S  X"
          using int_split_to_segs [OF [fX|Qx..c..Qz]] by auto
        thus ?thesis
          using QyX y_int by blast
      qed
      then obtain s where "s?S" "Qys" by blast

      have "i. i{1..(card X)-1}  [(f(i-1)); Qy; f i]"
      proof -
        obtain i' where i'_def: "i' < N-1" "s = segment (f i') (f (i' + 1))"
          using Qys s?S N=card X
          by (smt 2  N N  2 le_antisym mem_Collect_eq not_less)
        show ?thesis
        proof (rule exI, rule conjI)
          show "(i'+1)  {1..card X - 1}"
            using i'_def(1)
            by (simp add: N = card X)
          show "[f((i'+1) - 1); Qy; f(i'+1)]"
            using i'_def(2) Qys seg_betw by simp
        qed
      qed
      then obtain i where i_def: "i{1..(card X)-1}" "[(f(i-1)); Qy; f i]"
        by blast

      show ?thesis
        by (meson X_def(4) i_def event_y)
    qed
  qed
qed

subsection ‹Theorem 14 (Second Existence Theorem)›

lemma (*for 14i*) union_of_bounded_sets_is_bounded:
  assumes "xA. [a;x;b]" "xB. [c;x;d]" "AQ" "BQ" "Q𝒫"
    "card A > 1  infinite A" "card B > 1  infinite B"
  shows "lQ. uQ. xAB. [l;x;u]"
proof -
  let ?P = "λ A B. lQ. uQ. xAB. [l;x;u]"
  let ?I = "λ A a b. (card A > 1  infinite A)  (xA. [a;x;b])"
  let ?R = "λA. a b. ?I A a b"

  have on_path: "a b A. A  Q  ?I A a b  b  Q  a  Q"
  proof -
    fix a b A assume "AQ" "?I A a b"
    show "bQaQ"
    proof (cases)
      assume "card A  1  finite A"
      thus ?thesis
        using ?I A a b by auto
    next
      assume "¬ (card A  1  finite A)"
      hence asmA: "card A > 1  infinite A"
        by linarith
      then obtain x y where "xA" "yA" "xy"
      proof 
        assume "1 < card A" "x y. x  A; y  A; x  y  thesis"
        then show ?thesis 
          by (metis One_nat_def Suc_le_eq card_le_Suc_iff insert_iff)
      next
        assume "infinite A" "x y. x  A; y  A; x  y  thesis"
        then show ?thesis 
        using infinite_imp_nonempty by (metis finite_insert finite_subset singletonI subsetI)
    qed
      have "xQ" "yQ"
        using A  Q x  A y  A by auto
      have "[a;x;b]" "[a;y;b]"
        by (simp add: (1 < card A  infinite A)  (xA. [a;x;b]) x  A y  A)+ 
      hence "betw4 a x y b  betw4 a y x b"
        using x  y abd_acd_abcdacbd by blast
      hence "aQ  bQ"
        using Q𝒫 xQ xy xQ yQ betw_a_in_path betw_c_in_path by blast
      thus ?thesis by simp
    qed
  qed

  show ?thesis
  proof (cases)
    assume "ab  ac  ad  bc  bd  cd"
    show "?P A B"
    proof (rule_tac P="?P" and A=Q in wlog_endpoints_distinct)

      text ‹First, some technicalities: the relations $P, I, R$ have the symmetry required.›
      show "a b I. ?I I a b  ?I I b a" using abc_sym by blast
      show "a b A. A  Q  ?I A a b  b  Q  a  Q" using on_path assms(5) by blast
      show "I J. ?R I  ?R J  ?P I J  ?P J I" by (simp add: Un_commute)

      text ‹Next, the lemma/case assumptions have to be repeated for Isabelle.›
      show "?I A a b" "?I B c d" "AQ" "BQ" "Q𝒫"
        using assms by simp+ 
      show "ab  ac  ad  bc  bd  cd"
        using ab  ac  ad  bc  bd  cd by simp

      text ‹Finally, the important bit: proofs for the necessary cases of betweenness.›
      show "?P I J"
        if "?I I a b" "?I J c d"  "IQ" "JQ"
          and "[a;b;c;d]  [a;c;b;d]  [a;c;d;b]"
        for I J a b c d
      proof -
        consider "[a;b;c;d]"|"[a;c;b;d]"|"[a;c;d;b]"
          using [a;b;c;d]  [a;c;b;d]  [a;c;d;b] by fastforce
        thus ?thesis
        proof (cases)
          assume asm: "[a;b;c;d]" show "?P I J"
          proof -
            have "x IJ. [a;x;d]"
              by (metis Un_iff asm betw4_strong betw4_weak that(1) that(2))
            moreover have "aQ" "dQ"
              using assms(5) on_path that(1-4) by blast+ 
            ultimately show ?thesis by blast
          qed
        next
          assume "[a;c;b;d]" show "?P I J"
          proof -
            have "x IJ. [a;x;d]"
              by (metis Un_iff betw4 a c b d abc_bcd_abd abc_bcd_acd betw4_weak that(1,2))
            moreover have "aQ" "dQ"
              using assms(5) on_path that(1-4) by blast+ 
            ultimately show ?thesis by blast
          qed
        next
          assume "[a;c;d;b]" show "?P I J"
          proof -
            have "x IJ. [a;x;b]"
              using betw4 a c d b abc_bcd_abd abc_bcd_acd abe_ade_bcd_ace
              by (meson UnE that(1,2))
            moreover have "aQ" "bQ"
              using assms(5) on_path that(1-4) by blast+
            ultimately show ?thesis by blast
          qed
        qed
      qed
    qed
  next
    assume "¬(ab  ac  ad  bc  bd  cd)"
  
    show "?P A B"
    proof (rule_tac P="?P" and A=Q in wlog_endpoints_degenerate)

      text ‹
        This case follows the same pattern as above: the next five show› statements
        are effectively bookkeeping.›
      show "a b I. ?I I a b  ?I I b a" using abc_sym by blast
      show "a b A. A  Q  ?I A a b  b  Q  a  Q" using on_path Q𝒫 by blast
      show "I J. ?R I  ?R J  ?P I J  ?P J I" by (simp add: Un_commute)

      show "?I A a b" "?I B c d" "AQ" "BQ" "Q𝒫"
        using assms by simp+
      show "¬ (a  b  b  c  c  d  a  d  a  c  b  d)"
        using ¬ (a  b  a  c  a  d  b  c  b  d  c  d) by blast

      text ‹Again, this is the important bit: proofs for the necessary cases of degeneracy.›
      show "(a = b  b = c  c = d  ?P I J)  (a = b  b  c  c = d  ?P I J) 
          (a = b  b = c  c  d  ?P I J)  (a = b  b  c  c  d  a  d  ?P I J) 
          (a  b  b = c  c  d  a = d  ?P I J) 
          ([a;b;c]  a = d  ?P I J)  ([b;a;c]  a = d  ?P I J)"
      if "?I I a b" "?I J c d" "I  Q" "J  Q"
      for I J a b c d
      proof (intro conjI impI)
        assume "a = b  b = c  c = d"
        show "lQ. uQ. xI  J. [l;x;u]"
          using a = b  b = c  c = d abc_ac_neq assms(5) ex_crossing_path that(1,2)
          by fastforce
      next
        assume "a = b  b  c  c = d"
        show "lQ. uQ. xI  J. [l;x;u]"
          using a = b  b  c  c = d abc_ac_neq assms(5) ex_crossing_path that(1,2)
          by (metis Un_iff)
      next
        assume "a = b  b = c  c  d"
        hence "x IJ. [c;x;d]"
          using abc_abc_neq that(1,2) by fastforce
        moreover have "cQ" "dQ"
          using on_path a = b  b = c  c  d that(1,3) abc_abc_neq by metis+ 
        ultimately show "lQ. uQ. xI  J. [l;x;u]" by blast
      next
        assume "a = b  b  c  c  d  a  d"
        hence "x IJ. [c;x;d]"
          using abc_abc_neq that(1,2) by fastforce
        moreover have "cQ" "dQ"
          using on_path a = b  b  c  c  d  a  d that(1,3) abc_abc_neq by metis+ 
        ultimately show "lQ. uQ. xI  J. [l;x;u]" by blast
      next
        assume "a  b  b = c  c  d  a = d"
        hence "x IJ. [c;x;d]"
          using abc_sym that(1,2) by auto
        moreover have "cQ" "dQ"
          using on_path a  b  b = c  c  d  a = d that(1,3) abc_abc_neq by metis+ 
        ultimately show "lQ. uQ. xI  J. [l;x;u]" by blast
      next
        assume "[a;b;c]  a = d"
        hence "x IJ. [c;x;d]"
          by (metis UnE abc_acd_abd abc_sym that(1,2))
        moreover have "cQ" "dQ"
          using on_path that(2,4) by blast+ 
        ultimately show "lQ. uQ. xI  J. [l;x;u]" by blast
      next
        assume "[b;a;c]  a = d"
        hence "x IJ. [c;x;b]"
          using  abc_sym abd_bcd_abc betw4_strong that(1,2) by (metis Un_iff)
        moreover have "cQ" "bQ"
          using on_path that by blast+ 
        ultimately show "lQ. uQ. xI  J. [l;x;u]" by blast
      qed
    qed
  qed
qed


lemma (*for 14i*) union_of_bounded_sets_is_bounded2:
  assumes "xA. [a;x;b]" "xB. [c;x;d]" "AQ" "BQ" "Q𝒫"
      "1<card A  infinite A" "1<card B  infinite B"
    shows "lQ-(AB). uQ-(AB). xAB. [l;x;u]"
  using assms union_of_bounded_sets_is_bounded
    [where A=A and a=a and b=b and B=B and c=c and d=d and Q=Q]
  by (metis Diff_iff abc_abc_neq)

text ‹
  Schutz proves a mildly stronger version of this theorem than he states. Namely, he gives an
  additional condition that has to be fulfilled by the bounds $y,z$ in the proof (y,z∉unreach-on Q from ab›).
  This condition is trivial given abc_abc_neq›. His stating it in the proof makes me wonder
  whether his (strictly speaking) undefined notion of bounded set is somehow weaker than the
  version using strict betweenness in his theorem statement and used here in Isabelle.
  This would make sense, given the obvious analogy with sets on the real line.
›

theorem (*14i*) second_existence_thm_1:
  assumes path_Q: "Q𝒫"
      and events: "aQ" "bQ"
      and reachable: "path_ex a q1" "path_ex b q2" "q1Q" "q2Q"
    shows "yQ. zQ. (xunreach-on Q from a. [y;x;z])  (xunreach-on Q from b. [y;x;z])"
proof -
  text ‹Slightly annoying: Schutz implicitly extends bounded› to sets, so his statements are neater.›

(* alternative way of saying reachable *)
  have "qQ. q(unreach-on Q from a)" "qQ. q(unreach-on Q from b)"
    using cross_in_reachable reachable by blast+

  text ‹This is a helper statement for obtaining bounds in both directions of both unreachable sets.
  Notice this needs Theorem 13 right now, Schutz claims only Theorem 4. I think this is necessary?›
  have get_bds: "laQ. uaQ. launreach-on Q from a  uaunreach-on Q from a  (xunreach-on Q from a. [la;x;ua])"
    if asm: "aQ" "path_ex a q" "qQ"
    for a q
  proof -
    obtain Qy where "Qyunreach-on Q from a"
      using asm(2) a  Q in_path_event path_Q two_in_unreach by blast
    then obtain la where "la  Q - unreach-on Q from a"
      using asm(2,3) cross_in_reachable by blast
    then obtain ua where "ua  Q - unreach-on Q from a" "[la;Qy;ua]" "la  ua"
      using unreachable_set_bounded [where Q=Q and b=a and Qx=la and Qy=Qy]
      using Qy  unreach-on Q from a asm in_path_event path_Q by blast
    have "la  unreach-on Q from a  ua  unreach-on Q from a  (xunreach-on Q from a. (xla  xua)  [la;x;ua])"
    proof (intro conjI)
      show "la  unreach-on Q from a"
        using la  Q - unreach-on Q from a by force
    next
      show "ua  unreach-on Q from a"
        using ua  Q - unreach-on Q from a by force
    next show "xunreach-on Q from a. x  la  x  ua  [la;x;ua]"
    proof (safe)
      fix x assume "xunreach-on Q from a" "xla" "xua"
      {
        assume "x=Qy" hence "[la;x;ua]" by (simp add: [la;Qy;ua])
      } moreover {
        assume "xQy"
        have "[Qy;x;la]  [la;Qy;x]"
        proof -
          { assume "[x;la;Qy]"
            hence "launreach-on Q from a"
              using unreach_connected Qyunreach-on Q from axunreach-on Q from axQy in_path_event path_Q that by blast
            hence False
              using la  Q - unreach-on Q from a by blast }
          thus "[Qy;x;la]  [la;Qy;x]"
            using some_betw [where Q=Q and a=x and b=la and c=Qy] path_Q unreach_on_path
            using Qy  unreach-on Q from a la  Q - unreach-on Q from a x  unreach-on Q from a x  Qy x  la by force
        qed
        hence "[la;x;ua]"
        proof
          assume "[Qy;x;la]"
          thus ?thesis using [la;Qy;ua] abc_acd_abd abc_sym by blast
        next
          assume "[la;Qy;x]"
          hence "[la;x;ua]  [la;ua;x]"
            using [la;Qy;ua] x  ua abc_abd_acdadc by auto
          have "¬[la;ua;x]"
            using unreach_connected that abc_abc_neq abc_acd_bcd in_path_event path_Q
            by (metis DiffD2 Qy  unreach-on Q from a [la;Qy;ua] ua  Q - unreach-on Q from a x  unreach-on Q from a)
          show ?thesis
            using [la;x;ua]  [la;ua;x] ¬ [la;ua;x] by linarith
        qed
      }
      ultimately show "[la;x;ua]" by blast
    qed
  qed
    thus ?thesis using la  Q - unreach-on Q from a ua  Q - unreach-on Q from a by force
  qed

  have "yQ. zQ. (x(unreach-on Q from a)(unreach-on Q from b). [y;x;z])"
  proof -
    obtain la ua where "xunreach-on Q from a. [la;x;ua]"
      using events(1) get_bds reachable(1,3) by blast
    obtain lb ub where "xunreach-on Q from b. [lb;x;ub]"
      using events(2) get_bds reachable(2,4) by blast
    have "unreach-on Q from a  Q" "unreach-on Q from b  Q"
      by (simp add: subsetI unreach_on_path)+
    moreover have "1 < card (unreach-on Q from a)  infinite (unreach-on Q from a)"
      using two_in_unreach events(1) in_path_event path_Q reachable(1)
      by (metis One_nat_def card_le_Suc0_iff_eq not_less)
    moreover have "1 < card (unreach-on Q from b)  infinite (unreach-on Q from b)"
      using two_in_unreach events(2) in_path_event path_Q reachable(2)
      by (metis One_nat_def card_le_Suc0_iff_eq not_less)
    ultimately show ?thesis
      using union_of_bounded_sets_is_bounded [where Q=Q and A="unreach-on Q from a" and B="unreach-on Q from b"]
      using get_bds assms xunreach-on Q from a. [la;x;ua] xunreach-on Q from b. [lb;x;ub]
      by blast
  qed

  then obtain y z where "yQ" "zQ" "(x(unreach-on Q from a)(unreach-on Q from b). [y;x;z])"
    by blast
  show ?thesis
  proof (rule bexI)+
    show "yQ" by (simp add: y  Q)
    show "zQ" by (simp add: z  Q)
    show "(xunreach-on Q from a. [z;x;y])  (xunreach-on Q from b. [z;x;y])"
      by (simp add: xunreach-on Q from a  unreach-on Q from b. [y;x;z] abc_sym)
  qed
qed


theorem (*14*) second_existence_thm_2:
  assumes path_Q: "Q𝒫"
      and events: "aQ" "bQ" "cQ" "dQ" "cd"
      and reachable: "P𝒫. qQ. path P a q" "P𝒫. qQ. path P b q"
    shows "eQ. ae𝒫. be𝒫. path ae a e  path be b e  [c;d;e]"
proof -
  obtain y z where bounds_yz: "(xunreach-on Q from a. [z;x;y])  (xunreach-on Q from b. [z;x;y])"
               and yz_inQ: "yQ" "zQ"
    using second_existence_thm_1 [where Q=Q and a=a and b=b]
    using path_Q events(1,2) reachable by blast
  have "y(unreach-on Q from a)(unreach-on Q from b)" "z(unreach-on Q from a)(unreach-on Q from b)"
    by (meson Un_iff (xunreach-on Q from a. [z;x;y])  (xunreach-on Q from b. [z;x;y]) abc_abc_neq)+ 
  let ?P = "λe ae be. (eQ  path ae a e  path be b e  [c;d;e])"

  have exist_ay: "ay. path ay a y"
    if "aQ" "P𝒫. qQ. path P a q" "y(unreach-on Q from a)" "yQ"
    for a y
    using in_path_event path_Q that unreachable_bounded_path_only
    by blast

  have "[c;d;y]  y;c;d]  [c;y;d"
    by (meson y  Q abc_sym events(3-5) path_Q some_betw)
  moreover have "[c;d;z]  z;c;d]  [c;z;d"
    by (meson z  Q abc_sym events(3-5) path_Q some_betw)
  ultimately consider "[c;d;y]" | "[c;d;z]" |
                      "((y;c;d]  [c;y;d)  (z;c;d]  [c;z;d))"
    by auto
  thus ?thesis
  proof (cases)
    assume "[c;d;y]"
    have "y(unreach-on Q from a)" "y(unreach-on Q from b)"
      using y  unreach-on Q from a  unreach-on Q from b by blast+
    then obtain ay yb where "path ay a y" "path yb b y"
      using yQ exist_ay events(1,2) reachable(1,2) by blast
    have "?P y ay yb"
      using [c;d;y] path ay a y path yb b y y  Q by blast
    thus ?thesis by blast
  next
    assume "[c;d;z]"
    have "z(unreach-on Q from a)" "z(unreach-on Q from b)"
      using z  unreach-on Q from a  unreach-on Q from b by blast+ 
    then obtain az bz where "path az a z" "path bz b z"
      using zQ exist_ay events(1,2) reachable(1,2) by blast
    have "?P z az bz"
      using [c;d;z] path az a z path bz b z z  Q by blast
    thus ?thesis by blast
  next
    assume "(y;c;d]  [c;y;d)  (z;c;d]  [c;z;d)"
    have "e. [c;d;e]"
      using prolong_betw (* works as Schutz says! *)
      using events(3-5) path_Q by blast
    then obtain e where "[c;d;e]" by auto
    have "¬[y;e;z]"
    proof (rule notI)
      text ‹Notice Theorem 10 is not needed for this proof, and does not seem to help sledgehammer›.
        I think this is because it cannot be easily/automatically reconciled with non-strict
        notation.›
      assume "[y;e;z]"
      moreover consider "(y;c;d]  z;c;d])" | "(y;c;d]  [c;z;d)" |
               "([c;y;d  z;c;d])" | "([c;y;d  [c;z;d)"
        using (y;c;d]  [c;y;d)  (z;c;d]  [c;z;d) by linarith
      ultimately show False
        by (smt [c;d;e] abc_ac_neq betw4_strong betw4_weak)
    qed
    have "eQ"
      using [c;d;e] betw_c_in_path events(3-5) path_Q by blast
    have "e unreach-on Q from a" "e unreach-on Q from b"
      using bounds_yz ¬ [y;e;z] abc_sym by blast+ 
    hence ex_aebe: "ae be. path ae a e  path be b e"
      using e  Q events(1,2) in_path_event path_Q reachable(1,2) unreachable_bounded_path_only
      by metis
    thus ?thesis
      using [c;d;e] e  Q by blast
  qed
qed

text ‹
  The assumption Q≠R› in Theorem 14(iii) is somewhat implicit in Schutz.
  If Q=R›, unreach-on Q from a› is empty, so the third conjunct of the conclusion is meaningless.
›

theorem (*14*) second_existence_thm_3:
  assumes paths: "Q𝒫" "R𝒫" "QR"
      and events: "xQ" "xR" "aR" "ax" "bQ"
      and reachable: "P𝒫. qQ. path P b q"
    shows "e. ae𝒫. be𝒫. path ae a e  path be b e  (yunreach-on Q from a. [x;y;e])"
proof -
  have "aQ"
    using events(1-4) paths eq_paths by blast
  hence "unreach-on Q from a  {}"
    by (metis events(3) ex_in_conv in_path_event paths(1,2) two_in_unreach)

  then obtain d where "d unreach-on Q from a" (*as in Schutz*)
    by blast
  have "xd"
    using d  unreach-on Q from a cross_in_reachable events(1) events(2) events(3) paths(2) by auto
  have "dQ"
    using d  unreach-on Q from a unreach_on_path by blast

  have "eQ. ae be. [x;d;e]  path ae a e  path be b e"
    using second_existence_thm_2 [where c=x and Q=Q and a=a and b=b and d=d] (*as in Schutz*)
    using a  Q d  Q x  d events(1-3,5) paths(1,2) reachable by blast
  then obtain e ae be where conds: "[x;d;e]  path ae a e  path be b e" by blast
  have "y(unreach-on Q from a). [x;y;e]"
  proof
    fix y assume "y(unreach-on Q from a)"
    hence "yQ"
      using unreach_on_path by blast
    show "[x;y;e]"
    proof (rule ccontr)
      assume "¬[x;y;e]"
      then consider "y=x" | "y=e" | "[y;x;e]" | "[x;e;y]"
        by (metis dQ yQ abc_abc_neq abc_sym betw_c_in_path conds events(1) paths(1) some_betw)
      thus False
      proof (cases)
        assume "y=x" thus False
          using y  unreach-on Q from a events(2,3) paths(1,2) same_empty_unreach unreach_equiv unreach_on_path
          by blast
      next
        assume "y=e" thus False
          by (metis yQ assms(1) conds empty_iff same_empty_unreach unreach_equiv y  unreach-on Q from a)
      next
        assume "[y;x;e]"
        hence "[y;x;d]"
          using abd_bcd_abc conds by blast
        hence "x(unreach-on Q from a)"
          using unreach_connected [where Q=Q and Qx=y and Qy=x and Qz=d and b=a]
          using ¬[x;y;e] aQ dunreach-on Q from a yunreach-on Q from a conds in_path_event paths(1) by blast
        thus False
          using empty_iff events(2,3) paths(1,2) same_empty_unreach unreach_equiv unreach_on_path
          by metis
      next
        assume "[x;e;y]"
        hence "[d;e;y]"
          using abc_acd_bcd conds by blast
        hence "e(unreach-on Q from a)"
          using unreach_connected [where Q=Q and Qx=y and Qy=e and Qz=d and b=a]
          using a  Q d  unreach-on Q from a y  unreach-on Q from a
            abc_abc_neq abc_sym events(3) in_path_event paths(1,2)
          by blast
        thus False
          by (metis conds empty_iff paths(1) same_empty_unreach unreach_equiv unreach_on_path)
      qed
    qed
  qed
  thus ?thesis
    using conds in_path_event by blast 
qed


end (* context MinkowskiSpacetime *)

section "Theorem 11 - with path density assumed"
locale MinkowskiDense = MinkowskiSpacetime +
  assumes path_dense: "path ab a b  x. [a;x;b]"
begin

text ‹
  Path density: if $a$ and $b$ are connected by a path, then the segment between them is nonempty.
  Since Schutz insists on the number of segments in his segmentation (Theorem 11), we prove it here,
  showcasing where his missing assumption of path density fits in (it is used three times
  in number_of_segments›, once in each separate meaningful termlocal_ordering case).
›

lemma segment_nonempty:
  assumes "path ab a b"
  obtains x where "x  segment a b"
  using path_dense by (metis seg_betw assms)


lemma (*for 11*) number_of_segments:
  assumes path_P: "P𝒫"
      and Q_def: "QP"
      and f_def: "[fQ|a..b..c]"
    shows "card {segment (f i) (f (i+1)) | i. i<(card Q-1)} = card Q - 1"
proof -
  let ?S = "{segment (f i) (f (i+1)) | i. i<(card Q-1)}"
  let ?N = "card Q"
  let ?g = "λ i. segment (f i) (f (i+1))"
  have "?N  3" using chain_defs f_def by (meson finite_long_chain_with_card)
  have "?g ` {0..?N-2} = ?S"
  proof (safe)
    fix i assume "i{(0::nat)..?N-2}"
    show "ia. segment (f i) (f (i+1)) = segment (f ia) (f (ia+1))  ia<card Q - 1"
    proof
      have "i<?N-1"
        using assms i{(0::nat)..?N-2} ?N3
        by (metis One_nat_def Suc_diff_Suc atLeastAtMost_iff le_less_trans lessI less_le_trans
            less_trans numeral_2_eq_2 numeral_3_eq_3)
      then show "segment (f i) (f (i + 1)) = segment (f i) (f (i + 1))  i<?N-1"
        by blast
    qed
  next
    fix x i assume "i < card Q - 1"
    let ?s = "segment (f i) (f (i + 1))"
    show "?s  ?g ` {0..?N - 2}"
    proof -
      have "i{0..?N-2}"
        using i < card Q - 1 by force
      thus ?thesis by blast
    qed
  qed
  moreover have "inj_on ?g {0..?N-2}"
  proof
    fix i j assume asm: "i{0..?N-2}" "j{0..?N-2}" "?g i = ?g j"
    show "i=j"
    proof (rule ccontr)
      assume "ij"
      hence "f i  f j"
        using asm(1,2) f_def assms(3) indices_neq_imp_events_neq
          [where X=Q and f=f and a=a and b=b and c=c and i=i and j=j]
        by auto
      show False
      proof (cases)
        assume "j=i+1" hence "j=Suc i" by linarith
        have "Suc(Suc i) < ?N" using asm(1,2) eval_nat_numeral j = Suc i by auto
        hence "[f i; f (Suc i); f (Suc (Suc i))]"
          using assms short_ch_card ?N3 chain_defs local_ordering_def
          by (metis short_ch_alt(1) three_in_set3)
        hence "[f i; f j; f (j+1)]" by (simp add: j = i + 1)
        obtain e where "e?g j" using segment_nonempty abc_ex_path asm(3)
          by (metis [f i; f j; f (j+1)] f i  f j j = i + 1)
        hence "e?g i"
          using asm(3) by blast
        have "[f i; f j; e]"
          using abd_bcd_abc [f i; f j; f (j+1)]
          by (meson e  segment (f j) (f (j + 1)) seg_betw) 
        thus False
          using e  segment (f i) (f (i + 1)) j = i + 1 abc_only_cba(2) seg_betw
          by auto
      next assume "ji+1"
        have "i < card Q  j < card Q  (i+1) < card Q"
          using add_mono_thms_linordered_field(3) asm(1,2) assms ?N3 by auto
        hence "f i  Q  f j  Q  f (i+1)  Q"
          using f_def unfolding chain_defs local_ordering_def
          by (metis One_nat_def Suc_diff_le Suc_eq_plus1 3  card Q add_Suc card_1_singleton_iff
            card_gt_0_iff card_insert_if diff_Suc_1 diff_Suc_Suc less_natE less_numeral_extra(1)
            nat.discI numeral_3_eq_3)
        hence "f i  P  f j  P  f (i+1)  P"
          using path_is_union assms
          by (simp add: subset_iff)
        then consider "[f i; (f(i+1)); f j]" | "[f i; f j; (f(i+1))]" |
                      "[(f(i+1)); f i; f j]"
          using some_betw path_P f_def indices_neq_imp_events_neq
            f i  f j i < card Q  j < card Q  i + 1 < card Q j  i + 1
          by (metis abc_sym less_add_one less_irrefl_nat)
        thus False
        proof (cases)
          assume "[(f(i+1)); f i; f j]"
          then obtain e where "e?g i" using segment_nonempty
            by (metis f i  P  f j  P  f (i + 1)  P abc_abc_neq path_P)
          hence "[e; f j; (f(j+1))]"
            using [(f(i+1)); f i; f j]
            by (smt abc_acd_abd abc_acd_bcd abc_only_cba abc_sym asm(3) seg_betw)
          moreover have "e?g j"
            using e  ?g i asm(3) by blast
          ultimately show False
            by (simp add: abc_only_cba(1) seg_betw)
        next
          assume "[f i; f j; (f(i+1))]"
          thus False
            using abc_abc_neq [where b="f j" and a="f i" and c="f(i+1)"] asm(3) seg_betw [where x="f j"]
            using ends_notin_segment by blast
        next
          assume "[f i; (f(i+1)); f j]"
          then obtain e where "e?g i" using segment_nonempty
            by (metis f i  P  f j  P  f (i + 1)  P abc_abc_neq path_P)
          hence "[e; f j; (f(j+1))]"
          proof -
            have "f (i+1)  f j"
              using [f i; (f(i+1)); f j] abc_abc_neq by presburger
            then show ?thesis
              using e  segment (f i) (f (i+1)) [f i; (f(i+1)); f j] asm(3) seg_betw
              by (metis (no_types) abc_abc_neq abc_acd_abd abc_acd_bcd abc_sym)
          qed
          moreover have "e?g j"
            using e  ?g i asm(3) by blast
          ultimately show False
            by (simp add: abc_only_cba(1) seg_betw)
        qed
      qed
    qed
  qed
  ultimately have "bij_betw ?g {0..?N-2} ?S"
    using inj_on_imp_bij_betw by fastforce
  thus ?thesis
    using assms(2) bij_betw_same_card numeral_2_eq_2 numeral_3_eq_3 ?N3
    by (metis (no_types, lifting) One_nat_def Suc_diff_Suc card_atLeastAtMost le_less_trans
        less_Suc_eq_le minus_nat.diff_0 not_less not_numeral_le_zero)
qed

theorem (*11*) segmentation_card:
  assumes path_P: "P𝒫"
      and Q_def: "QP"
      and f_def: "[fQ|a..b]" (* This always exists given card Q > 2 *)
    fixes P1 defines P1_def: "P1  prolongation b a"
    fixes P2 defines  P2_def: "P2  prolongation a b"
    fixes S defines S_def: "S  {segment (f i) (f (i+1)) | i. i<card Q-1}"
    shows "P = ((S)  P1  P2  Q)"
        (* The union of these segments and prolongations with the separating points is the path. *)
          "card S = (card Q-1)  (xS. is_segment x)"
        (* There are N-1 segments. *)
        (* There are two prolongations. *)
          "disjoint (S{P1,P2})" "P1P2" "P1S" "P2S"
        (* The prolongations and all the segments are disjoint. *)
proof -
  let ?N = "card Q"
  have "2  card Q"
    using f_def fin_chain_card_geq_2 by blast
  have seg_facts: "P = (S  P1  P2  Q)" "(xS. is_segment x)"
    "disjoint (S{P1,P2})" "P1P2" "P1S" "P2S"
    using show_segmentation [OF path_P Q_def f_def]
    using P1_def P2_def S_def by fastforce+
  show "P = S  P1  P2  Q" by (simp add: seg_facts(1))
  show "disjoint (S{P1,P2})" "P1P2" "P1S" "P2S"
    using seg_facts(3-6) by blast+
  have "card S = (?N-1)"
  proof (cases)
    assume "?N=2"
    hence "card S = 1"
      by (simp add: S_def)
    thus ?thesis
      by (simp add: ?N = 2)
  next
    assume "?N2"
    hence "?N3"
      using 2  card Q by linarith
    then obtain c where "[fQ|a..c..b]"
      using assms chain_defs short_ch_card_2 2  card Q card Q  2
      by (metis three_in_set3)
    show ?thesis
      using number_of_segments [OF assms(1,2) [fQ|a..c..b]]
      using S_def card Q  2 by presburger
  qed
  thus "card S = card Q - 1  Ball S is_segment"
    using seg_facts(2) by blast
qed


end (* context MinkowskiDense *)

(*
context MinkowskiSpacetime begin
interpretation is_dense: MinkowskiDense apply unfold_locales oops
end
*)

end