Theory Minkowski

(*  Title:      Schutz_Spacetime/Minkowski.thy
    Authors:    Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot
                University of Edinburgh, 2021          
*)
theory Minkowski
imports TernaryOrdering
begin

(* It is best to avoid `distinct` because it makes proof more difficult. *)
(* If it has to be used, the lemma: distinct_length_2_or_more is essential. *)
(* For proofs involving the cardinality of concrete, finite sets see ‹card_insert_if›. *)

text ‹
  Primitives and axioms as given in cite‹pp.~9-17› in "schutz1997".
›

text ‹
  I've tried to do little to no proofs in this file, and keep that in other files.
  So, this is mostly locale and other definitions, except where it is nice to prove something
  about definitional equivalence and the like (plus the intermediate lemmas that are necessary for
  doing so).
›

text ‹
  Minkowski spacetime = $(\mathcal{E}, \mathcal{P}, [\dots])$
  except in the notation here I've used $[[\dots]]$ for $[\dots]$ as Isabelle uses $[\dots]$ for lists.

  Except where stated otherwise all axioms are exactly as they appear in Schutz97.
  It is the independent axiomatic system provided in the main body of the book.
  The axioms O1-O6 are the axioms of order, and largely concern properties of the betweenness
  relation.
  I1-I7 are the axioms of incidence.
  I1-I3 are similar to axioms found in systems for Euclidean geometry. As compared to Hilbert's
  Foundations (HIn), our incidence axioms (In) are loosely identifiable as
   I1 →› HI3, HI8;
   I2 →› HI1;
   I3 →› HI2.
  I4 fixes the dimension of the space.
  I5-I7 are what makes our system non-Galilean, and lead (I think)
  to Lorentz transforms (together with S?) and the ultimate speed limit.
  Axioms S and C and the axioms of symmetry and continuity, where the latter is what makes the
  system second order. Symmetry replaces all of Hilbert's axioms of congruence, when considered
  in the context of I5-I7.
›

section "MinkowskiPrimitive: I1-I3"

text ‹
  Events ℰ›, paths 𝒫›, and sprays. Sprays only need to refer to ℰ› and 𝒫›.
  Axiom in_path_event› is covered in English by saying "a path is a set of events",
  but is necessary to have explicitly as an axiom as the types do not force it to be the case.
›

text ‹
  I think part of why Schutz has I1, together with the trickery
  ⟦ ℰ≠{} ⟧ ⟹› $\dots$ in I4, is that then I4 talks \emph{only} about dimension, and results such
  as no_empty_paths› can be proved using only existence of elements and unreachable sets.
  In our case, it's also a question of ordering the sequence of axiom introductions:
  dimension should really go at the end, since it is not needed for quite a while;
  but many earlier proofs rely on the set of events being non-empty.
  It may be nice to have the existence of paths as a separate axiom too, which currently
  still relies on the axiom of dimension (Schutz has no such axiom either).
›

locale MinkowskiPrimitive =
  fixes  :: "'a set"
    and 𝒫 :: "('a set) set"
  assumes in_path_event [simp]: "Q  𝒫; a  Q  a  "
      (* I1 *)
      and nonempty_events [simp]: "  {}"
      (* I2 *)
      (* It's still possible to have 1 event and 0 paths. *)
      and events_paths: "a  ; b  ; a  b  R𝒫. S𝒫. a  R  b  S  R  S  {}"
      (* I3 *)
      and eq_paths [intro]: "P  𝒫; Q  𝒫; a  P; b  P; a  Q; b  Q; a  b  P = Q"
begin

text ‹This should be ensured by the additional axiom.›

lemma path_sub_events:
  "Q  𝒫  Q  "
by (simp add: subsetI)

lemma paths_sub_power:
  "𝒫  Pow "
by (simp add: path_sub_events subsetI)

text ‹
  Define path› for more terse statements.
  $a \neq b$ because $a$ and $b$ are being used to identify the path, and $a = b$ would not do that.
›

abbreviation path :: "'a set  'a  'a  bool" where
  "path ab a b  ab  𝒫  a  ab  b  ab  a  b"

abbreviation path_ex :: "'a  'a  bool" where
  "path_ex a b  Q. path Q a b"

lemma path_permute:
  "path ab a b = path ab b a"
  by auto

abbreviation path_of :: "'a  'a  'a set" where
  "path_of a b  THE ab. path ab a b"

lemma path_of_ex: "path (path_of a b) a b  path_ex a b"
  using theI' [where P="λx. path x a b"] eq_paths by blast

lemma path_unique:
  assumes "path ab a b" and "path ab' a b"
    shows "ab = ab'"
  using eq_paths assms by blast

lemma paths_cross_once:
  assumes path_Q: "Q  𝒫"
      and path_R: "R  𝒫"
      and Q_neq_R: "Q  R"
      and QR_nonempty: "QR  {}"
  shows "∃!a. QR = {a}"
proof -
  have ab_inQR: "a. aQR" using QR_nonempty in_path_event path_Q by auto
  then obtain a where a_event: "a  " and a_inQR: "a  QR" by auto
  have "QR = {a}"
  proof (rule ccontr)
    assume "QR  {a}"
    then have "bQR. b  a" using a_inQR by blast
    then have "Q = R" using eq_paths a_inQR path_Q path_R by auto
    thus False using Q_neq_R by simp
  qed
  thus ?thesis using a_event by blast
qed


section "Primitives: Unreachable Subset (from an Event)"

text ‹
  The Q ∈ 𝒫 ∧ b ∈ ℰ› constraints are necessary as the types as not expressive enough to do it on
  their own. Schutz's notation is: $Q(b,\emptyset)$.
›

definition unreachable_subset :: "'a set  'a  'a set" ("unreach-on _ from _" [100, 100]) where
  "unreach-on Q from b  {xQ. Q  𝒫  b    b  Q  ¬(path_ex b x)}"


section "Primitives: Kinematic Triangle"

definition kinematic_triangle :: "'a  'a  'a  bool" (" _ _ _" [100, 100, 100] 100) where
    "kinematic_triangle a b c 
       a    b    c    a  b  a  c  b  c
        (Q𝒫. R𝒫. Q  R  (S𝒫. Q  S  R  S
                                        a  Q  b  Q
                                        a  R  c  R
                                        b  S  c  S))"

text ‹A fuller, more explicit equivalent of △›, to show that the above definition is sufficient.›
lemma tri_full:
  " a b c = (a    b    c    a  b  a  c  b  c
               (Q𝒫. R𝒫. Q  R  (S𝒫. Q  S  R  S
                                               a  Q  b  Q  c  Q
                                               a  R  c  R  b  R
                                               b  S  c  S  a  S)))"
unfolding kinematic_triangle_def by (meson path_unique)


section "Primitives: SPRAY"

text ‹
  It's okay to not require $x \in \E$ because if $x \notin \E$ the SPRAY› will be empty anyway,
  and if it's nonempty then $x \in \E$ is derivable.›
definition SPRAY :: "'a  ('a set) set" where
  "SPRAY x  {R𝒫. x  R}"

definition spray :: "'a  'a set" where
  "spray x  {y. RSPRAY x. y  R}"

(* Just for convenience. *)
definition is_SPRAY :: "('a set) set  bool" where
  "is_SPRAY S  x. S = SPRAY x"

definition is_spray :: "'a set  bool" where
  "is_spray S  x. S = spray x"

text ‹Some very simple SPRAY and spray lemmas below.›

lemma SPRAY_event:
  "SPRAY x  {}  x  "
proof (unfold SPRAY_def)
  assume nonempty_SPRAY: "{R  𝒫. x  R}  {}"
  then have x_in_path_R: "R  𝒫. x  R" by blast
  thus "x  " using in_path_event by blast
qed

lemma SPRAY_nonevent:
  "x    SPRAY x = {}"
using SPRAY_event by auto

lemma SPRAY_path:
  "P  SPRAY x  P  𝒫"
by (simp add: SPRAY_def)

lemma in_SPRAY_path:
  "P  SPRAY x  x  P"
by (simp add: SPRAY_def)

lemma source_in_SPRAY:
  "SPRAY x  {}  P  SPRAY x. x  P"
using in_SPRAY_path by auto

lemma spray_event:
  "spray x  {}  x  "
proof (unfold spray_def)
  assume "{y. R  SPRAY x. y  R}  {}"
  then have "y. R  SPRAY x. y  R" by simp
  then have "SPRAY x  {}" by blast
  thus "x  " using SPRAY_event by simp
qed

lemma spray_nonevent:
  "x    spray x = {}"
using spray_event by auto

lemma in_spray_event:
  "y  spray x  y  "
proof (unfold spray_def)
  assume "y  {y. RSPRAY x. y  R}"
  then have "RSPRAY x. y  R" by (rule CollectD)
  then obtain R where path_R: "R  𝒫"
                  and y_inR: "y  R" using SPRAY_path by auto
  thus "y  " using in_path_event by simp
qed

lemma source_in_spray:
  "spray x  {}  x  spray x"
proof -
  assume nonempty_spray: "spray x  {}"
  have spray_eq: "spray x = {y. RSPRAY x. y  R}" using spray_def by simp
  then have ex_in_SPRAY_path: "y. RSPRAY x. y  R" using nonempty_spray by simp
  show "x  spray x" using ex_in_SPRAY_path spray_eq source_in_SPRAY by auto
qed


section "Primitives: Path (In)dependence"

text ‹
  "A subset of three paths of a SPRAY is dependent if there is a path which does not belong
  to the SPRAY and which contains one event from each of the three paths: we also say any
  one of the three paths is dependent on the other two.
  Otherwise the subset is independent." [Schutz97]
›

text ‹The definition of SPRAY› constrains $x, Q, R, S$ to be in $\E$ and $\P$.›
definition "dep3_event Q R S x
   card {Q,R,S} = 3  {Q,R,S}  SPRAY x
   (T𝒫. T  SPRAY x  QT{}  RT{}  ST{})"

definition "dep3_spray Q R S SPR  x. SPRAY x = SPR  dep3_event Q R S x"

definition "dep3 Q R S  x. dep3_event Q R S x"

text ‹Some very simple lemmas related to dep3_event›.›

(* Nice to have this on its own. *)
lemma dep3_nonspray:
  assumes "dep3_event Q R S x"
    shows "P𝒫. P  SPRAY x"
  by (metis assms dep3_event_def)

lemma dep3_path:
  assumes dep3_QRSx: "dep3 Q R S"
  shows "Q  𝒫" "R  𝒫" "S  𝒫"
  using assms dep3_event_def dep3_def SPRAY_path insert_subset by auto

lemma dep3_distinct:
  assumes dep3_QRSx: "dep3 Q R S"
  shows "Q  R" "Q  S" "R  S"
  using assms dep3_def dep3_event_def by (simp_all add: card_3_dist)

lemma dep3_is_event:
  "dep3_event Q R S x  x  "
using SPRAY_event dep3_event_def by auto

lemma dep3_event_old:
  "dep3_event Q R S x  Q  R  Q  S  R  S  Q  SPRAY x  R  SPRAY x  S  SPRAY x
                        (T𝒫. T  SPRAY x  (yQ. y  T)  (yR. y  T)  (yS. y  T))"
  by (rule iffI; unfold dep3_event_def, (simp add: card_3_dist), blast)

lemma dep3_event_permute [no_atp]:
  assumes "dep3_event Q R S x"
    shows "dep3_event Q S R x" "dep3_event R Q S x" "dep3_event R S Q x"
     "dep3_event S Q R x" "dep3_event S R Q x"
using dep3_event_old assms by auto

lemma dep3_permute [no_atp]:
  assumes "dep3 Q R S"
  shows "dep3 Q S R" "dep3 R Q S" "dep3 R S Q"
    and "dep3 S Q R" "dep3 S R Q"
  using dep3_event_permute dep3_def assms by meson+

text ‹
  "We next give recursive definitions of dependence and independence which will be used to
  characterize the concept of dimension. A path $T$ is dependent on the set of $n$ paths (where $n\geq3$)

  $$S = \left\lbrace Q_i \colon i = 1, 2, ..., n; Q_i \in \spray x\right\rbrace$$

  if it is dependent on two paths $S_1$ and $S_2$, where each of these two paths is dependent
  on some subset of $n - 1$ paths from the set $S$." [Schutz97]›

inductive dep_path :: "'a set  ('a set) set  bool" where
  dep_3: "dep3 T A B  dep_path T {A, B}"
| dep_n: "dep3 T S1 S2; dep_path S1 S'; dep_path S2 S''; S  SPRAY x;
           S'  S; S''  S; Suc (card S') = card S; Suc (card S'') = card S  dep_path T S"

lemma card_Suc_ex:
  assumes "card A = Suc (card B)" "B  A"
  shows "b. A = insert b B  bB"
proof -
  have "finite A" using assms(1) card_ge_0_finite card.infinite by fastforce
  obtain b where "bA-B"
    by (metis Diff_eq_empty_iff all_not_in_conv assms n_not_Suc_n subset_antisym)
  show "b. A = insert b B  bB"
  proof
    show "A = insert b B  bB"
      using  bA-B finite A assms
      by (metis DiffD1 DiffD2 Diff_insert_absorb Diff_single_insert card_insert_disjoint
        card_subset_eq insert_absorb rev_finite_subset)
  qed
qed

lemma union_of_subsets_by_singleton:
  assumes "Suc (card S') = card S" "Suc (card S'') = card S"
    and "S'  S''" "S'  S" "S''  S"
  shows "S'  S'' = S"
proof -
  obtain x y where x: "insert x S' = S" "xS'" and y: "insert y S'' = S" "yS''"
    using assms(1,2,4,5) by (metis card_Suc_ex)
  have "xy" using x y assms(3) by (metis insert_eq_iff)
  thus ?thesis using x(1) y(1) by blast
qed

lemma dep_path_card_2: "dep_path T S  card S  2"
  by (induct rule: dep_path.induct, simp add: dep3_def dep3_event_old, linarith)

text ‹
  "We also say that the set of $n + 1$ paths $S\cup\{T\}$ is a dependent set." [Schutz97]
  Starting from this constructive definition, the below gives an analytical one.
›

definition dep_set :: "('a set) set  bool" where
  "dep_set S  S'S. P(S-S'). dep_path P S'"

text ‹
  Notice that the relation between dep_set› and dep_path› becomes somewhat meaningless
  in the case where we apply dep_path› to an element of the set. This is because sets have no
  duplicate members, and we do not mirror the idea that scalar multiples of vectors linearly depend
  on those vectors: paths in a SPRAY are (in the $\mathbb{R}^4$ model) already equivalence classes
  of vectors that are scalar multiples of each other.
›

lemma dep_path_imp_dep_set:
  assumes "dep_path P S" "PS"
  shows "dep_set (insert P S)"
  using assms dep_set_def by auto

lemma dep_path_for_set_members:
  assumes "PS"
  shows "dep_set S = dep_set (insert P S)"
  by (simp add: assms(1) insert_absorb)

lemma dependent_superset:
  assumes "dep_set A" and "AB"
  shows "dep_set B"
  using assms dep_set_def
  by (meson Diff_mono dual_order.trans in_mono order_refl)

lemma path_in_dep_set:
  assumes "dep3 P Q R"
  shows "dep_set {P,Q,R}"
  using dep_3 assms dep3_def dep_set_def dep3_event_old
  by (metis DiffI insert_iff singletonD subset_insertI)

lemma path_in_dep_set2a:
  assumes "dep3 P Q R"
  shows "dep_path P {P,Q,R}"
proof
  let ?S' = "{P,R}"
  let ?S'' = "{P,Q}"
  have all_neq: "PQ" "PR" "RQ" using assms dep3_def dep3_event_old by auto
  show "dep3 P Q R" using assms dep3_event_def by (simp add: dep_3)
  show "dep_path Q ?S'" using assms dep3_event_permute(2) dep_3 dep3_def by meson
  show "dep_path R ?S''" using assms dep3_event_permute(4) dep_3 dep3_def by meson
  show "?S'  {P, Q, R}" by simp
  show "?S''  {P, Q, R}" by simp
  show "Suc (card ?S') = card {P, Q, R}" "Suc (card ?S'') = card {P, Q, R}"
    using all_neq card_insert_disjoint by auto
  show "{P, Q, R}  SPRAY (SOME x. dep3_event P Q R x)"
    using assms dep3_def dep3_event_def by (metis some_eq_ex)
qed

definition indep_set :: "('a set) set  bool" where
  "indep_set S  ¬ dep_set S"

lemma no_dep_in_indep: "indep_set S  ¬(T  S. dep_set T)"
  using indep_set_def dependent_superset by blast

lemma indep_set_alt_intro: "¬(T  S. dep_set T)  indep_set S"
  using indep_set_def by blast

lemma indep_set_alt: "indep_set S  ¬(S'  S. dep_set S')"
  using no_dep_in_indep indep_set_alt_intro by blast

lemma "dep_set S  indep_set S"
  by (simp add: indep_set_def)

section "Primitives: 3-SPRAY"

text ‹
  "We now make the following definition which enables us to specify the dimensions of Minkowski
  space-time. A SPRAY is a 3-SPRAY if:
    i) it contains four independent paths, and
   ii) all paths of the SPRAY are dependent on these four paths." [Schutz97]
›

definition n_SPRAY_basis :: "nat  'a set set  'a  bool"  where
  "n_SPRAY_basis n S x  SSPRAY x  card S = (Suc n)  indep_set S  (PSPRAY x. dep_path P S)"

definition n_SPRAY ("_-SPRAY _" [100,100]) where
  "n-SPRAY x  SSPRAY x. card S = (Suc n)  indep_set S  (PSPRAY x. dep_path P S)"

abbreviation "three_SPRAY x  3-SPRAY x"

lemma n_SPRAY_intro:
  assumes "SSPRAY x" "card S = (Suc n)" "indep_set S" "PSPRAY x. dep_path P S"
  shows "n-SPRAY x"
  using assms n_SPRAY_def by blast

lemma three_SPRAY_alt:
  "three_SPRAY x = (S1 S2 S3 S4.
    S1  S2  S1  S3  S1  S4  S2  S3  S2  S4  S3  S4
     S1  SPRAY x  S2  SPRAY x  S3  SPRAY x  S4  SPRAY x
     (indep_set {S1, S2, S3, S4})
     (SSPRAY x. dep_path S {S1,S2,S3,S4}))"
  (is "three_SPRAY x  ?three_SPRAY' x")
proof
  assume "three_SPRAY x"
  then obtain S where ns: "SSPRAY x" "card S = 4" "indep_set S" "PSPRAY x. dep_path P S"
    using n_SPRAY_def by auto
  then obtain S1 S2 S3 S4 where
    "S = {S1, S2, S3, S4}" and
    "S1  S2  S1  S3  S1  S4  S2  S3  S2  S4  S3  S4" and
    "S1  SPRAY x  S2  SPRAY x  S3  SPRAY x  S4  SPRAY x"
    using card_4_eq by (smt (verit) insert_subset ns)
  thus "?three_SPRAY' x"
    by (metis ns(3,4))
next
  assume "?three_SPRAY' x"
  then obtain S1 S2 S3 S4 where ns:
    "S1  S2  S1  S3  S1  S4  S2  S3  S2  S4  S3  S4"
    "S1  SPRAY x  S2  SPRAY x  S3  SPRAY x  S4  SPRAY x"
    "indep_set {S1, S2, S3, S4}"
    "SSPRAY x. dep_path S {S1,S2,S3,S4}"
    by metis
  show "three_SPRAY x"
    apply (intro n_SPRAY_intro[of "{S1, S2, S3, S4}"])
    by (simp add: ns)+
qed

lemma three_SPRAY_intro:
  assumes "S1  S2  S1  S3  S1  S4  S2  S3  S2  S4  S3  S4"
    and "S1  SPRAY x  S2  SPRAY x  S3  SPRAY x  S4  SPRAY x"
    and "indep_set {S1, S2, S3, S4}"
    and "SSPRAY x. dep_path S {S1,S2,S3,S4}"
  shows "three_SPRAY x"
  unfolding three_SPRAY_alt by (metis assms)

text ‹
  Lemma is_three_SPRAY› says "this set of sets of elements is a set of paths which is a 3-$\spray$".
  Lemma three_SPRAY_ge4› just extracts a bit of the definition.
›

definition is_three_SPRAY :: "('a set) set  bool" where
  "is_three_SPRAY S   x. S = SPRAY x  3-SPRAY x"

lemma three_SPRAY_ge4:
  assumes "three_SPRAY x"
  shows "Q1𝒫. Q2𝒫. Q3𝒫. Q4𝒫. Q1  Q2  Q1  Q3  Q1  Q4  Q2  Q3  Q2  Q4  Q3  Q4"
using assms three_SPRAY_alt SPRAY_path by meson

end (* MinkowskiPrimitive *)

section "MinkowskiBetweenness: O1-O5"

text ‹
  In O4, I have removed the requirement that $a \neq d$ in order to prove negative
  betweenness statements as Schutz does. For example, if we have $[abc]$
  and $[bca]$ we want to conclude $[aba]$ and claim "contradiction!", but
  we can't as long as we mandate that $a \neq d$.
›

locale MinkowskiBetweenness = MinkowskiPrimitive +
  fixes betw :: "'a  'a  'a  bool" ("[_;_;_]")
      (* O1 *) (*notice this is not only for events, but all things with same data type*)
  assumes abc_ex_path: "[a;b;c]  Q𝒫. a  Q  b  Q  c  Q"
      (* O2 *)
      and abc_sym: "[a;b;c]  [c;b;a]"
      (* O3, relaxed, as O3 can be proven from this. *)
      and abc_ac_neq: "[a;b;c]  a  c"
      (* O4 *)
      and abc_bcd_abd [intro]: "[a;b;c]; [b;c;d]  [a;b;d]"
      (* O5, relaxed; exhausting all six options is not necessary thanks to abc_sym. *)
      and some_betw: "Q  𝒫; a  Q; b  Q; c  Q; a  b; a  c; b  c
                [a;b;c]  [b;c;a]  [c;a;b]"
begin

text ‹
  The next few lemmas either provide the full axiom from the text derived from a new simpler
  statement, or provide some very simple fundamental additions which make sense to prove
  immediately before starting, usually related to set-level things that should be true which
  fix the type-level ambiguity of 'a.
›

lemma betw_events:
  assumes abc: "[a;b;c]"
  shows "a    b    c  "
proof -
  have "Q𝒫. a  Q  b  Q  c  Q" using abc_ex_path abc by simp
  thus ?thesis using in_path_event by auto
qed

text ‹This shows the shorter version of O5 is equivalent.›

lemma O5_still_O5 [no_atp]:
  "((Q  𝒫  {a,b,c}  Q  a    b    c    a  b  a  c  b  c)
      [a;b;c]  [b;c;a]  [c;a;b])
   =
   ((Q  𝒫  {a,b,c}  Q  a    b    c    a  b  a  c  b  c)
      [a;b;c]  [b;c;a]  [c;a;b]  [c;b;a]  [a;c;b]  [b;a;c])"
by (auto simp add: abc_sym)

lemma some_betw_xor:
   "Q  𝒫; a  Q; b  Q; c  Q; a  b; a  c; b  c
                ([a;b;c]  ¬ [b;c;a]  ¬ [c;a;b])
                  ([b;c;a]  ¬ [a;b;c]  ¬ [c;a;b])
                  ([c;a;b]  ¬ [a;b;c]  ¬ [b;c;a])"
by (meson abc_ac_neq abc_bcd_abd some_betw)

text ‹The lemma abc_abc_neq› is the full O3 as stated by Schutz.›
lemma abc_abc_neq:
  assumes abc: "[a;b;c]"
  shows "a  b  a  c  b  c"
using abc_sym abc_ac_neq assms abc_bcd_abd by blast


lemma abc_bcd_acd:
  assumes abc: "[a;b;c]"
      and bcd: "[b;c;d]"
  shows "[a;c;d]"
proof -
  have cba: "[c;b;a]" using abc_sym abc by simp
  have dcb: "[d;c;b]" using abc_sym bcd by simp
  have "[d;c;a]" using abc_bcd_abd dcb cba by blast
  thus ?thesis using abc_sym by simp
qed

lemma abc_only_cba:
  assumes "[a;b;c]"
    shows "¬ [b;a;c]" "¬ [a;c;b]" "¬ [b;c;a]" "¬ [c;a;b]"
using abc_sym abc_abc_neq abc_bcd_abd assms by blast+


section "Betweenness: Unreachable Subset Via a Path"

definition unreachable_subset_via :: "'a set  'a  'a set  'a  'a set" where
  "unreachable_subset_via Q Qa R x  {Qy. [x;Qy;Qa]  (RwR. Qa  unreach-on Q from Rw  Qy  unreach-on Q from Rw)}"

definition unreachable_subset_via_notation ("unreach-via _ on _ from _ to _" [100, 100, 100, 100] 100)
  where "unreach-via P on Q from a to x  unreachable_subset_via Q a P x"



section "Betweenness: Chains"
named_theorems chain_defs
named_theorems chain_alts

subsection "Locally ordered chains with indexing"

text ‹Definitions for Schutz's chains, with local order only.›

text‹A chain can be:
  (i) a set of two distinct events connected by a path, or ...›
definition short_ch :: "'a set  bool" where
  "short_ch X  card X = 2  (P𝒫. X  P)"

lemma short_ch_alt[chain_alts]:
  "short_ch X = (xX. yX. path_ex x y  ¬(zX. zx  zy))"
  "short_ch X = (x y. X = {x,y}  path_ex x y)"
  unfolding short_ch_def
  apply (simp add: card_2_iff', smt (verit, ccfv_SIG) in_mono subsetI)
  by (metis card_2_iff empty_subsetI insert_subset)

lemma short_ch_intros:
  "xX; yX; path_ex x y; ¬(zX. zx  zy)  short_ch X"
  "X = {x,y}; path_ex x y  short_ch X"
  by (auto simp: short_ch_alt)

lemma short_ch_path: "short_ch {x,y}  path_ex x y"
  unfolding short_ch_def by force

text ‹... a set of at least three events such that any three adjacent events are ordered.
  Notice infinite sets have card 0, because card gives a natural number always.›

definition local_long_ch_by_ord :: "(nat  'a)  'a set  bool" where
  "local_long_ch_by_ord f X  (infinite X  card X  3)  local_ordering f betw X"

lemma local_long_ch_by_ord_alt [chain_alts]:
  "local_long_ch_by_ord f X =
    (xX. yX. zX. xy  yz  xz  local_ordering f betw X)"
  (is "_ = ?ch f X")
proof
  assume asm: "local_long_ch_by_ord f X"
  {
    assume "card X  3"
    then have "x y z. xy  yz  xz  {x,y,z}X"
      apply (simp add: eval_nat_numeral)
      by (auto simp add: card_le_Suc_iff)
  } moreover {
    assume "infinite X"
    then have "x y z. xy  yz  xz  {x,y,z}X"
      using inf_3_elms bot.extremum by fastforce
  }
  ultimately show "?ch f X" using asm unfolding local_long_ch_by_ord_def by auto
next
  assume asm: "?ch f X"
  then obtain x y z where xyz: "{x,y,z}X  x  y  y  z  x  z"
    apply (simp add: eval_nat_numeral) by auto
  hence "card X  3  infinite X"
    apply (simp add: eval_nat_numeral)
    by (smt (z3) xyz card.empty card_insert_if card_subset finite.emptyI finite_insert insertE
      insert_absorb insert_not_empty)
  thus "local_long_ch_by_ord f X" unfolding local_long_ch_by_ord_def using asm by auto
qed

lemma short_xor_long:
  shows "short_ch Q  f. local_long_ch_by_ord f Q"
    and "local_long_ch_by_ord f Q  ¬ short_ch Q"
  unfolding chain_alts by (metis)+

text ‹Any short chain can have an ``ordering'' defined on it: this isn't the ternary ordering betw›
  that is used for triplets of elements, but merely an indexing function that fixes the
  ``direction'' of the chain, i.e. maps 0› to one element and 1› to the other.
  We define this in order to be able to unify chain definitions with those for long chains.
  Thus the indexing function $f$ of short_ch_by_ord f Q› has a similar status to the ordering
  on a long chain in many regards: e.g. it implies that $f({0 \dots |Q|-1}) \subseteq Q$.
›

definition short_ch_by_ord :: "(nat'a)  'a set  bool"
  where "short_ch_by_ord f Q  Q = {f 0, f 1}  path_ex (f 0) (f 1)"

lemma short_ch_equiv [chain_alts]: "f. short_ch_by_ord f Q  short_ch Q"
proof -
  { assume asm: "short_ch Q"
    obtain x y where xy: "{x,y}Q" "path_ex x y"
      using asm short_ch_alt(2) by (auto simp: short_ch_def)
    let ?f = "λn::nat. if n=0 then x else y"
    have "f. (x y. Q = {x, y}  f (0::nat) = x  f 1 = y  (Q. path Q x y))"
      apply (rule exI[of _ "?f"]) using asm xy short_ch_alt(2) by auto
  } moreover {
    fix f assume asm: "short_ch_by_ord f Q"
    have "card Q = 2" "(P𝒫. Q  P)"
      using asm short_ch_by_ord_def by auto
  } ultimately show ?thesis by (metis short_ch_by_ord_def short_ch_def)
qed

lemma short_ch_card:
  "short_ch_by_ord f Q  card Q = 2"
  "short_ch Q  card Q = 2"
  using short_ch_by_ord_def short_ch_def short_ch_equiv by auto

lemma short_ch_sym:
  assumes "short_ch_by_ord f Q"
  shows "short_ch_by_ord (λn. if n=0 then f 1 else f 0) Q"
  using assms unfolding short_ch_by_ord_def by auto

lemma short_ch_ord_in:
  assumes "short_ch_by_ord f Q"
  shows "f 0  Q" "f 1  Q"
  using assms unfolding short_ch_by_ord_def by auto


text ‹Does this restrict chains to lie on paths? Proven in TemporalOrderingOnPath›'s Interlude!›

definition ch_by_ord ("[__]") where
  "[fX]  short_ch_by_ord f X  local_long_ch_by_ord f X"

definition ch :: "'a set  bool" where "ch X  f. [fX]"

declare short_ch_def [chain_defs]
  and local_long_ch_by_ord_def [chain_defs]
  and ch_by_ord_def [chain_defs]
  and short_ch_by_ord_def [chain_defs]

text ‹We include alternative definitions in the chain_defs› set, because we do not want
  arbitrary orderings to appear on short chains. Unless an ordering for a short chain
  is explicitly written down by the user, we shouldn't introduce a short_ch_by_ord›
  when e.g. unfolding.›

lemma ch_alt[chain_defs]: "ch X  short_ch X  (f. local_long_ch_by_ord f X)"
  unfolding ch_def ch_by_ord_def using chain_defs short_ch_intros(2)
  by (smt (verit) short_ch_equiv)

text ‹
  Since $f(0)$ is always in the chain, and plays a special role particularly for infinite chains
  (as the 'endpoint', the non-finite edge) let us fix it straight in the definition.
  Notice we require both infinite X› and long_ch_by_ord›, thus circumventing infinite
  Isabelle sets having cardinality $0$.
›
definition infinite_chain :: "(nat  'a)  'a set  bool" where
  "infinite_chain f Q  infinite Q  [fQ]"

declare infinite_chain_def [chain_defs]

lemma infinite_chain_alt[chain_alts]:
  "infinite_chain f Q  infinite Q  local_ordering f betw Q"
  unfolding chain_defs by fastforce

definition infinite_chain_with :: "(nat  'a)  'a set  'a  bool" ("[__|_ ..]") where
  "infinite_chain_with f Q x  infinite_chain f Q  f 0 = x"

declare infinite_chain_with_def [chain_defs]

lemma "infinite_chain f Q  [fQ|f 0..]"
  by (simp add: infinite_chain_with_def)

definition finite_chain :: "(nat  'a)  'a set  bool" where
  "finite_chain f Q  finite Q  [fQ]"

declare finite_chain_def [chain_defs]

lemma finite_chain_alt[chain_alts]: "finite_chain f Q  short_ch_by_ord f Q  (finite Q  local_long_ch_by_ord f Q)"
  unfolding chain_defs by auto

definition finite_chain_with :: "(nat  'a)  'a set  'a  'a  bool" ("[__|_ .. _]") where
  "[fQ|x..y]  finite_chain f Q  f 0 = x  f (card Q - 1) = y"

declare finite_chain_with_def [chain_defs]

lemma "finite_chain f Q  [fQ|f 0 .. f (card Q - 1)]"
  by (simp add: finite_chain_with_def)

lemma finite_chain_with_alt [chain_alts]:
  "[fQ|x..z]  (short_ch_by_ord f Q  (card Q  3  local_ordering f betw Q)) 
    x = f 0  z = f (card Q - 1)"
  unfolding chain_defs
  by (metis card.infinite finite.emptyI finite.insertI not_numeral_le_zero)

lemma finite_chain_with_cases:
  assumes "[fQ|x..z]"
  obtains
    (short) "x = f 0" "z = f (card Q - 1)" "short_ch_by_ord f Q"
  | (long) "x = f 0" "z = f (card Q - 1)" "card Q  3" "local_long_ch_by_ord f Q"
  using assms finite_chain_with_alt by (meson local_long_ch_by_ord_def)


definition finite_long_chain_with:: "(nat'a)  'a set  'a  'a  'a  bool" ("[__|_.._.._]")
  where "[fQ|x..y..z]  [fQ|x..z]  xy  yz  yQ"

declare finite_long_chain_with_def [chain_defs]

lemma points_in_chain:
  assumes "[fQ|x..z]"
  shows "xQ  zQ"
  apply (cases rule: finite_chain_with_cases[OF assms])
  using short_ch_card(1) short_ch_ord_in by (simp add: chain_defs local_ordering_def[of f betw Q])+

lemma points_in_long_chain:
  assumes "[fQ|x..y..z]"
  shows "xQ" and "yQ" and "zQ"
  using points_in_chain finite_long_chain_with_def assms by meson+

lemma finite_chain_with_card_less3:
  assumes "[fQ|x..z]"
    and "card Q < 3"
  shows "short_ch_by_ord f Q" "z = f 1"
proof -
  show 1: "short_ch_by_ord f Q"
    using finite_chain_with_alt assms by simp
  thus "z = f 1"
    using assms(1) by (auto simp: eval_nat_numeral chain_defs)
qed

lemma ch_long_if_card_geq3:
  assumes "ch X"
      and "card X  3"
    shows "f. local_long_ch_by_ord f X"
proof -
  show "f. local_long_ch_by_ord f X"
  proof (rule ccontr)
    assume "f. local_long_ch_by_ord f X"
    hence "short_ch X"
      using assms(1) unfolding chain_defs by auto
    obtain x y z where "xX  yX  zX" and "xy  yz  xz"
      using assms(2) by (auto simp add: card_le_Suc_iff numeral_3_eq_3)
    thus False
      using short_ch X by (metis short_ch_alt(1))
  qed
qed

lemma ch_short_if_card_less3:
  assumes "ch Q"
      and "card Q < 3"
      and "finite Q"
    shows "f. short_ch_by_ord f Q"
  using short_ch_equiv finite_chain_with_card_less3
  by (metis assms ch_alt diff_is_0_eq' less_irrefl_nat local_long_ch_by_ord_def zero_less_diff)


lemma three_in_long_chain:
  assumes "local_long_ch_by_ord f X"
  obtains x y z where "xX" and "yX" and "zX" and "xy" and "xz" and "yz"
  using assms(1) local_long_ch_by_ord_alt by auto


lemma short_ch_card_2:
  assumes "ch_by_ord f X"
  shows "short_ch X  card X = 2"
  using assms unfolding chain_defs using card_2_iff' card_gt_0_iff by fastforce


lemma long_chain_card_geq:
  assumes "local_long_ch_by_ord f X" and fin: "finite X"
  shows "card X  3"
proof -
  obtain x y z where xyz: "xX" "yX" "zX" and neq: "xy" "xz" "yz"
    using three_in_long_chain assms by blast
  let ?S = "{x,y,z}"
  have "?S  X"
    by (simp add: xyz)
  moreover have "card ?S  3"
    using antisym x  y x  z y  z by auto
  ultimately show ?thesis
    by (meson neq fin three_subset)
qed


lemma fin_chain_card_geq_2:
  assumes "[fX|a..b]"
  shows "card X  2"
  using finite_chain_with_def apply (cases "short_ch X")
  using short_ch_card_2
  apply (metis dual_order.eq_iff short_ch_def)
  using assms chain_defs not_less by fastforce



section "Betweenness: Rays and Intervals"

text ‹
  ``Given any two distinct events $a,b$ of a path we define the segment
  $(ab) = \left\lbrace x : \ord{a}{x}{b},\; x \in ab \right\rbrace$" [Schutz97]
  Our version is a little different, because it is defined for any $a,b$ of type 'a›.
  Thus we can have empty set segments, while Schutz can prove (once he proves path density)
  that segments are never empty.
›
definition segment :: "'a  'a  'a set"
  where "segment a b  {x::'a. ab. [a;x;b]  xab  path ab a b}"

abbreviation is_segment :: "'a set  bool"
  where "is_segment ab  (a b. ab = segment a b)"

definition interval :: "'a  'a  'a set"
  where "interval a b  insert b (insert a (segment a b))"

abbreviation is_interval :: "'a set  bool"
  where "is_interval ab  (a b. ab = interval a b)"

definition prolongation :: "'a  'a  'a set"
  where "prolongation a b  {x::'a. ab. [a;b;x]  xab  path ab a b}"

abbreviation is_prolongation :: "'a set  bool"
  where "is_prolongation ab  a b. ab = prolongation a b"

text ‹
  I think this is what Schutz actually meant, maybe there is a typo in the text?
  Notice that b ∈ ray a b› for any $a$, always. Cf the comment on segment_def›.
  Thus ∃ray a b ≠ {}› is no guarantee that a path $ab$ exists.
›
definition ray :: "'a  'a  'a set"
  where "ray a b  insert b (segment a b  prolongation a b)"

abbreviation is_ray :: "'a set  bool"
  where "is_ray R  a b. R = ray a b"

definition is_ray_on :: "'a set  'a set  bool"
  where "is_ray_on R P  P𝒫  RP  is_ray R"

text ‹This is as in Schutz. Notice $b$ is not in the ray through $b$?›
definition ray_Schutz :: "'a  'a  'a set"
  where "ray_Schutz a b  insert a (segment a b  prolongation a b)"

lemma ends_notin_segment: "a  segment a b  b  segment a b"
  using abc_abc_neq segment_def by fastforce

lemma ends_in_int: "a  interval a b  b  interval a b"
  using interval_def by auto

lemma seg_betw: "x  segment a b  [a;x;b]"
  using segment_def abc_abc_neq abc_ex_path by fastforce

lemma pro_betw: "x  prolongation a b  [a;b;x]"
  using prolongation_def abc_abc_neq abc_ex_path by fastforce

lemma seg_sym: "segment a b = segment b a"
  using abc_sym segment_def by auto

lemma empty_segment: "segment a a = {}"
  by (simp add: segment_def)

lemma int_sym: "interval a b = interval b a"
  by (simp add: insert_commute interval_def seg_sym)

lemma seg_path:
  assumes "x  segment a b" (* thus segment ≠ {} *)
  obtains ab where "path ab a b" "segment a b  ab"
proof -
  obtain ab where "path ab a b"
    using abc_abc_neq abc_ex_path assms seg_betw
    by meson
  have "segment a b  ab"
    using path ab a b abc_ex_path path_unique seg_betw
    by fastforce
  thus ?thesis
    using path ab a b that by blast
qed

lemma seg_path2:
  assumes "segment a b  {}"
  obtains ab where "path ab a b" "segment a b  ab"
  using assms seg_path by force

text ‹Path density (theorem 17) will extend this by weakening the assumptions to segment a b ≠ {}›.›
lemma seg_endpoints_on_path:
  assumes "card (segment a b)  2" "segment a b  P" "P𝒫"
  shows "path P a b"
proof -
  have non_empty: "segment a b  {}" using assms(1) numeral_2_eq_2 by auto
  then obtain ab where "path ab a b" "segment a b  ab"
    using seg_path2 by force
  have "ab" by (simp add: path ab a b)
  obtain x y where "xsegment a b" "ysegment a b" "xy"
    using assms(1) numeral_2_eq_2
    by (metis card.infinite card_le_Suc0_iff_eq not_less_eq_eq not_numeral_le_zero)
  have "[a;x;b]"
    using x  segment a b seg_betw by auto
  have "[a;y;b]"
    using y  segment a b seg_betw by auto
  have "xP  yP"
    using x  segment a b y  segment a b assms(2) by blast
  have "xab  yab"
    using segment a b  ab x  segment a b y  segment a b by blast
  have "ab=P"
    using path ab a b x  P  y  P x  ab  y  ab x  y assms(3) path_unique by auto
  thus ?thesis
    using path ab a b by auto
qed

lemma pro_path:
  assumes "x  prolongation a b" (* thus prolongation ≠ {} *)
  obtains ab where "path ab a b" "prolongation a b  ab"
proof -
  obtain ab where "path ab a b"
    using abc_abc_neq abc_ex_path assms pro_betw
    by meson
  have "prolongation a b  ab"
    using path ab a b abc_ex_path path_unique pro_betw
    by fastforce
  thus ?thesis
    using path ab a b that by blast
qed

lemma ray_cases:
  assumes "x  ray a b"
  shows "[a;x;b]  [a;b;x]  x = b"
proof -
  have "xsegment a b  x prolongation a b  x=b"
    using assms ray_def by auto
  thus "[a;x;b]  [a;b;x]  x = b"
    using pro_betw seg_betw by auto
qed

lemma ray_path:
  assumes "x  ray a b" "xb"
  obtains ab where "path ab a b  ray a b  ab"
proof -
  let ?r = "ray a b"
  have "?r  {b}"
    using assms by blast
  have "ab. path ab a b  ray a b  ab"
  proof -
    have betw_cases: "[a;x;b]  [a;b;x]" using ray_cases assms
      by blast
    then obtain ab where "path ab a b"
      using abc_abc_neq abc_ex_path by blast
    have "?r  ab" using betw_cases
    proof (rule disjE)
      assume "[a;x;b]"
      show "?r  ab"
      proof
        fix x assume "x?r"
        show "xab"
          by (metis path ab a b x  ray a b abc_ex_path eq_paths ray_cases)
      qed
    next assume "[a;b;x]"
      show "?r  ab"
      proof
        fix x assume "x?r"
        show "xab"
          by (metis path ab a b x  ray a b abc_ex_path eq_paths ray_cases)
      qed
    qed
    thus ?thesis
      using path ab a b by blast
  qed
  thus ?thesis
    using that by blast
qed

end (* MinkowskiBetweenness *)

section "MinkowskiChain: O6"

text ‹O6 supposedly serves the same purpose as Pasch's axiom.›

locale MinkowskiChain = MinkowskiBetweenness +
  assumes O6: "{Q,R,S,T}  𝒫; card{Q,R,S} = 3; a  QR; b  QS; c  RS; dST; eRT; [b;c;d]; [c;e;a]
                fTQ. g X. [gX|a..f..b]"
begin

lemma O6_old: "Q  𝒫; R  𝒫; S  𝒫; T  𝒫; Q  R; Q  S; R  S; a  QR  b  QS  c  RS;
                dS. [b;c;d]  (eR. d  T  e  T  [c;e;a])
                fTQ. g X. [gX|a..f..b]"
  using O6[of Q R S T a b c] by (metis IntI card_3_dist empty_subsetI insert_subset)


section "Chains: (Closest) Bounds"

definition is_bound_f :: "'a  'a set  (nat'a)  bool" where
  "is_bound_f Qb Q f 
    i j ::nat. [fQ|(f 0)..]  (i<j  [f i; f j; Qb])"


definition is_bound where
  "is_bound Qb Q 
    f::(nat'a). is_bound_f Qb Q f"

text ‹
  $Q_b$ has to be on the same path as the chain $Q$.
  This is left implicit in the betweenness condition (as is $Q_b\in\E$).
  So this is equivalent to Schutz only if we also assume his axioms,
  i.e. the statement of the continuity axiom is no longer independent of other axioms.
›


definition all_bounds where
  "all_bounds Q = {Qb. is_bound Qb Q}"

definition bounded where
  "bounded Q   Qb. is_bound Qb Q"

lemma bounded_imp_inf:
  assumes "bounded Q"
  shows "infinite Q"
  using assms bounded_def is_bound_def is_bound_f_def chain_defs by meson


definition closest_bound_f where
  "closest_bound_f Qb Q f 
⌦‹Q is an infinite chain indexed by f bound by Qb
    is_bound_f Qb Q f 
⌦‹Any other bound must be further from the start of the chain than the closest bound›
    ( Qb'. (is_bound Qb' Q  Qb'  Qb)  [f 0; Qb; Qb'])"


definition closest_bound where
  "closest_bound Qb Q 
    f. is_bound_f Qb Q f
      ( Qb'. (is_bound Qb' Q  Qb'  Qb)  [f 0; Qb; Qb'])"

lemma "closest_bound Qb Q = (f. closest_bound_f Qb Q f)"
  unfolding closest_bound_f_def closest_bound_def by simp

end (*MinkowskiChain*)

section "MinkowskiUnreachable: I5-I7"

locale MinkowskiUnreachable = MinkowskiChain +
  assumes I5: "Q  𝒫; b  -Q  x y. {x,y}  unreach-on Q from b  x  y"
      and I6: "Q  𝒫; b  -Q; {Qx,Qz}  unreach-on Q from b; QxQz
       X f. [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))"
      and I7: "Q  𝒫; b  -Q; Qx  Q - unreach-on Q from b; Qy  unreach-on Q from b
                g X Qn. [gX|Qx..Qy..Qn]  Qn  Q - unreach-on Q from b"
begin

lemma two_in_unreach:
  "Q  𝒫; b  ; b  Q  xunreach-on Q from b. yunreach-on Q from b. x  y"
  using I5 by fastforce

lemma I6_old:
  assumes "Q  𝒫" "b  Q" "b  " "Qx  (unreach-on Q from b)" "Qz  (unreach-on Q from b)" "QxQz"
  shows "X. f. 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  QxX  QzX  (Qy. [Qx;Qy;Qz]  Qy  unreach-on Q from b))"
proof -
  from assms I6[of Q b Qx Qz] obtain f X
    where fX: "[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))"
    using DiffI Un_Diff_cancel by blast
  show ?thesis
  proof ((rule exI)+, intro conjI, rule_tac[4] ballI, rule_tac[5] impI; (intro conjI)?)
    show 1: "[fX]" "f 0 = Qx" "f (card X - 1) = Qz"
      using fX(1) chain_defs by meson+
    { 
      fix i assume i_asm: "i{1..card X - 1}"
      show 2: "f i  unreach-on Q from b"
        using fX(2) i_asm by fastforce
      show 3: "Qy. [f (i - 1);Qy;f i]  Qy  unreach-on Q from b"
        using fX(2) i_asm by blast
    } {
      assume X_asm: "short_ch X"
      show 4: "Qx  X" "Qz  X"
        using fX(1) points_in_chain by auto
      have "{1..card X-1} = {1}"
        using X_asm short_ch_alt(2) by force
      thus 5: "Qy. [Qx;Qy;Qz]  Qy  unreach-on Q from b"
        using fX(2) 1(2,3) by auto
    }
  qed
qed

lemma I7_old:
  assumes "Q  𝒫" "b  Q" "b  " "Qx  Q - unreach-on Q from b" "Qy  unreach-on Q from b"
  shows "g X Qn. [gX|Qx..Qy..Qn]  Qn  Q - unreach-on Q from b"
  using I7 assms by auto

lemma card_unreach_geq_2:
  assumes "Q𝒫" "b-Q"
  shows "2  card (unreach-on Q from b)  (infinite (unreach-on Q from b))"
  using DiffD1 assms(1) assms(2) card_le_Suc0_iff_eq two_in_unreach by fastforce


text ‹In order to more faithfully capture Schutz' definition of unreachable subsets via a path,
  we show that intersections of distinct paths are unique, and then define a new notation that
  doesn't carry the intersection of two paths around.›

lemma unreach_empty_on_same_path:
  assumes "P𝒫" "Q𝒫" "P=Q"
  shows "x. unreach-via P on Q from a to x = {}"
  unfolding unreachable_subset_via_notation_def unreachable_subset_via_def unreachable_subset_def
  by (simp add: assms(3))

definition unreachable_subset_via_notation_2 ("unreach-via _ on _ from _" [100, 100, 100] 100)
  where "unreach-via P on Q from a  unreachable_subset_via Q a P (THE x. xQP)"

lemma unreach_via_for_crossing_paths:
  assumes "P𝒫" "Q𝒫" "PQ = {x}"
  shows "unreach-via P on Q from a to x = unreach-via P on Q from a"
  unfolding unreachable_subset_via_notation_2_def is_singleton_def unreachable_subset_via_notation_def
  using the_equality assms by (metis Int_commute empty_iff insert_iff)

end

section "MinkowskiSymmetry: Symmetry"

locale MinkowskiSymmetry = MinkowskiUnreachable +
  assumes Symmetry: "{Q,R,S}  𝒫; card {Q,R,S} = 3;
               x  QRS; Qa  Q; Qa  x;
               unreach-via R on Q from Qa = unreach-via S on Q from Qa
                θ::'a'a.                               ⌦‹i) there is a map θ:ℰ⇒ℰ›
                     bij_betw (λP. {θ y | y. yP}) 𝒫 𝒫      ⌦‹ii) which induces a bijection Θ›
                      (yQ  θ y = y)                    ⌦‹iii) θ leaves Q invariant›
                      (λP. {θ y | y. yP}) R = S           ⌦‹iv) Θ maps R to S›"
begin

lemma Symmetry_old:
  assumes "Q  𝒫" "R  𝒫" "S  𝒫" "Q  R" "Q  S" "R  S"
    and "x  QRS" "Qa  Q" "Qa  x"
    and "unreach-via R on Q from Qa to x = unreach-via S on Q from Qa to x"
  shows "θ::'a'a. bij_betw (λP. {θ y | y. yP}) 𝒫 𝒫
                   (yQ  θ y = y)
                   (λP. {θ y | y. yP}) R = S"
proof -
  have QS: "QS = {x}" and QR: "QR = {x}"
    using assms(1-7) paths_cross_once by (metis Int_iff empty_iff insertE)+
  have "unreach-via R on Q from Qa = unreach-via R on Q from Qa to x"
    using unreach_via_for_crossing_paths QR by (simp add: Int_commute assms(1,2))
  moreover have "unreach-via S on Q from Qa = unreach-via S on Q from Qa to x"
    using unreach_via_for_crossing_paths QS by (simp add: Int_commute assms(1,3))
  ultimately show ?thesis
    using Symmetry assms by simp
qed

end

section "MinkowskiContinuity: Continuity"

locale MinkowskiContinuity = MinkowskiSymmetry +
  assumes Continuity: "bounded Q  Qb. closest_bound Qb Q"



section "MinkowskiSpacetime: Dimension (I4)"

locale MinkowskiSpacetime = MinkowskiContinuity +
  (* I4 *)
  assumes ex_3SPRAY [simp]: "  {}  x. 3-SPRAY x"
begin

(* substitute for I1, if I1 is omitted *)
(* lemma nonempty_events:
  "ℰ ≠ {}"
using ex_3SPRAY by blast *)

text ‹
  There exists an event by nonempty_events›, and by ex_3SPRAY› there is a three-$\spray$, which by
  three_SPRAY_ge4› means that there are at least four paths.
›
lemma four_paths:
  "Q1𝒫. Q2𝒫. Q3𝒫. Q4𝒫. Q1  Q2  Q1  Q3  Q1  Q4  Q2  Q3  Q2  Q4  Q3  Q4"
using nonempty_events ex_3SPRAY three_SPRAY_ge4 by blast

end

end