diff --git a/thys/Schutz_Spacetime/Minkowski.thy b/thys/Schutz_Spacetime/Minkowski.thy --- a/thys/Schutz_Spacetime/Minkowski.thy +++ b/thys/Schutz_Spacetime/Minkowski.thy @@ -1,892 +1,1233 @@ -(* Title: Schutz_Spacetime/Minkowski.thy - Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot - University of Edinburgh, 2021 -*) -theory Minkowski -imports Main 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]{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 \ - 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 - - -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" ("\ _ _" [100, 100] 100) where - "unreachable_subset Q b \ {x\Q. 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. \R\SPRAY 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. \R\SPRAY x. y \ R}" - then have "\R\SPRAY 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. \R\SPRAY x. y \ R}" using spray_def by simp - then have ex_in_SPRAY_path: "\y. \R\SPRAY 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 :: "'a set \ 'a set \ 'a set \ 'a \ bool" where - "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 \ (\y\Q. y \ T) \ (\y\R. y \ T) \ (\y\S. y \ T))" - -definition dep3_spray :: "'a set \ 'a set \ 'a set \ ('a set) set \ bool" where - "dep3_spray Q R S SPR \ \x. SPRAY x = SPR \ dep3_event Q R S x" - -definition dep3 :: "'a set \ 'a set \ 'a set \ bool" where - "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_event Q R S x" - shows "Q \ \

" "R \ \

" "S \ \

" -proof - - have "{Q,R,S} \ SPRAY x" using dep3_event_def using dep3_QRSx by simp - thus "Q \ \

" "R \ \

" "S \ \

" using SPRAY_path by auto -qed - -lemma dep3_is_event: - "dep3_event Q R S x \ x \ \" -using SPRAY_event dep3_event_def by auto - -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_def assms by auto - -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 \ 'a \ bool" where - dep_two: "dep3_event T A B x \ dep_path T {A, B} x" -| dep_n: "\S \ SPRAY x; card S \ 3; dep_path T {S1, S2} x; - S' \ S; S'' \ S; card S' = card S - 1; card S'' = card S - 1; - dep_path S1 S' x; dep_path S2 S'' x\ \ dep_path T S x" - -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 \ \x. \S'\S. \P\(S-S'). dep_path P S' x" - -lemma dependent_superset: - assumes "dep_set A" and "A\B" - shows "dep_set B" - using assms(1) assms(2) dep_set_def - by (meson Diff_mono dual_order.trans in_mono order_refl) - -lemma path_in_dep_set: - assumes "dep3_event P Q R x" - shows "dep_set {P,Q,R}" - using dep_two assms dep3_event_def dep_set_def - by (metis DiffI insertE insertI1 singletonD subset_insertI) - -lemma path_in_dep_set2: - assumes "dep3_event P Q R x" - shows "dep_path P {P,Q,R} x" -proof - let ?S1 = "Q" - let ?S2 = "R" - let ?S' = "{P,R}" - let ?S'' = "{P,Q}" - show "{P, Q, R} \ SPRAY x" using assms dep3_event_def by blast - show "3 \ card {P, Q, R}" using assms dep3_event_def by auto - show "dep_path P {?S1, ?S2} x" using assms dep3_event_def by (simp add: dep_two) - show "?S' \ {P, Q, R}" by simp - show "?S'' \ {P, Q, R}" by simp - show "card ?S' = card {P, Q, R} - 1" using assms dep3_event_def by auto - show "card ?S'' = card {P, Q, R} - 1" using assms dep3_event_def by auto - show "dep_path ?S1 ?S' x" by (simp add: assms dep3_event_permute(2) dep_two) - show "dep_path ?S2 ?S'' x" using assms dep3_event_permute(2,4) dep_two by blast -qed - - -definition indep_set :: "('a set) set \ bool" where - "indep_set S \ \(\T \ S. dep_set T)" - - -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 three_SPRAY :: "'a \ bool" where - "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}) - \ (\S\SPRAY x. dep_path S {S1,S2,S3,S4} x)" - -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 SPR \ \ x. SPR = SPRAY x \ three_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_def 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" - ("\ _ from _ via _ at _" [100, 100, 100, 100] 100) where - "unreachable_subset_via Q Qa R x \ {Qy. [[x Qy Qa]] \ (\Rw\R. Qa \ \ Q Rw \ Qy \ \ Q Rw)}" - - - -section "Betweenness: Chains" - -subsection "Totally ordered chains with indexing" - -definition short_ch :: "'a set \ bool" where - "short_ch X \ - \\EITHER two distinct events connected by a path\ - \x\X. \y\X. path_ex x y \ \(\z\X. z\x \ z\y)" - -text \Infinite sets have card 0, because card gives a natural number always.\ - -definition long_ch_by_ord :: "(nat \ 'a) \ 'a set \ bool" where - "long_ch_by_ord f X \ - \\OR at least three events such that any three events are ordered\ - \x\X. \y\X. \z\X. x\y \ y\z \ x\z \ ordering f betw X" - -text \Does this restrict chains to lie on paths? Proven in Ch3's Interlude!\ -definition ch_by_ord :: "(nat \ 'a) \ 'a set \ bool" where - "ch_by_ord f X \ short_ch X \ long_ch_by_ord f X" - -definition ch :: "'a set \ bool" where - "ch X \ \f. ch_by_ord f X" - -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 semifin_chain:: "(nat \ 'a) \ 'a \ 'a set \ bool" ("[_[_ ..]_]") where - "semifin_chain f x Q \ - infinite Q \ long_ch_by_ord f Q - \ f 0 = x" - -definition fin_long_chain:: "(nat \ 'a) \ 'a \ 'a \ 'a \ 'a set \ bool" - ("[_[_ .. _ .. _]_]") where - "fin_long_chain f x y z Q \ - x\y \ x\z \ y\z - \ finite Q \ long_ch_by_ord f Q - \ f 0 = x \ y\Q \ f (card Q - 1) = z" - -lemma index_middle_element: - assumes "[f[a..b..c]X]" - shows "\n. 0 n<(card X - 1) \ f n = b" -proof - - obtain n where n_def: "n < card X" "f n = b" - by (metis TernaryOrdering.ordering_def assms fin_long_chain_def long_ch_by_ord_def) - have "0 n<(card X - 1) \ f n = b" - using assms fin_long_chain_def n_def - by (metis Suc_pred' gr_implies_not0 less_SucE not_gr_zero) - thus ?thesis by blast -qed - -lemma fin_ch_betw: - assumes "[f[a..b..c]X]" - shows "[[a b c]]" -proof - - obtain nb where n_def: "nb\0" "nbn. f (card X - 1 - n))" -using ordering_sym assms(1) unfolding fin_long_chain_def long_ch_by_ord_def -by (metis (mono_tags, lifting) abc_sym diff_self_eq_0 diff_zero) - -lemma chain_sym: - assumes "[f[a..b..c]X]" - shows "[\n. f (card X - 1 - n)[c..b..a]X]" - using chain_sym_obtain [where f=f and a=a and b=b and c=c and X=X] - using assms(1) by blast - -definition fin_long_chain_2:: "'a \ 'a \ 'a \ 'a set \ bool" where - "fin_long_chain_2 x y z Q \ \f. [f[x..y..z]Q]" - -definition fin_chain:: "(nat \ 'a) \ 'a \ 'a \ 'a set \ bool" ("[_[_ .. _]_]") where - "fin_chain f x y Q \ - (short_ch Q \ x\Q \ y\Q \ x\y) - \ (\z\Q. [f[x..z..y]Q])" - -lemma points_in_chain: - assumes "[f[x..y..z]Q]" - shows "x\Q \ y\Q \ z\Q" -proof - - have "x\Q" - using ordering_def assms card_gt_0_iff emptyE fin_long_chain_def long_ch_by_ord_def - by metis - moreover have "y\Q" - using assms fin_long_chain_def - by auto - moreover have "z\Q" - using ordering_def assms card_gt_0_iff emptyE fin_long_chain_def long_ch_by_ord_def - by (metis (no_types, opaque_lifting) Suc_diff_1 lessI) - ultimately show ?thesis - by blast -qed - -lemma ch_long_if_card_ge3: - assumes "ch X" - and "card X \ 3" - shows "\f. long_ch_by_ord f X" -proof (rule ccontr) - assume "\f. long_ch_by_ord f X" - hence "short_ch X" - using assms(1) ch_by_ord_def ch_def - by auto - obtain x y z where "x\X \ y\X \ z\X" and "x\y \ y\z \ x\z" - using assms(2) - by (auto simp add: card_le_Suc_iff numeral_3_eq_3) - thus False - using \short_ch X\ short_ch_def - by metis -qed - -subsection "Locally ordered chains with indexing" - -text \Definition for Schutz-like chains, with local order only.\ -definition long_ch_by_ord2 :: "(nat \ 'a) \ 'a set \ bool" where - "long_ch_by_ord2 f X \ - \x\X. \y\X. \z\X. x\y \ y\z \ x\z \ ordering2 f betw X" - -subsection "Chains using betweenness" - -text \Old definitions of chains. Shown equivalent to \fin_long_chain_2\ in TemporalOrderOnPath.thy.\ - -definition chain_with :: "'a \ 'a \ 'a \ 'a set \ bool" ("[[.. _ .. _ .. _ ..]_]") where - "chain_with x y z X \ [[x y z]] \ x \ X \ y \ X \ z \ X \ (\f. ordering f betw X)" -definition finite_chain_with3 :: "'a \ 'a \ 'a \ 'a set \ bool" ("[[_ .. _ .. _]_]") where - "finite_chain_with3 x y z X \ [[..x..y..z..]X] \ \(\w\X. [[w x y]] \ [[y z w]])" - -lemma long_chain_betw: "[[..a..b..c..]X] \ [[a b c]]" -by (simp add: chain_with_def) - -lemma finite_chain3_betw: "[[a..b..c]X] \ [[a b c]]" -by (simp add: chain_with_def finite_chain_with3_def) - -definition finite_chain_with2 :: "'a \ 'a \ 'a set \ bool" ("[[_ .. _]_]") where - "finite_chain_with2 x z X \ \y\X. [[x..y..z]X]" - -lemma finite_chain2_betw: "[[a..c]X] \ \b. [[a b c]]" - using finite_chain_with2_def finite_chain3_betw by meson - - -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]] \ x\ab \ 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]] \ x\ab \ 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\\

\ R\P \ 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 "a\b" by (simp add: \path ab a b\) - obtain x y where "x\segment a b" "y\segment a b" "x\y" - 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 "x\P \ y\P" - using \x \ segment a b\ \y \ segment a b\ assms(2) by blast - have "x\ab \ y\ab" - 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 "x\segment 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" "x\b" - 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 "x\ab" - 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 "x\ab" - 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 \ \

; Q \ R; Q \ S; R \ S; a \ Q\R \ b \ Q\S \ c \ R\S; - \d\S. [[b c d]] \ (\e\R. d \ T \ e \ T \ [[c e a]])\ - \ \f\T\Q. \X. [[a..f..b]X]" -begin - - -section "Chains: (Closest) Bounds" - -definition is_bound_f :: "'a \ 'a set \ (nat\'a) \ bool" where - "is_bound_f Q\<^sub>b Q f \ - \i j ::nat. [f[(f 0)..]Q] \ (i [[(f i) (f j) Q\<^sub>b]])" - - -definition is_bound :: "'a \ 'a set \ bool" where - "is_bound Q\<^sub>b Q \ - \f::(nat\'a). is_bound_f Q\<^sub>b 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 :: "'a set \ 'a set" where - "all_bounds Q = {Q\<^sub>b. is_bound Q\<^sub>b Q}" - -definition bounded :: "'a set \ bool" where - "bounded Q \ \ Q\<^sub>b. is_bound Q\<^sub>b Q" - -text \Just to make sure Continuity is not too strong.\ -lemma bounded_imp_inf: - assumes "bounded Q" - shows "infinite Q" - using assms bounded_def is_bound_def is_bound_f_def semifin_chain_def by blast - - -definition closest_bound_f :: "'a \ 'a set \ (nat\'a) \ bool" where - "closest_bound_f Q\<^sub>b Q f \ -\<^cancel>\Q is an infinite chain indexed by f bound by Q\<^sub>b\ - is_bound_f Q\<^sub>b Q f \ -\<^cancel>\Any other bound must be further from the start of the chain than the closest bound\ - (\ Q\<^sub>b'. (is_bound Q\<^sub>b' Q \ Q\<^sub>b' \ Q\<^sub>b) \ [[(f 0) Q\<^sub>b Q\<^sub>b']])" - - -definition closest_bound :: "'a \ 'a set \ bool" where - "closest_bound Q\<^sub>b Q \ -\<^cancel>\Q is an infinite chain indexed by f bound by Q\<^sub>b\ - \f. is_bound_f Q\<^sub>b Q f \ -\<^cancel>\Any other bound must be further from the start of the chain than the closest bound\ - (\ Q\<^sub>b'. (is_bound Q\<^sub>b' Q \ Q\<^sub>b' \ Q\<^sub>b) \ [[(f 0) Q\<^sub>b Q\<^sub>b']])" - -end (*MinkowskiChain*) - -section "MinkowskiUnreachable: I5-I7" - -locale MinkowskiUnreachable = MinkowskiChain + - (* I5 *) - assumes two_in_unreach: "\Q \ \

; b \ \; b \ Q\ \ \x\\ Q b. \y\\ Q b. x \ y" - - and I6: "\Q \ \

; b \ Q; b \ \; Qx \ (\ Q b); Qz \ (\ Q b); Qx\Qz\ - \ \X. \f. ch_by_ord f X \ f 0 = Qx \ f (card X - 1) = Qz - \ (\i\{1 .. card X - 1}. (f i) \ \ Q b - \ (\Qy\\. [[(f(i-1)) Qy (f i)]] \ Qy \ \ Q b)) - \ (short_ch X \ Qx\X \ Qz\X \ (\Qy\\. [[Qx Qy Qz]] \ Qy \ \ Q b))" - and I7: "\Q \ \

; b \ Q; b \ \; Qx \ Q - \ Q b; Qy \ \ Q b\ - \ \g X Qn. [g[Qx..Qy..Qn]X] \ Qn \ Q - \ Q b" -begin - -lemma card_unreach_geq_2: - assumes "Q\\

" "b\\-Q" - shows "2 \ card (\ Q b) \ (infinite (\ Q b))" - using DiffD1 assms(1) assms(2) card_le_Suc0_iff_eq two_in_unreach by fastforce - -end - -section "MinkowskiSymmetry: Symmetry" - -locale MinkowskiSymmetry = MinkowskiUnreachable + - assumes Symmetry: "\Q \ \

; R \ \

; S \ \

; Q \ R; Q \ S; R \ S; - x \ Q\R\S; Q\<^sub>a \ Q; Q\<^sub>a \ x; - \ Q from Q\<^sub>a via R at x = \ Q from Q\<^sub>a via S at x\ - \ \\::'a\'a. \<^cancel>\i) there is a map \:\\\\ - bij_betw (\P. {\ y | y. y\P}) \

\

\<^cancel>\ii) which induces a bijection \\ - \ (y\Q \ \ y = y) \<^cancel>\iii) \ leaves Q invariant\ - \ (\P. {\ y | y. y\P}) R = S \<^cancel>\iv) \ maps R to S\" - - -section "MinkowskiContinuity: Continuity" - -locale MinkowskiContinuity = MinkowskiSymmetry + - assumes Continuity: "bounded Q \ (\Q\<^sub>b. closest_bound Q\<^sub>b Q)" - - - -section "MinkowskiSpacetime: Dimension (I4)" - -locale MinkowskiSpacetime = MinkowskiContinuity + - (* I4 *) - assumes ex_3SPRAY [simp]: "\\ \ {}\ \ \x\\. three_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 +(* 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]{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: "Q\R \ {}" + shows "\!a\\. Q\R = {a}" +proof - + have ab_inQR: "\a\\. a\Q\R" using QR_nonempty in_path_event path_Q by auto + then obtain a where a_event: "a \ \" and a_inQR: "a \ Q\R" by auto + have "Q\R = {a}" + proof (rule ccontr) + assume "Q\R \ {a}" + then have "\b\Q\R. 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 \ {x\Q. 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. \R\SPRAY 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. \R\SPRAY x. y \ R}" + then have "\R\SPRAY 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. \R\SPRAY x. y \ R}" using spray_def by simp + then have ex_in_SPRAY_path: "\y. \R\SPRAY 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 \ Q\T\{} \ R\T\{} \ S\T\{})" + +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 \ (\y\Q. y \ T) \ (\y\R. y \ T) \ (\y\S. 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 \ b\B" +proof - + have "finite A" using assms(1) card_ge_0_finite card.infinite by fastforce + obtain b where "b\A-B" + by (metis Diff_eq_empty_iff all_not_in_conv assms n_not_Suc_n subset_antisym) + show "\b. A = insert b B \ b\B" + proof + show "A = insert b B \ b\B" + using \b\A-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" "x\S'" and y: "insert y S'' = S" "y\S''" + using assms(1,2,4,5) by (metis card_Suc_ex) + have "x\y" 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" "P\S" + shows "dep_set (insert P S)" + using assms dep_set_def by auto + +lemma dep_path_for_set_members: + assumes "P\S" + shows "dep_set S = dep_set (insert P S)" + by (simp add: assms(1) insert_absorb) + +lemma dependent_superset: + assumes "dep_set A" and "A\B" + 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: "P\Q" "P\R" "R\Q" 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 \ S\SPRAY x \ card S = (Suc n) \ indep_set S \ (\P\SPRAY x. dep_path P S)" + +definition n_SPRAY ("_-SPRAY _" [100,100]) where + "n-SPRAY x \ \S\SPRAY x. card S = (Suc n) \ indep_set S \ (\P\SPRAY x. dep_path P S)" + +abbreviation "three_SPRAY x \ 3-SPRAY x" + +lemma n_SPRAY_intro: + assumes "S\SPRAY x" "card S = (Suc n)" "indep_set S" "\P\SPRAY 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}) + \ (\S\SPRAY 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: "S\SPRAY x" "card S = 4" "indep_set S" "\P\SPRAY x. dep_path P S" + using n_SPRAY_def by auto + then obtain S\<^sub>1 S\<^sub>2 S\<^sub>3 S\<^sub>4 where + "S = {S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4}" and + "S\<^sub>1 \ S\<^sub>2 \ S\<^sub>1 \ S\<^sub>3 \ S\<^sub>1 \ S\<^sub>4 \ S\<^sub>2 \ S\<^sub>3 \ S\<^sub>2 \ S\<^sub>4 \ S\<^sub>3 \ S\<^sub>4" and + "S\<^sub>1 \ SPRAY x \ S\<^sub>2 \ SPRAY x \ S\<^sub>3 \ SPRAY x \ S\<^sub>4 \ 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 S\<^sub>1 S\<^sub>2 S\<^sub>3 S\<^sub>4 where ns: + "S\<^sub>1 \ S\<^sub>2 \ S\<^sub>1 \ S\<^sub>3 \ S\<^sub>1 \ S\<^sub>4 \ S\<^sub>2 \ S\<^sub>3 \ S\<^sub>2 \ S\<^sub>4 \ S\<^sub>3 \ S\<^sub>4" + "S\<^sub>1 \ SPRAY x \ S\<^sub>2 \ SPRAY x \ S\<^sub>3 \ SPRAY x \ S\<^sub>4 \ SPRAY x" + "indep_set {S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4}" + "\S\SPRAY x. dep_path S {S\<^sub>1,S\<^sub>2,S\<^sub>3,S\<^sub>4}" + by metis + show "three_SPRAY x" + apply (intro n_SPRAY_intro[of "{S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4}"]) + 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 "\S\SPRAY 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] \ (\Rw\R. 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 = (\x\X. \y\X. path_ex x y \ \(\z\X. z\x \ z\y))" + "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: + "\x\X; y\X; path_ex x y; \(\z\X. z\x \ z\y)\ \ 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 = + (\x\X. \y\X. \z\X. x\y \ y\z \ x\z \ 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. x\y \ y\z \ x\z \ {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. x\y \ y\z \ x\z \ {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 + "[f\X] \ short_ch_by_ord f X \ local_long_ch_by_ord f X" + +definition ch :: "'a set \ bool" where "ch X \ \f. [f\X]" + +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 \ [f\Q]" + +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 \ [f\Q|f 0..]" + by (simp add: infinite_chain_with_def) + +definition finite_chain :: "(nat \ 'a) \ 'a set \ bool" where + "finite_chain f Q \ finite Q \ [f\Q]" + +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 + "[f\Q|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 \ [f\Q|f 0 .. f (card Q - 1)]" + by (simp add: finite_chain_with_def) + +lemma finite_chain_with_alt [chain_alts]: + "[f\Q|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 "[f\Q|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 "[f\Q|x..y..z] \ [f\Q|x..z] \ x\y \ y\z \ y\Q" + +declare finite_long_chain_with_def [chain_defs] + +lemma points_in_chain: + assumes "[f\Q|x..z]" + shows "x\Q \ z\Q" + 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 "[f\Q|x..y..z]" + shows "x\Q" and "y\Q" and "z\Q" + using points_in_chain finite_long_chain_with_def assms by meson+ + +lemma finite_chain_with_card_less3: + assumes "[f\Q|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 "x\X \ y\X \ z\X" and "x\y \ y\z \ x\z" + 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 "x\X" and "y\X" and "z\X" and "x\y" and "x\z" and "y\z" + 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: "x\X" "y\X" "z\X" and neq: "x\y" "x\z" "y\z" + 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 "[f\X|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] \ x\ab \ 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] \ x\ab \ 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\\

\ R\P \ 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 "a\b" by (simp add: \path ab a b\) + obtain x y where "x\segment a b" "y\segment a b" "x\y" + 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 "x\P \ y\P" + using \x \ segment a b\ \y \ segment a b\ assms(2) by blast + have "x\ab \ y\ab" + 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 "x\segment 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" "x\b" + 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 "x\ab" + 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 "x\ab" + 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 \ Q\R; b \ Q\S; c \ R\S; d\S\T; e\R\T; [b;c;d]; [c;e;a]\ + \ \f\T\Q. \g X. [g\X|a..f..b]" +begin + +lemma O6_old: "\Q \ \

; R \ \

; S \ \

; T \ \

; Q \ R; Q \ S; R \ S; a \ Q\R \ b \ Q\S \ c \ R\S; + \d\S. [b;c;d] \ (\e\R. d \ T \ e \ T \ [c;e;a])\ + \ \f\T\Q. \g X. [g\X|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 Q\<^sub>b Q f \ + \i j ::nat. [f\Q|(f 0)..] \ (i [f i; f j; Q\<^sub>b])" + + +definition is_bound where + "is_bound Q\<^sub>b Q \ + \f::(nat\'a). is_bound_f Q\<^sub>b 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 = {Q\<^sub>b. is_bound Q\<^sub>b Q}" + +definition bounded where + "bounded Q \ \ Q\<^sub>b. is_bound Q\<^sub>b 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 Q\<^sub>b Q f \ +\<^cancel>\Q is an infinite chain indexed by f bound by Q\<^sub>b\ + is_bound_f Q\<^sub>b Q f \ +\<^cancel>\Any other bound must be further from the start of the chain than the closest bound\ + (\ Q\<^sub>b'. (is_bound Q\<^sub>b' Q \ Q\<^sub>b' \ Q\<^sub>b) \ [f 0; Q\<^sub>b; Q\<^sub>b'])" + + +definition closest_bound where + "closest_bound Q\<^sub>b Q \ + \f. is_bound_f Q\<^sub>b Q f + \ (\ Q\<^sub>b'. (is_bound Q\<^sub>b' Q \ Q\<^sub>b' \ Q\<^sub>b) \ [f 0; Q\<^sub>b; Q\<^sub>b'])" + +lemma "closest_bound Q\<^sub>b Q = (\f. closest_bound_f Q\<^sub>b 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; Qx\Qz\ + \ \X f. [f\X|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. [g\X|Qx..Qy..Qn] \ Qn \ Q - unreach-on Q from b" +begin + +lemma two_in_unreach: + "\Q \ \

; b \ \; b \ Q\ \ \x\unreach-on Q from b. \y\unreach-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)" "Qx\Qz" + 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 \ Qx\X \ Qz\X \ (\Qy\\. [Qx;Qy;Qz] \ Qy \ unreach-on Q from b))" +proof - + from assms I6[of Q b Qx Qz] obtain f X + where fX: "[f\X|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: "[f\X]" "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. [g\X|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. x\Q\P)" + +lemma unreach_via_for_crossing_paths: + assumes "P\\

" "Q\\

" "P\Q = {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 \ Q\R\S; Q\<^sub>a \ Q; Q\<^sub>a \ x; + unreach-via R on Q from Q\<^sub>a = unreach-via S on Q from Q\<^sub>a\ + \ \\::'a\'a. \<^cancel>\i) there is a map \:\\\\ + bij_betw (\P. {\ y | y. y\P}) \

\

\<^cancel>\ii) which induces a bijection \\ + \ (y\Q \ \ y = y) \<^cancel>\iii) \ leaves Q invariant\ + \ (\P. {\ y | y. y\P}) R = S \<^cancel>\iv) \ maps R to S\" +begin + +lemma Symmetry_old: + assumes "Q \ \

" "R \ \

" "S \ \

" "Q \ R" "Q \ S" "R \ S" + and "x \ Q\R\S" "Q\<^sub>a \ Q" "Q\<^sub>a \ x" + and "unreach-via R on Q from Q\<^sub>a to x = unreach-via S on Q from Q\<^sub>a to x" + shows "\\::'a\'a. bij_betw (\P. {\ y | y. y\P}) \

\

+ \ (y\Q \ \ y = y) + \ (\P. {\ y | y. y\P}) R = S" +proof - + have QS: "Q\S = {x}" and QR: "Q\R = {x}" + using assms(1-7) paths_cross_once by (metis Int_iff empty_iff insertE)+ + have "unreach-via R on Q from Q\<^sub>a = unreach-via R on Q from Q\<^sub>a 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 Q\<^sub>a = unreach-via S on Q from Q\<^sub>a 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 \ \Q\<^sub>b. closest_bound Q\<^sub>b 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 diff --git a/thys/Schutz_Spacetime/ROOT b/thys/Schutz_Spacetime/ROOT --- a/thys/Schutz_Spacetime/ROOT +++ b/thys/Schutz_Spacetime/ROOT @@ -1,13 +1,13 @@ -chapter AFP - -session "Schutz_Spacetime" (AFP) = HOL + - options [timeout = 600] - theories [document = false] - Util - theories - TernaryOrdering - Minkowski - TemporalOrderOnPath - document_files - "root.bib" - "root.tex" +chapter AFP + +session "Schutz_Spacetime" (AFP) = "HOL-Library" + + options [timeout = 600] + theories [document = false] + Util + theories + TernaryOrdering + Minkowski + TemporalOrderOnPath + document_files + "root.bib" + "root.tex" diff --git a/thys/Schutz_Spacetime/TemporalOrderExtra.thy b/thys/Schutz_Spacetime/TemporalOrderExtra.thy new file mode 100644 --- /dev/null +++ b/thys/Schutz_Spacetime/TemporalOrderExtra.thy @@ -0,0 +1,88 @@ +(* Title: Schutz_Spacetime/TemporalOrderExtra.thy + Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot + University of Edinburgh, 2021 +*) +theory TemporalOrderExtra + imports Main Minkowski TernaryOrdering TemporalOrderOnPath "HOL-Algebra.Order" +begin + +section \Ternary-to-binary\ + +context MinkowskiSpacetime +begin + +\ \The cases we need to cover where x <= y if we take a < b as our ordering basis: + [abxy], [axby], [axyb], + [xaby], [xayb], [xyab]. +\ +abbreviation betw_gorder :: "'a \ 'a \ 'a set \ 'a gorder" where + "betw_gorder a b Q \ \carrier = Q, eq = (=), + le = (\x y. x = y + \ \Base case. Everything from here is < really.\ + \ x = a \ y = b + \ \a < y\b case (\ fact is covered by [___].\ + \ x = a \ ([a;b;y] \ [a;y;b]) + \ \b < y case. Can't have [b a y].\ + \ x = b \ [a;b;y] + \ \x < a\ + \ y = a \ [x;a;b] + \ \x < b\ + \ y = b \ ([x;a;b] \ [a;x;b]) + \ \4-event cases.\ + \ [a;b;x;y] \ [a;x;b;y] \ [a;x;y;b] + \ [x;a;b;y] \ [x;a;y;b] \ [x;y;a;b] + ) + \" + +lemma betw_total_order_on_path: + fixes a b :: 'a + and Q :: "'a set" + assumes "Q \ \

" + and "a \ Q" + and "b \ Q" + and "a \ b" + shows "total_order (betw_gorder a b Q)" (is "total_order ?o") +proof + show "\x. x \ carrier ?o \ x .=\<^bsub>?o\<^esub> x" + by simp +next + show "\x y. \x .=\<^bsub>?o\<^esub> y; x \ carrier ?o; y \ carrier ?o\ \ y .=\<^bsub>?o\<^esub> x" + by simp +next + show "\x y z. \x .=\<^bsub>?o\<^esub> y; y .=\<^bsub>?o\<^esub> z; x \ carrier ?o; + y \ carrier ?o; z \ carrier ?o\ + \ x .=\<^bsub>?o\<^esub> z" + by simp +next + show "\x. x \ carrier ?o \ x \\<^bsub>?o\<^esub> x" + by simp +next + show "\x y. \x \\<^bsub>?o\<^esub> y; y \\<^bsub>?o\<^esub> x; x \ carrier ?o; y \ carrier ?o\ \ x .=\<^bsub>?o\<^esub> y" + by (smt abc_abc_neq abcd_dcba_only(18) betw4_sym betw4_weak eq_object.select_convs(1) + gorder.select_convs(1)) +next + show "\x y z. + \x \\<^bsub>?o\<^esub> y; y \\<^bsub>?o\<^esub> z; x \ carrier ?o; + y \ carrier ?o; z \ carrier ?o\ + \ x \\<^bsub>?o\<^esub> z" + by (smt (z3) abc_acd_abd abc_acd_bcd abc_bcd_acd abc_only_cba(1,4) assms(1) + gorder.select_convs(1) partial_object.select_convs(1) some_betw2) +next + show "\x y z w. + \x .=\<^bsub>?o\<^esub> y; z .=\<^bsub>?o\<^esub> w; x \ carrier ?o; + y \ carrier ?o; z \ carrier ?o; w \ carrier ?o\ + \ (x \\<^bsub>?o\<^esub> z) = (y \\<^bsub>?o\<^esub> w)" + by simp +next + show "(.=\<^bsub>?o\<^esub>) = (=)" + by simp +next + show "\x y. \x \ carrier (betw_gorder a b Q); y \ carrier (betw_gorder a b Q)\ + \ x \\<^bsub>betw_gorder a b Q\<^esub> y \ y \\<^bsub>betw_gorder a b Q\<^esub> x" + by (smt (z3) abc_sym assms gorder.select_convs(1) partial_object.select_convs(1) + some_betw2 some_betw4b) +qed + +end + +end \ No newline at end of file diff --git a/thys/Schutz_Spacetime/TemporalOrderOnPath.thy b/thys/Schutz_Spacetime/TemporalOrderOnPath.thy --- a/thys/Schutz_Spacetime/TemporalOrderOnPath.thy +++ b/thys/Schutz_Spacetime/TemporalOrderOnPath.thy @@ -1,7720 +1,6677 @@ -(* Title: Schutz_Spacetime/TemporalOrderOnPath.thy - Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot - University of Edinburgh, 2021 -*) -theory TemporalOrderOnPath -imports Main Minkowski TernaryOrdering -begin - -text \ - In Schutz \cite[pp.~18-30]{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. -\ - -text \Disable list syntax.\ -no_translations - "[x, xs]" == "x#[xs]" - "[x]" == "x#[]" -no_syntax - \ \list Enumeration\ - "_list" :: "args => 'a list" ("[(_)]") -no_notation Cons (infixr "#" 65) -no_notation Nil ("[]") - - -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 three_in_set3: - assumes "card X \ 3" - obtains x y z where "x\X" and "y\X" and "z\X" and "x\y" and "x\z" and "y\z" - using assms by (auto simp add: card_le_Suc_iff numeral_3_eq_3) - -lemma paths_cross_once: - assumes path_Q: "Q \ \

" - and path_R: "R \ \

" - and Q_neq_R: "Q \ R" - and QR_nonempty: "Q\R \ {}" - shows "\!a\\. Q\R = {a}" -proof - - have ab_inQR: "\a\\. a\Q\R" using QR_nonempty in_path_event path_Q by auto - then obtain a where a_event: "a \ \" and a_inQR: "a \ Q\R" by auto - have "Q\R = {a}" - proof (rule ccontr) - assume "Q\R \ {a}" - then have "\b\Q\R. 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 - -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. long_ch_by_ord f {a,b,c}" - using abc_abc_neq ord_ordered long_ch_by_ord_def assms - by (smt insertI1 insert_commute) - -lemma between_chain: "[[a b c]] \ ch {a,b,c}" -proof - - assume "[[a b c]]" - hence "\f. ordering f betw {a,b,c}" - by (simp add: abc_abc_neq ord_ordered) - hence "\f. long_ch_by_ord f {a,b,c}" - using \[[a b c]]\ abc_abc_neq long_ch_by_ord_def by auto - thus ?thesis - by (simp add: ch_by_ord_def ch_def) -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. 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) - qed - hence "\f. long_ch_by_ord f {a,b,c,d}" - using \[[a b c]]\ abc_abc_neq long_ch_by_ord_def by auto - thus ?thesis - by (simp add: ch_by_ord_def ch_def) - 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. - Splitting it up into the proof of: - ``there is a betweenness relation for each ordered triple", and - ``all events of a chain are distinct" - The first part is obvious with total chains (using \ordering\), and will be proved using the - local definition as well (\ordering2\), following Schutz' proof. - The second part is proved as injectivity of the indexing function (see \index_injective\). -\ - -text \ - For the case of two-element chains: the elements are distinct by definition, - and the statement on ordering is void (respectively, \False \ P\ for any \P\). -\ - -(* Theorem 2 for total chains *) -theorem order_finite_chain: - assumes chX: "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)]]" - by (metis chX long_ch_by_ord_def ordered_nats ordering_ord_ijk) - -(* Theorem 2 (with helper lemmas for induction) for local chains *) -lemma thm2_ind1: - assumes chX: "long_ch_by_ord2 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 ?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 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 "ii 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: "long_ch_by_ord2 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 long_ch_by_ord2_def ordering2_ord_ijk) - 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 long_ch_by_ord2_def ordering2_def - by (meson asm(3) less_imp_diff_less) - 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: "long_ch_by_ord2 f X" - and finiteX: "finite X" - and ordered_nats: "0 k 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_chain2: - assumes chX: "long_ch_by_ord2 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: "0\i \ i jk. 0 k [[(f (k-1)) (f k) (f l)]]" - using thm2_ind2b chX finiteX ordered_nats - by blast - have "j j=l-1" - using ordered_nats by linarith - thus ?thesis - proof - assume "jj < 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 - - -lemma three_in_long_chain2: - assumes "long_ch_by_ord2 f X" - obtains x y z where "x\X" and "y\X" and "z\X" and "x\y" and "x\z" and "y\z" - using assms(1) long_ch_by_ord2_def by auto - - -lemma short_ch_card_2: - assumes "ch_by_ord f X" - shows "short_ch X \ card X = 2" - by (metis assms card_2_iff' ch_by_ord_def long_ch_by_ord_def short_ch_def) - - -lemma long_chain2_card_geq: - assumes "long_ch_by_ord2 f X" and fin: "finite X" - shows "card X \ 3" -proof - - obtain x y z where xyz: "x\X" "y\X" "z\X" and neq: "x\y" "x\z" "y\z" - using three_in_long_chain2 assms(1) 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 "[f[a..b]X]" - shows "card X \ 2" - using fin_chain_def apply (cases "short_ch X") - using short_ch_card_2 - apply (metis card_2_iff' dual_order.eq_iff short_ch_def) - using assms fin_long_chain_def not_less by fastforce - - -(*TODO: make index_neq_ results use this, much simpler!?*) -theorem (*2ii*) index_injective: - fixes i::nat and j::nat - assumes chX: "long_ch_by_ord2 f X" - and finiteX: "finite X" - and indices: "i f j" -proof (cases) - assume "Suc i < j" - then have "[[(f i) (f (Suc(i))) (f j)]]" - using order_finite_chain2 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 "0Suc i = j\ \Suc j = card X\) - have "card X \ 3" - using assms(1) finiteX long_chain2_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_chain2 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_chain2 by blast - thus ?thesis - using abc_abc_neq by blast - qed -qed - -end (* context MinkowskiBetweenness *) - - -section \Finite chain equivalence: local \\\ global\ -context MinkowskiBetweenness begin - - -lemma ch_equiv1: - assumes "long_ch_by_ord f X" "finite X" - shows "long_ch_by_ord2 f X" - using assms - unfolding long_ch_by_ord_def long_ch_by_ord2_def ordering_def ordering2_def - by (metis lessI) - - -lemma ch_equiv2: - assumes "long_ch_by_ord2 f X" "finite X" - shows "long_ch_by_ord f X" - using order_finite_chain2 assms - unfolding long_ch_by_ord_def long_ch_by_ord2_def ordering_def ordering2_def - apply safe by blast - - -lemma ch_equiv: - assumes "finite X" - shows "long_ch_by_ord f X \ long_ch_by_ord2 f X" - using ch_equiv1 ch_equiv2 assms by blast - - -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 "\f\de\ab. [[a f b]]" -proof - - have "\f\ab \ de. \X. [[a..f..b]X]" - 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 [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 finite_chain3_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 \ (\f\de\ab. [[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 "\f\ab \ de. \X. [[a..f..b]X]" - 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 [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 finite_chain3_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 "(\f\de\(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. [[a..f..b]X]" - 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 [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 finite_chain3_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\\

. \a\Q. a \ \" - by simp - -lemma same_empty_unreach: - "\Q \ \

; a \ Q\ \ \ Q a = {}" -apply (unfold unreachable_subset_def) -by simp - -lemma same_path_reachable: - "\Q \ \

; a \ Q; b \ Q\ \ a \ Q - \ Q 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 - \ R 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 - \ R 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 - \ Q 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 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\{}" -proof - - obtain a where "a\\" using nonempty_events by blast - have "a\Q \ a\Q" by auto - thus ?thesis - proof - assume "a\Q" - thus ?thesis by blast - next - assume "a\Q" - then obtain b where "b\\ Q 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" -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 "\x\\ Q b. \y\\ Q b. x \ y" - shows "\x\Q. \y\Q. 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: "b\Q" "b\\" - and unreach: "Q\<^sub>x \ \ Q b" "Q\<^sub>z \ \ Q b" "Q\<^sub>x \ Q\<^sub>z" - and X_def: "ch_by_ord f X" "f 0 = Q\<^sub>x" "f (card X - 1) = Q\<^sub>z" - "(\i\{1 .. card X - 1}. (f i) \ \ Q b \ (\Q\<^sub>y\\. [[(f(i-1)) Q\<^sub>y (f i)]] \ Q\<^sub>y \ \ Q b))" - "(short_ch X \ Q\<^sub>x\X \ Q\<^sub>z\X \ (\Q\<^sub>y\\. [[Q\<^sub>x Q\<^sub>y Q\<^sub>z]] \ Q\<^sub>y \ \ Q b))" - shows "X\Q" - (* by (smt X_def(1) chain_on_path eq_paths in_Q in_X path_Q subset_eq unreach(3)) *) - (* smt has a really easy time of this, but no other method does (legacy code from thm13) *) -proof - - have in_Q: "Q\<^sub>x\Q \ Q\<^sub>z\Q" - using unreachable_subset_def unreach(1,2) by blast - have fin_X: "finite X" - using unreach(3) not_less X_def by fastforce - { - assume "short_ch X" - hence ?thesis - by (metis X_def(5) in_Q short_ch_def subsetI unreach(3)) - } moreover { - assume asm: "long_ch_by_ord f X" - have ?thesis - proof - fix x assume "x\X" - then obtain i where "f i = x" "i < card X" - using asm unfolding ch_by_ord_def long_ch_by_ord_def ordering_def - using fin_X by auto - show "x\Q" - proof (cases) - assume "x=Q\<^sub>x \ x=Q\<^sub>z" - thus ?thesis - using in_Q by blast - next - assume "\(x=Q\<^sub>x \ x=Q\<^sub>z)" - hence "x\Q\<^sub>x" "x\Q\<^sub>z" by linarith+ - have "i>0" - using X_def(2) \x\Q\<^sub>x\ \f i = x\ gr_zeroI by force - have "if i = x\ \i < card X\ \x \ Q\<^sub>z\ less_imp_diff_less less_SucE - by (metis Suc_pred' cancel_comm_monoid_add_class.diff_cancel) - have "[[Q\<^sub>x (f i) Q\<^sub>z]]" - using X_def(2,3) \0 < i\ \i < card X - 1\ asm fin_X order_finite_chain - by auto - thus ?thesis - by (simp add: \f i = x\ betw_b_in_path in_Q path_Q unreach(3)) - qed - qed - } - ultimately show ?thesis - using X_def(1) ch_by_ord_def by blast -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 \ R\Q\ \ a \ b" -by blast - -lemma nocross_events_neq: - "\Q \ \

; R \ \

; a \ Q; b \ R; R\Q = {}\ \ 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: \\X Q0 Qm Qn. [[Q0 .. Qm .. Qn]X] \ Q0 = ?Qx \ Qm = ?Qy \ Qn \ ?Q - \ ?Q ?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 - \ Q b" - and Qy_unreachable: "Qy \ \ Q b" - shows "\Qz\Q - \ Q b. [[Qx Qy Qz]] \ Qx \ Qz" - using assms I7 order_finite_chain fin_long_chain_def - by (metis fin_ch_betw) - -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 "R\Q" - using R_neq_Q path_R path_sub_events by auto - obtain c where "c\R" "c\Q" - 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 "a\Q" "b\Q" - 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 "\b\Q. 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 "\a\Q. \b\Q. 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 - \ R 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 \ \ R a" - using two_in_unreach a_inQ in_path_event path_Q path_R by blast - then obtain c where c_reachable: "c \ R - \ R 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 \ \ R 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 \ \ Q b \ a \ Q" -using unreachable_subset_def by simp - -lemma (in MinkowskiUnreachable) unreach_equiv: - "\Q \ \

; R \ \

; a \ Q; b \ R; a \ \ Q b\ \ b \ \ R 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 \ \ ae b" - using two_in_unreach b_inQ in_path_event path_Q path_ae by blast - then have b_unreachable: "b \ \ Q f" using unreach_equiv - by (metis (mono_tags, lifting) CollectD b_inQ path_Q unreachable_subset_def) - have a_reachable: "a \ Q - \ Q 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 "\c\Q. [[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 "\c\Q. \d\Q. [[a b c]] \ [[a b d]] \ c\d" - 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 "\a\X. \b\X. a \ b \ (\c\X. 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 (* jep: sledgehammer had trouble with this case without explicitly telling it what - assumptions to use, for some reason, and even then only one prover figured out "by auto". *) - 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: "\c\Y. (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: "Q\X" - obtains a b where "a\b" and "a\Q" and "b\Q" and "\c\Q. (a\c \ b\c) \ [[a c b]]" -proof - - obtain n where "n\0" and "card Q = n+3" - using card_Q nat_le_iff_add - by auto - then obtain a b where "a\b" and "a\Q" and "b\Q" and "\c\Q. (a\c \ b\c) \ [[a c b]]" - using finite_path_has_ends assms \n\0\ - 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 "n\1" - using \card Q \ 0\ by linarith - then consider (n1) "n=1" | (n2) "n=2" | (n3) "n\3" - 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 "a\b" and "a\Q" and "b\Q" - using ge2_events_lax assms(1) by blast - then obtain c where "c\Q" and "c\a" and "c\b" - 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 "Q\Q" by simp - then obtain a b where "a\Q" and "b\Q" and "a\b" - and acb: "\c\Q. (c\a \ c\b) \ [[a c b]]" - using obtain_fin_path_ends card_Q fin_Q assms(1) - by metis - then obtain x where "[[a b x]]" and "x\Q" - 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\de. [[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 "dc\ce" - 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 "fe\ea" - 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)) - 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 - proof (cases) - assume a'b'c': "[[a' b' c']]" - then have c'b'a': "[[c' b' a']]" using abc_sym 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 path_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 "\x\a'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 False using ca'b abc_only_cba by blast - next - assume b'c'a': "[[b' c' a']]" - then have a'c'b': "[[a' c' b']]" using abc_sym by simp - have nopath_a'cb': "\ (\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 ab_inQ: "a \ Q \ b \ Q" - using eq_paths - by (metis ab'c abc_a'b'c'_neq betw_a_in_path betw_c_in_path ca'b) - thus False using c_inQ path_Q tri_abc triangle_diff_paths by blast - qed - then have tri_a'cb': "\ a' c b'" - by (smt ab'c abc_ex_path_unique b'c'a' ca'b paths_tri) - obtain bc' where path_bc': "path bc' b c'" - using abc_a'b'c'_neq abc_ex_path_unique bc'a - by blast - obtain b'c where path_b'c: "path b'c b' c" using tri_a'cb' triangle_paths(3) by blast - then have "\x\b'c. [[b' x c]] \ [[b c' x]]" - using collinearity2 [where a = b' and b = c and c = a' - and e = c' and d = b and de = bc'] - bc'a betw_b_in_path a'c'b' path_bc' tri_a'cb' - by (meson ca'b triangle_permutes(5)) - then obtain x where x_in_b'c: "x \ b'c" - and b'xc: "[[b' x c]]" - and bc'x: "[[b c' x]]" by blast - have a_in_bc': "a \ bc'" using bc'a betw_c_in_path path_bc' by blast - have a_in_b'c: "a \ b'c" using ab'c betw_a_in_path path_b'c by blast - have bc'_b'c_distinct: "bc' \ b'c" - using a_in_bc' path_b'c path_bc' tri_abc triangle_diff_paths by blast - have "bc' \ b'c = {a}" - using paths_cross_at bc'_b'c_distinct a_in_b'c a_in_bc' path_b'c path_bc' by auto - then have "x = a" using bc'x betw_c_in_path path_bc' x_in_b'c by auto - then have "[[b' a c]]" using b'xc by auto - thus False using ab'c abc_only_cba by blast - next - assume c'a'b': "[[c' a' b']]" - then have b'a'c': "[[b' a' c']]" using abc_sym by simp - have nopath_c'ab': "\ (\Q\\

. c' \ Q \ a \ Q \ b' \ Q)" - proof (rule notI) - assume "\Q\\

. c' \ Q \ a \ Q \ b' \ Q" - then obtain Q where path_Q: "Q \ \

" - and c'_inQ: "c' \ Q" - and a_inQ: "a \ Q" - and b'_inQ: "b' \ Q" by blast - then have bc_inQ: "b \ Q \ c \ Q" - using eq_paths ab'c abc_a'b'c'_neq bc'a betw_a_in_path betw_c_in_path by blast - thus False using a_inQ path_Q tri_abc triangle_diff_paths by blast - qed - then have tri_a'cb': "\ b' a c'" - by (smt bc'a abc_ex_path_unique c'a'b' ab'c paths_tri) - obtain ca' where path_ca': "path ca' c a'" - using abc_a'b'c'_neq abc_ex_path_unique ca'b - by blast - obtain c'a where path_c'a: "path c'a c' a" using tri_a'cb' triangle_paths(3) by blast - then have "\x\c'a. [[c' x a]] \ [[c a' x]]" - using collinearity2 [where a = c' and b = a and c = b' - and e = a' and d = c and de = ca'] - ab'c b'a'c' betw_b_in_path path_ca' tri_a'cb' triangle_permutes(5) by blast - then obtain x where x_in_c'a: "x \ c'a" - and c'xa: "[[c' x a]]" - and ca'x: "[[c a' x]]" by blast - have b_in_ca': "b \ ca'" using betw_c_in_path ca'b path_ca' by blast - have b_in_c'a: "b \ c'a" using bc'a betw_a_in_path path_c'a by auto - have ca'_c'a_distinct: "ca' \ c'a" - using b_in_c'a path_c'a path_ca' tri_abc triangle_diff_paths by blast - have "ca' \ c'a = {b}" - using b_in_c'a b_in_ca' ca'_c'a_distinct path_c'a path_ca' paths_cross_at by auto - then have "x = b" using betw_c_in_path ca'x path_ca' x_in_c'a by auto - then have "[[c' b a]]" using c'xa by auto - thus False using abc_only_cba bc'a by blast - qed - 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'\ \ ab 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 \ e\R \ d'\e)" - by blast - hence "\(\R\\

. e\R \ d'\R)" - using d'_def(3) by blast - moreover have "ab\\

\ e\\ \ e\ab" - by (simp add: e_event e_notin_S path_ab) - ultimately have "d'\ \ ab 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 "e\ab" - using S_neq_ab a_inS e_inS e_neq_a eq_paths path_S path_ab by auto - have "ab\\

\ e\ab" - using S_neq_ab a_inS e_inS e_neq_a eq_paths path_S path_ab - by auto - have "b \ ab - \ ab e" - using cross_in_reachable path_ab path_be - by blast - have "d \ \ ab e" - using no_de abd path_ab e_event \e\ab\ - 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'\ \ ab e" "b\d'" "e\d'" - using unreachable_set_bounded \b \ ab - \ ab e\ \d \ \ ab e\ e_event \e\ab\ path_ab - by (metis DiffE) - then obtain d'e where "path d'e d' e" - using unreachable_bounded_path_only e_event \e\ab\ 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]] \ d\ab" - 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 "c\d" 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]]" "f\c'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 \ \ c'e 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 - \ c'e b" - using c'd'_def cross_in_reachable path_ab by blast - moreover have "b\\ \ b\c'e" - using \f \ \ c'e b\ betw_events c'd'_def same_empty_unreach by auto - ultimately obtain f' where f'_def: "[[c' f f']]" "f'\c'e" "f'\ \ c'e b" "c'\f'" "b\f'" - 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 "f\c'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 \ \ c'e b" - using no_path assms(1,4-9) path_S_2 \f\c'e\ \c'\ab\ \d'\ab\ - abc_abc_neq betw_events eq_paths unreachable_bounded_path_only - by metis - moreover have "c' \ c'e - \ c'e b" - using c'd'_def cross_in_reachable path_ab \c' \ ab\ by blast - moreover have "b\\ \ b\c'e" - using \f \ \ c'e b\ betw_events c'd'_def same_empty_unreach by auto - ultimately obtain f' where f'_def: "[[c' f f']]" "f'\c'e" "f'\ \ c'e b" "c'\f'" "b\f'" - 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 "\e\S. 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 \ \ S b" - and y_in_unreach: "y \ \ S b" - and x_neq_y: "x \ y" - and z_in_reach: "z \ S - \ S 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 - \ S 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: "f\c'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 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 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 - (* This goes through so easily because the overlapping betweenness axiom is an intro rule. - Adding it here to more closely match Schutz. *) - 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 - (* Schutz implicitly assumes this. *) - 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: "b\c" - shows "[[a b c]] \ [[a c b]]" -proof - - obtain P where P_def: "P\\

" "a\P" "b\P" "d\P" - using abd abc_ex_path by blast - hence "c\P" - 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 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 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 between_chain by auto - show ?thesis - using abc_betw ch1 ch2 ch3 by (metis insert_commute) -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 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 - - -end (* context MinkowskiSpacetime *) - - -section "Interlude - Chains and Equivalences" -text \ - This section is meant for our alternative definitions of chains, and proofs of equivalence. - If we want to regain full independence of our axioms, we probably need to shuffle a few - things around. Some of this may be redundant, but is kept for compatibility with legacy proofs. - - Three definitions are given (cf `Betweenness: Chains` in Minkowski.thy): - - one relying on explicit betweenness conditions - - one relying on a total ordering and explicit indexing - - one equivalent to the above except for use of the weaker, local-only ordering2 -\ - -context MinkowskiChain begin - -subsection "Proofs for totally ordered index-chains" - -subsubsection "General results" - -lemma inf_chain_is_long: - assumes "semifin_chain f x X" - shows "long_ch_by_ord f X \ f 0 = x \ infinite X" -proof - - have "infinite X \ card X \ 2" using card.infinite by simp - hence "semifin_chain f x X \ long_ch_by_ord f X" - using long_ch_by_ord_def semifin_chain_def short_ch_def - by simp - thus ?thesis using assms semifin_chain_def by blast -qed - -text \A reassurance that the starting point $x$ is implied.\ -lemma long_inf_chain_is_semifin: - assumes "long_ch_by_ord f X \ infinite X" - shows "\ x. [f[x..]X]" - by (simp add: assms semifin_chain_def) - -lemma endpoint_in_semifin: - assumes "semifin_chain f x X" - shows "x\X" - using assms semifin_chain_def zero_into_ordering inf_chain_is_long long_ch_by_ord_def - by (metis finite.emptyI) - -lemma three_in_long_chain: - assumes "long_ch_by_ord f X" and fin: "finite X" - obtains x y z where "x\X" and "y\X" and "z\X" and "x\y" and "x\z" and "y\z" - using assms(1) long_ch_by_ord_def by auto - -subsubsection "Index-chains lie on paths" - -lemma all_aligned_on_semifin_chain: - assumes "[f[x..]X]" - and a: "y\X" and b:"z\X" and xy: "x\y" and xz: "x\z" and yz: "y\z" - shows "[[x y z]] \ [[x z y]]" -proof - - obtain n\<^sub>y n\<^sub>z where "f n\<^sub>y = y" and "f n\<^sub>z = z" - by (metis TernaryOrdering.ordering_def a assms(1) b inf_chain_is_long long_ch_by_ord_def) - have "(0y \ n\<^sub>yz) \ (0z \ n\<^sub>zy)" - using \f n\<^sub>y = y\ \f n\<^sub>z = z\ assms less_linear semifin_chain_def xy xz yz by auto - hence "[[(f 0) (f n\<^sub>y) (f n\<^sub>z)]] \ [[(f 0) (f n\<^sub>z) (f n\<^sub>y)]]" - using ordering_def assms(1) long_ch_by_ord_def semifin_chain_def - by (metis long_ch_by_ord_def) - thus "[[x y z]] \ [[x z y]]" - using \f n\<^sub>y = y\ \f n\<^sub>z = z\ assms semifin_chain_def by auto - qed - - -lemma semifin_chain_on_path: - assumes "[f[x..]X]" - shows "\P\\

. X\P" -proof - - obtain y where "y\X" and "y\x" - using assms inf_chain_is_long - by (metis Diff_iff all_not_in_conv finite_Diff2 finite_insert infinite_imp_nonempty insert_iff) - have path_exists: "\P\\

. path P x y" - proof - - obtain e where "e\X" and "e\x" and "e\y" and "[[x y e]] \ [[x e y]]" - using all_aligned_on_semifin_chain inf_chain_is_long long_ch_by_ord_def assms - ordering_def lessI \y \ X\ \y \ x\ finite.emptyI finite_insert - finite_subset insert_iff subsetI - by smt - obtain P where "path P x y" - using \[[x y e]] \ [[x e y]]\ abc_abc_neq abc_ex_path - by blast - show ?thesis - using \path P x y\ - by blast - qed - obtain P where "path P x y" - using path_exists - by blast - have "X\P" - proof - fix e - assume "e\X" - show "e\P" - proof - - have "e=x \ e=y \ (e\x \ e\y)" by auto - moreover { assume "e\x \ e\y" - have "[[x y e]] \ [[x e y]]" - using all_aligned_on_semifin_chain assms - \e \ X\ \e \ x \ e \ y\ \y \ X\ \y \ x\ - by blast - hence ?thesis - using \path P x y\ abc_ex_path path_unique - by blast - } moreover { assume "e=x" - have ?thesis - by (simp add: \e = x\ \path P x y\) - } moreover { assume "e=y" - have "e\P" - by (simp add: \e = y\ \path P x y\) - } - ultimately show ?thesis by blast - qed - qed - thus ?thesis - using \path P x y\ - by blast -qed - - -lemma card2_either_elt1_or_elt2: - assumes "card X = 2" and "x\X" and "y\X" and "x\y" - and "z\X" and "z\x" - shows "z=y" -by (metis assms card_2_iff') - -lemma short_chain_on_path: - assumes "short_ch X" - shows "\P\\

. X\P" -proof - - obtain x y where "x\y" and "x\X" and "y\X" - using assms short_ch_def by auto - obtain P where "path P x y" - using \x \ X\ \x \ y\ \y \ X\ assms short_ch_def - by metis - have "X\P" - proof - fix z - assume "z\X" - show "z\P" - proof cases - assume "z=x" - show "z\P" using \path P x y\ by (simp add: \z=x\) - next - assume "z\x" - have "z=y" - using \x\X\ \y\X\ \z\x\ \z\X\ \x\y\ assms short_ch_def - by metis - thus "z\P" using \path P x y\ by (simp add: \z=y\) - qed - qed - thus ?thesis - using \path P x y\ by blast -qed - - -lemma all_aligned_on_long_chain: - assumes "long_ch_by_ord f X" and "finite X" - and a: "x\X" and b: "y\X" and c:"z\X" and xy: "x\y" and xz: "x\z" and yz: "y\z" -shows "[[x y z]] \ [[x z y]] \ [[z x y]]" -proof - - obtain n\<^sub>x n\<^sub>y n\<^sub>z where fx: "f n\<^sub>x = x" and fy: "f n\<^sub>y = y" and fz: "f n\<^sub>z = z" - and xx: "n\<^sub>x < card X" and yy: "n\<^sub>y < card X" and zz: "n\<^sub>z < card X" - proof - - assume a1: "\n\<^sub>x n\<^sub>y n\<^sub>z. \f n\<^sub>x = x; f n\<^sub>y = y; f n\<^sub>z = z; n\<^sub>x < card X; n\<^sub>y < card X; n\<^sub>z < card X\ \ thesis" - obtain nn :: "'a set \ (nat \ 'a) \ 'a \ nat" where - "\a A f p pa. (a \ A \ \ ordering f p A \ f (nn A f a) = a) - \ (infinite A \ a \ A \ \ ordering f pa A \ nn A f a < card A)" - by (metis (no_types) ordering_def) - then show ?thesis - using a1 by (metis a assms(1) assms(2) b c long_ch_by_ord_def) - qed - have less_or: "(n\<^sub>xy \ n\<^sub>yz) \ (n\<^sub>xz \ n\<^sub>zy) \ (n\<^sub>zx \ n\<^sub>xy) \ - (n\<^sub>zy \ n\<^sub>yx) \ (n\<^sub>yz \ n\<^sub>zx) \ (n\<^sub>yx \ n\<^sub>xz)" - using fx fy fz assms less_linear - by metis - have int_imp_1: "(n\<^sub>xy \ n\<^sub>yz) \ long_ch_by_ord f X \ n\<^sub>z < card X \ [[(f n\<^sub>x) (f n\<^sub>y) (f n\<^sub>z)]]" - using assms long_ch_by_ord_def ordering_def - by metis - hence "[[(f n\<^sub>x) (f n\<^sub>y) (f n\<^sub>z)]] \ [[(f n\<^sub>x) (f n\<^sub>z) (f n\<^sub>y)]] \ [[(f n\<^sub>z) (f n\<^sub>x) (f n\<^sub>y)]] \ - [[(f n\<^sub>z) (f n\<^sub>y) (f n\<^sub>x)]] \ [[(f n\<^sub>y) (f n\<^sub>z) (f n\<^sub>x)]] \ [[(f n\<^sub>y) (f n\<^sub>x) (f n\<^sub>z)]]" - proof - - have f1: "\n na nb. \ n < na \ \ nb < n \ \ na < card X \ [[(f nb) (f n) (f na)]]" - by (metis (no_types) ordering_def \long_ch_by_ord f X\ long_ch_by_ord_def) - then have f2: "\ n\<^sub>z < n\<^sub>y \ \ n\<^sub>x < n\<^sub>z \ [[x z y]]" - using fx fy fz yy - by blast - have "\ n\<^sub>x < n\<^sub>y \ \ n\<^sub>z < n\<^sub>x \ [[z x y]]" - using f1 fx fy fz yy by blast - then show ?thesis - using f2 f1 fx fy fz less_or xx zz by auto - qed - hence "[[x y z]] \ [[x z y]] \ [[z x y]] \ - [[z y x]] \ [[y z x]] \ [[y x z]]" - using fx fy fz assms semifin_chain_def long_ch_by_ord_def - by metis - thus ?thesis - using abc_sym - by blast -qed - - -lemma long_chain_on_path: - assumes "long_ch_by_ord f X" and "finite X" - shows "\P\\

. X\P" -proof - - obtain x y where "x\X" and "y\X" and "y\x" - using long_ch_by_ord_def assms - by (metis (mono_tags, opaque_lifting)) - obtain z where "z\X" and "x\z" and "y\z" - using long_ch_by_ord_def assms - by metis - have "[[x y z]] \ [[x z y]] \ [[z x y]]" - using all_aligned_on_long_chain assms - using \x \ X\ \x \ z\ \y \ X\ \y \ x\ \y \ z\ \z \ X\ - by auto - then have path_exists: "\P\\

. path P x y" - using all_aligned_on_long_chain abc_ex_path - by (metis \y \ x\) - obtain P where "path P x y" - using path_exists - by blast - have "X\P" - proof - fix e - assume "e\X" - show "e\P" - proof - - have "e=x \ e=y \ (e\x \ e\y)" by auto - moreover { - assume "e\x \ e\y" - have "[[x y e]] \ [[x e y]] \ [[e x y]]" - using all_aligned_on_long_chain all_aligned_on_long_chain assms - \e \ X\ \e \ x \ e \ y\ \y \ X\ \y \ x\ \x\X\ - by metis - hence ?thesis - using \path P x y\ abc_ex_path path_unique - by blast - } - moreover { - assume "e=x" - have ?thesis - by (simp add: \e = x\ \path P x y\) - } - moreover { - assume "e=y" - have "e\P" - by (simp add: \e = y\ \path P x y\) - } - ultimately show ?thesis by blast - qed - qed - thus ?thesis - using \path P x y\ - by blast -qed - -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 chain_on_path: - assumes "ch_by_ord f X" - shows "\P\\

. X\P" - using assms ch_by_ord_def - using semifin_chain_on_path long_chain_on_path short_chain_on_path long_inf_chain_is_semifin - by meson - -subsubsection "More general results" - -(* In fact, it is xor. *) -lemma ch_some_betw: "\x \ X; y \ X; z \ X; x \ y; x \ z; y \ z; ch X\ - \ [[x y z]] \ [[y x z]] \ [[y z x]]" -proof - - assume asm: "x \ X" "y \ X" "z \ X" "x \ y" "x \ z" "y \ z" "ch X" - { - fix f assume f_def: "long_ch_by_ord f X" - assume evts: "x \ X" "y \ X" "z \ X" "x \ y" "x \ z" "y \ z" - assume ords: "\ [[x y z]]" "\ [[y z x]]" - obtain P where "X\P" "P\\

" - using chain_on_path f_def ch_by_ord_def - by meson - have "[[y x z]]" - proof - - have f1: "\A Aa a. \ A \ Aa \ (a::'a) \ A \ a \ Aa" - by blast - have f2: "y \ P" - using \X \ P\ evts(2) by blast - have f3: "x \ P" - using f1 by (metis \X \ P\ evts(1)) - have "z \ P" - using \X \ P\ evts(3) by blast - then show ?thesis - using f3 f2 by (metis some_betw_xor \P \ \

\ abc_sym evts(4,5,6) ords) - qed - } - thus ?thesis - unfolding ch_def long_ch_by_ord_def ch_by_ord_def ordering_def short_ch_def - using asm ch_by_ord_def ch_def short_ch_def - by (metis \\f. \long_ch_by_ord f X; x \ X; y \ X; z \ X; x \ y; x \ z; y \ z; - \ [[x y z]]; \ [[y z x]]\ \ [[y x z]]\) -qed - - -lemma ch_all_betw_f: - assumes "[f[x..yy..z]X]" and "y\X" and "y\x" and "y\z" - shows "[[x y z]]" -proof (rule ccontr) - assume asm: "\ [[x y z]]" - obtain Q where "Q\\

" and "x\Q \ y\Q \ z\Q" - using chain_on_path assms ch_by_ord_def asm fin_ch_betw fin_long_chain_def - by auto - hence "[[x y z]] \ [[y x z]] \ [[y z x]]" - using some_betw assms - by (metis abc_sym fin_long_chain_def) - hence "[[y x z]] \ [[x z y]]" - using asm abc_sym - by blast - thus False - using fin_long_chain_def long_ch_by_ord_def asm assms fin_ch_betw - by (metis (no_types, opaque_lifting)) -qed - -(* potential misnomer: Schutz defines bounds only for infinite chains *) -lemma get_fin_long_ch_bounds: - assumes "long_ch_by_ord f X" - and "finite X" - shows "\x\X. \y\X. \z\X. [f[x..y..z]X]" -proof - - obtain x where "x = f 0" by simp - obtain z where "z = f (card X - 1)" by simp - obtain y where y_def: "y\x \ y\z \ y\X" - by (metis assms(1) long_ch_by_ord_def) - have "x\X" - using ordering_def \x = f 0\ assms(1) long_ch_by_ord_def - by (metis card_gt_0_iff equals0D) - have "z\X" - using ordering_def \z = f (card X - 1)\ assms(1) long_ch_by_ord_def - by (metis card_gt_0_iff equals0D Suc_diff_1 lessI) - obtain n where "n0" - using y_def \f n = y\ \x = f 0\ - using neq0_conv by blast - moreover have "nf n = y\ \n < card X\ \z = f (card X - 1)\ assms(2) - by (metis card.remove card_Diff_singleton less_SucE) - ultimately have "[f[x..y..z]X]" - using long_ch_by_ord_def y_def \x = f 0\ \z = f (card X - 1)\ abc_abc_neq assms ordering_ord_ijk - unfolding fin_long_chain_def - by (metis (no_types, lifting) card_gt_0_iff diff_less equals0D zero_less_one) - thus ?thesis - using points_in_chain - by blast -qed - -lemma get_fin_long_ch_bounds2: - assumes "long_ch_by_ord f X" - and "finite X" - obtains x y z n\<^sub>x n\<^sub>y n\<^sub>z - where "x\X \ y\X \ z\X \ [f[x..y..z]X] \ f n\<^sub>x = x \ f n\<^sub>y = y \ f n\<^sub>z = z" - by (meson assms(1) assms(2) fin_long_chain_def get_fin_long_ch_bounds index_middle_element) - -lemma long_ch_card_ge3: - assumes "ch_by_ord f X" "finite X" - shows "long_ch_by_ord f X \ card X \ 3" -proof - assume "long_ch_by_ord f X" - then obtain a b c where "[f[a..b..c]X]" - using get_fin_long_ch_bounds assms(2) by blast - thus "3 \ card X" - by (metis (no_types, opaque_lifting) One_nat_def card_eq_0_iff diff_Suc_1 empty_iff - fin_long_chain_def index_middle_element leI less_3_cases less_one) -next - assume "3 \ card X" - hence "\short_ch X" - using assms(1) short_ch_card_2 by auto - thus "long_ch_by_ord f X" - using assms(1) ch_by_ord_def by auto -qed - -lemma chain_bounds_unique: - assumes "[f[a..b..c]X]" "[g[x..y..z]X]" - shows "(a=x \ c=z) \ (a=z \ c=x)" -proof - - have "\p\X. (a = p \ p = c) \ [[a p c]]" - using assms(1) ch_all_betw_f by force - then show ?thesis - by (metis (full_types) abc_abc_neq abc_bcd_abd abc_sym assms(1,2) ch_all_betw_f points_in_chain) -qed - -lemma chain_bounds_unique2: - assumes "[f[a..c]X]" "[g[x..z]X]" "card X \ 3" - shows "(a=x \ c=z) \ (a=z \ c=x)" - using chain_bounds_unique - by (metis abc_ac_neq assms(1,2) ch_all_betw_f fin_chain_def points_in_chain short_ch_def) - -subsection "Chain Equivalences" - -subsubsection "Betweenness-chains and strong index-chains" - -lemma equiv_chain_1a: - assumes "[[..a..b..c..]X]" - shows "\f. ch_by_ord f X \ a\X \ b\X \ c\X \ a\b \ a\c \ b\c" -proof - - have in_X: "a\X \ b\X \ c\X" - using assms chain_with_def by auto - have all_neq: "a\c \ a\b \ b\c" - using abc_abc_neq assms chain_with_def by auto - obtain f where "ordering f betw X" - using assms chain_with_def by auto - hence "long_ch_by_ord f X" - using in_X all_neq long_ch_by_ord_def by blast - hence "ch_by_ord f X" - by (simp add: ch_by_ord_def) - thus ?thesis - using all_neq in_X by blast -qed - - -lemma equiv_chain_1b: - assumes "ch_by_ord f X \ a\X \ b\X \ c\X \ a\b \ a\c \ b\c \ [[a b c]]" - shows "[[..a..b..c..]X]" - using assms chain_with_def ch_by_ord_def - by (metis long_ch_by_ord_def short_ch_def) - - -lemma equiv_chain_1: - "[[..a..b..c..]X] \ (\f. ch_by_ord f X \ a\X \ b\X \ c\X \ a\b \ a\c \ b\c \ [[a b c]])" - using equiv_chain_1a equiv_chain_1b long_chain_betw - by meson - - -lemma index_order: - assumes "chain_with x y z X" - and "ch_by_ord f X" and "f a = x" and "f b = y" and "f c = z" - and "finite X \ a b c b (c b (a < b \ b < c \ c < b \ b < a)" - hence "(a\b \ b\c) \ (c\b \ b\a)" - by auto - have all_neq: "x\y \ x\z \ y\z" - using assms(1) equiv_chain_1 by blast - hence is_long: "long_ch_by_ord f X" - by (metis assms(1) assms(2) ch_by_ord_def equiv_chain_1 short_ch_def) - have "a\b \ a\c \ b\c" - using assms(3) assms(4) assms(5) all_neq by blast - hence "(a>b \ b>c) \ (c>b \ b>a)" - using a1 linorder_neqE_nat by blast - hence "(a>b \ c>b) \ (b>c \ b>a)" - using not_less_iff_gr_or_eq by blast - have "a>c \ c>a" - using \a \ b \ a \ c \ b \ c\ by auto - hence "(a>c \ c>b) \ (a>c \ b>a) \ (a>b \ c>a) \ (b>c \ c>a)" - using \(b < a \ c < b) \ (b < c \ a < b)\ by blast - hence o1: "(b c (c a (b a (a c c [[y z x]]" - using assms ordering_ord_ijk long_ch_by_ord_def is_long - by metis - moreover have "(c a [[z x y]]" - using assms ordering_ord_ijk long_ch_by_ord_def is_long - by metis - moreover have "(b a [[y x z]]" - using assms ordering_ord_ijk long_ch_by_ord_def is_long - by metis - moreover have "(a c [[x z y]]" - using assms ordering_ord_ijk long_ch_by_ord_def is_long - by metis - ultimately have "[[y z x]] \ [[z x y]] \ [[y x z]] \ [[x z y]]" - using assms long_ch_by_ord_def is_long o1 - by metis - thus False - by (meson abc_only_cba assms(1) chain_with_def) -qed - - -lemma old_fin_chain_finite: - assumes "finite_chain_with3 x y z X" - shows "finite X" -proof (rule ccontr) - assume "infinite X" - have "x\X" - using assms finite_chain_with3_def chain_with_def by simp - have "y\X" - using assms finite_chain_with3_def chain_with_def by simp - have "z\X" - using assms finite_chain_with3_def chain_with_def by simp - obtain f where "ch_by_ord f X" - using assms equiv_chain_1 finite_chain_with3_def - by auto - obtain a where "f a = x" - using equiv_chain_1 ordering_def \ch_by_ord f X\ assms - by (metis ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def short_ch_def) - obtain c where "f c = z" and "a\c" - using equiv_chain_1 ordering_def \ch_by_ord f X\ \f a = x\ assms - using ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def short_ch_def - by metis - obtain b where "f b = y" and "a\b" and "b\c" - using equiv_chain_1 ordering_def \ch_by_ord f X\ \f a = x\ \f c = z\ assms - using ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def short_ch_def - by metis - obtain n where "ach_by_ord f X\ \f a = x\ \f c = z\ assms equiv_chain_1 \infinite X\ - using ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def short_ch_def - by (metis less_Suc_eq_le not_le not_less_iff_gr_or_eq) - have "[[x y z]]" - using assms chain_with_def finite_chain_with3_def by auto - hence "(a b (c bf a = x\ \f b = y\ \f c = z\ \ch_by_ord f X\ \x\X\ \y\X\ \z\X\ index_order - using \infinite X\ assms finite_chain_with3_def - by blast - hence "(a b c (c b aa\c\ \a\b\ \b\c\ \a \c less_linear - by blast - hence acn_can: "(b c (b a X" - by (metis ordering_def \ch_by_ord f X\ \infinite X\ assms ch_by_ord_def equiv_chain_1 finite_chain_with3_def long_ch_by_ord_def short_ch_def) - hence outside: "[[y z (f n)]] \ [[(f n) x y]]" - using acn_can \ch_by_ord f X\ \f a = x\ \f c = z\ \infinite X\ assms equiv_chain_1 abc_sym - using ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def ordering_ord_ijk short_ch_def - by (metis \f b = y\) - thus False - using \f n \ X\ assms finite_chain_with3_def - by blast -qed - - -lemma index_from_with3: - assumes "finite_chain_with3 a b c X" - shows "\f. (f 0 = a \ f 0 = c) \ ch_by_ord f X" -proof - - obtain f where "ch_by_ord f X" - using assms equiv_chain_1 finite_chain_with3_def - by auto - have no_elt: "\(\w\X. [[w a b]] \ [[b c w]])" - using assms finite_chain_with3_def - by blast - obtain n\<^sub>a n\<^sub>b where "f n\<^sub>a = a" and "n\<^sub>a < card X" - and "f n\<^sub>b = b" and "n\<^sub>b < card X" - using assms old_fin_chain_finite ch_by_ord_def ordering_def - using \ch_by_ord f X\ equiv_chain_1 finite_chain_with3_def long_ch_by_ord_def short_ch_def - by metis - obtain n\<^sub>c where "f n\<^sub>c = c" and "n\<^sub>c < card X" - using assms old_fin_chain_finite ch_by_ord_def ordering_def - using \ch_by_ord f X\ equiv_chain_1 finite_chain_with3_def long_ch_by_ord_def short_ch_def - by metis - have "a\b \ b\c \ a\c" - using assms equiv_chain_1 finite_chain_with3_def by auto - have "a\b \ n\<^sub>a\n\<^sub>b \ b\c \ n\<^sub>a\n\<^sub>c \ a\c \ n\<^sub>b\n\<^sub>c" - using \f n\<^sub>a = a\ \f n\<^sub>b = b\ \f n\<^sub>c = c\ by blast - hence "n\<^sub>a\n\<^sub>b \ n\<^sub>a\n\<^sub>c \ n\<^sub>b\n\<^sub>c" - using \a \ b \ b \ c \ a \ c\ \f n\<^sub>a = a\ \f n\<^sub>b = b\ \f n\<^sub>c = c\ - by auto - have "n\<^sub>a = 0 \ n\<^sub>c = 0" - proof (rule ccontr) - assume "\ (n\<^sub>a = 0 \ n\<^sub>c = 0)" - hence not_0: "n\<^sub>a \ 0 \ n\<^sub>c \ 0" - by linarith - then obtain p where "f 0 = p" - by simp - hence "p\X" - using \ch_by_ord f X\ \n\<^sub>a < card X\ assms card_0_eq ch_by_ord_def zero_into_ordering - using equiv_chain_1 finite_chain_with3_def inf.strict_coboundedI2 inf.strict_order_iff less_one long_ch_by_ord_def old_fin_chain_finite short_ch_def - by metis - have "n\<^sub>ac \ n\<^sub>ca" - using \n\<^sub>a \ n\<^sub>b \ n\<^sub>a \ n\<^sub>c \ n\<^sub>b \ n\<^sub>c\ less_linear by blast - { - assume "n\<^sub>a < n\<^sub>c" - hence "n\<^sub>a < n\<^sub>b" - using index_order \ch_by_ord f X\ \f n\<^sub>a = a\ \f n\<^sub>b = b\ \f n\<^sub>c = c\ \n\<^sub>c < card X\ - using finite_chain_with3_def assms - by fastforce - have "0a \ n\<^sub>ab" - using index_order \n\<^sub>a < n\<^sub>b\ not_0 - by blast - hence "[[p a b]]" - using \ch_by_ord f X\ \f 0=p\ \f n\<^sub>a=a\ \f n\<^sub>b=b\ \n\<^sub>b assms equiv_chain_1 short_ch_def - by (metis ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def ordering_ord_ijk) - hence False - using finite_chain_with3_def \p\X\ - by (metis no_elt) - } - moreover { - assume "n\<^sub>c < n\<^sub>a" - hence "n\<^sub>c < n\<^sub>b" - using index_order \ch_by_ord f X\ \f n\<^sub>a = a\ \f n\<^sub>b = b\ \f n\<^sub>c = c\ \n\<^sub>a < card X\ - using finite_chain_with3_def assms - by fastforce - have "0c \ n\<^sub>cb" - using index_order \n\<^sub>c < n\<^sub>b\ not_0 - by blast - hence "[[p c b]]" - using \ch_by_ord f X\ \f 0=p\ \f n\<^sub>c=c\ \f n\<^sub>b=b\ \n\<^sub>b assms equiv_chain_1 short_ch_def - using ch_by_ord_def finite_chain_with3_def long_ch_by_ord_def ordering_ord_ijk - by metis - hence "[[b c p]]" - by (simp add: abc_sym) - hence False - using finite_chain_with3_def \p\X\ - by (metis no_elt) - } - ultimately show False - using \n\<^sub>a < n\<^sub>c \ n\<^sub>c < n\<^sub>a\ by blast - qed - thus ?thesis - using \ch_by_ord f X\ \f n\<^sub>a = a\ \f n\<^sub>c = c\ - by blast -qed - - -lemma (in MinkowskiSpacetime) with3_and_index_is_fin_chain: - assumes "f 0 = a" and "ch_by_ord f X" and "finite_chain_with3 a b c X" - shows "[f[a..b..c]X]" -proof - - have "finite X" - using ordering_def assms old_fin_chain_finite - by auto - moreover have "long_ch_by_ord f X" - using assms(2) assms(3) ch_by_ord_def equiv_chain_1 finite_chain_with3_def short_ch_def - by metis - moreover have "a\b \ a\c \ b\c \ f 0 = a \ b\X" - using assms(1) assms(3) equiv_chain_1 finite_chain_with3_def - by auto - moreover have "f (card X - 1) = c" - proof - - obtain n where "f n = c" and "n < card X" - using ordering_def equiv_chain_1 finite_chain_with3_def long_ch_by_ord_def - by (metis assms(3) calculation(1,2)) - { - assume "n < card X - 1" - then obtain m where "n (f m)\X" - proof - - have f1: "TernaryOrdering.ordering f betw X" - using \long_ch_by_ord f X\ long_ch_by_ord_def by blast - have f2: "\f A p na. ((p (f na::'a) (f n) (f m) \ \ m < card A) \ \ ordering f p A) - \ \ na < n" - by (metis ordering_def \n < m\) - have "f m \ X" - using f1 by (simp add: ordering_def \m < card X\) - then show ?thesis - using f2 f1 \a\b \ a\c \ b\c \ f 0 = a \ b\X\ \f n = c\ \m < card X\ - using gr_implies_not0 linorder_neqE_nat - by (metis (no_types)) - qed - hence "[[b c (f m)]]" using abc_acd_bcd - by (meson assms(3) chain_with_def finite_chain_with3_def) - hence False - using assms(3) \[[a c (f m)]] \ f m \ X\ - by (metis finite_chain_with3_def) - } - hence "n = card X - 1" - using \n < card X\ by fastforce - thus ?thesis - using \f n = c\ by blast - qed - ultimately show ?thesis - by (simp add: fin_long_chain_def) -qed - - -lemma (in MinkowskiSpacetime) g_from_with3: - assumes "finite_chain_with3 a b c X" - obtains g where "[g[a..b..c]X] \ [g[c..b..a]X]" -proof - - have old_chain_sym: "finite_chain_with3 c b a X" - by (metis abc_sym assms chain_with_def finite_chain_with3_def) - obtain f where f_def: "(f 0 = a \ f 0 = c) \ ch_by_ord f X" - using index_from_with3 assms - by blast - hence "f 0 = a \ [f[a..b..c]X]" - using with3_and_index_is_fin_chain f_def assms - by simp - moreover have "f 0 = c \ [f[c..b..a]X]" - using with3_and_index_is_fin_chain f_def assms old_chain_sym - by simp - ultimately show ?thesis - using f_def that - by auto -qed - - -lemma (in MinkowskiSpacetime) equiv_chain_2a: - assumes "finite_chain_with3 a b c X" - obtains f where "[f[a..b..c]X]" -proof - - obtain g where "[g[a..b..c]X] \ [g[c..b..a]X]" - using assms g_from_with3 by blast - thus ?thesis - proof (* in two cases *) - assume "[g[a..b..c]X]" - show ?thesis - using \[g[a .. b .. c]X]\ that - by blast - next - assume "[g[c..b..a]X]" - show ?thesis - using \[g[c .. b .. a]X]\ chain_sym that - by blast - qed -qed - - -lemma equiv_chain_2b: - assumes "[f[a..b..c]X]" - shows "finite_chain_with3 a b c X" -proof - - have aligned: "[[a b c]]" - using assms fin_ch_betw - by auto - hence some_chain: "[[..a..b..c..]X]" - using assms ch_by_ord_def equiv_chain_1b fin_long_chain_def points_in_chain - by metis - have "\(\w\X. [[w a b]] \ [[b c w]])" - proof (safe) - fix w assume "w\X" - { - assume case1: "[[w a b]]" - then obtain n where "f n = w" and "nw\X\ abc_bcd_abd abc_only_cba aligned assms fin_ch_betw fin_long_chain_def - by (metis (no_types, opaque_lifting)) - have "f 0 = a" - using assms fin_long_chain_def - by blast - hence "n < 0" - proof - - have f1: "f (card X - 1) = c" - by (meson MinkowskiBetweenness.fin_long_chain_def MinkowskiBetweenness_axioms assms) - have "\ [[a w c]]" - by (meson abc_bcd_abd abc_only_cba assms case1 fin_ch_betw) - thus ?thesis - using f1 fin_long_chain_def \w \ X\ abc_only_cba assms case1 fin_ch_betw - by (metis (no_types)) - qed - thus False - by simp - } - moreover { - assume case2: "[[b c w]]" - then obtain n where "f n = w" and "nw\X\ ordering_def abc_bcd_abd abc_only_cba aligned assms fin_ch_betw - using fin_long_chain_def long_ch_by_ord_def - by metis - have "f (card X - 1) = c" - using assms fin_long_chain_def - by blast - have "\ [[a w c]]" - using abc_bcd_abd abc_only_cba assms case2 fin_ch_betw abc_bcd_acd - by meson - hence "n > card X - 1" - using \\ [[a w c]]\ \w \ X\ abc_only_cba assms case2 fin_ch_betw - unfolding fin_long_chain_def - by (metis (no_types)) - thus False - using \n < card X\ - by linarith - } - qed - thus ?thesis - by (simp add: finite_chain_with3_def some_chain) -qed - - -lemma (in MinkowskiSpacetime) equiv_chain_2: - "\f. [f[a..b..c]X] \ [[a..b..c]X]" - using equiv_chain_2a equiv_chain_2b - by meson - -end (* context MinkowskiChain *) - -section "Results for segments, rays and chains" - -context MinkowskiChain begin - -lemma inside_not_bound: - assumes "[f[a..b..c]X]" - and "j0 \ f j \ a" "j f j \ c" -proof - - have bound_indices: "f 0 = a \ f (card X - 1) = c" - using assms(1) fin_long_chain_def by auto - show "f j \ a" if "j>0" - proof (cases) - assume "f j = c" - then have "[[(f 0) (f j) b]] \ [[(f 0) b (f j)]]" - using assms(1) fin_ch_betw fin_long_chain_def - by metis - thus ?thesis using abc_abc_neq bound_indices by blast - next - assume "f j \ c" - then have "[[(f 0) (f j) c]] \ [[(f 0) c (f j)]]" - using assms fin_ch_betw - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis abc_abc_neq assms that ch_all_betw_f nat_neq_iff) - thus ?thesis - using abc_abc_neq bound_indices by blast - qed - show "f j \ c" if "jf j = a\ assms(1) fin_long_chain_def - by blast - next - assume "f j \ a" - have "0 < card X" - using assms(2) by linarith - hence "[[a (f j) (f (card X - 1))]] \ [[(f j) a (f (card X - 1))]]" - using assms fin_ch_betw fin_long_chain_def order_finite_chain - by (metis \f j \ a\ diff_less le_numeral_extra(1-3) neq0_conv that) - thus "f j \ c" - using abc_abc_neq bound_indices by auto - qed -qed - - -lemma some_betw2: - assumes "[f[a..b..c]X]" - and "j0" "f j \ b" - shows "[[a b (f j)]] \ [[a (f j) b]]" -proof - - obtain ab where ab_def: "path ab a b" "X\ab" - by (metis fin_long_chain_def long_chain_on_path assms(1) points_in_chain subsetD) - have bound_indices: "f 0 = a \ f (card X - 1) = c" - using assms(1) fin_long_chain_def by auto - have "f j \ a" - using inside_not_bound(1) assms(1) assms(2) assms(3) - by blast - have "\[[(f j) a b]]" - using abc_bcd_abd abc_only_cba assms(1,2) fin_ch_betw fin_long_chain_def - by (metis ordering_def ch_all_betw_f long_ch_by_ord_def) - thus " [[a b (f j)]] \ [[a (f j) b]]" - using some_betw [where Q=ab and a=a and b=b and c="f j"] - using ab_def assms(4) \f j \ a\ - by (metis ordering_def abc_sym assms(1,2) fin_long_chain_def long_ch_by_ord_def subsetD) -qed - -lemma i_le_j_events_neq1: - assumes "[f[a..b..c]X]" - and "i b" (* this just means you need to pick b well *) - shows "f i \ f j" -proof - - have in_X: "f i \ X \ f j \ X" - by (metis ordering_def assms(1,2,3) fin_long_chain_def less_trans long_ch_by_ord_def) - have bound_indices: "f 0 = a \ f (card X - 1) = c" - using assms(1) fin_long_chain_def by auto - obtain ab where ab_def: "path ab a b" "X\ab" - by (metis fin_long_chain_def long_chain_on_path assms(1) points_in_chain subsetD) - show ?thesis - proof (cases) - assume "f i = a" - hence "[[a (f j) b]] \ [[a b (f j)]]" - using some_betw2 assms by blast - thus ?thesis - using \f i = a\ abc_abc_neq by blast - next assume "f i \ a" - hence "[[a (f i) (f j)]]" - using assms(1,2,3) ch_equiv fin_long_chain_def order_finite_chain2 - by (metis gr_implies_not_zero le_numeral_extra(3) less_linear) - thus ?thesis - using abc_abc_neq by blast - qed -qed - -lemma i_le_j_events_neq: - assumes "[f[a..b..c]X]" - and "i f j" -proof - - have in_X: "f i \ X \ f j \ X" - by (metis ordering_def assms(1,2,3) fin_long_chain_def less_trans long_ch_by_ord_def) - have bound_indices: "f 0 = a \ f (card X - 1) = c" - using assms(1) fin_long_chain_def by auto - obtain ab where ab_def: "path ab a b" "X\ab" - by (metis fin_long_chain_def long_chain_on_path assms(1) points_in_chain subsetD) - show ?thesis - proof (cases) - assume "f i = a" - show ?thesis - proof (cases) - assume "(f j) = b" - thus ?thesis - by (simp add: \(f i) = a\ ab_def(1)) - next assume "(f j) \ b" - have "[[a (f j) b]] \ [[a b (f j)]]" - using some_betw2 assms \(f j) \ b\ by blast - thus ?thesis - using \(f i) = a\ abc_abc_neq by blast - qed - next assume "(f i) \ a" - hence "[[a (f i) (f j)]]" - using assms(1,2,3) ch_equiv fin_long_chain_def order_finite_chain2 - by (metis gr_implies_not_zero le_numeral_extra(3) less_linear) - thus ?thesis - using abc_abc_neq by blast - qed -qed - -lemma indices_neq_imp_events_neq: - assumes "[f[a..b..c]X]" - and "i\j" "j f j" - by (metis assms i_le_j_events_neq less_linear) - - -lemma index_order2: - assumes "[f[x..y..z]X]" and "f a = x" and "f b = y" and "f c = z" - and "finite X \ a < card X" and "finite X \ b < card X" and "finite X \ c < card X" - shows "(a b (c b a < card X" and "finite X \ b < card X" and "finite X \ c < card X" - shows "(a b (c b\

" "[f[(f 0)..]X]" "X\Q" "is_bound_f b X f" - shows "b\Q" -proof - - obtain a c where "a\X" "c\X" "[[a c b]]" - using assms(4) - by (metis ordering_def inf_chain_is_long is_bound_f_def 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 "x\segment a b \ x\segment b c" for x - proof - fix x assume "x\segment 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 "x\segment 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}. c\a \ c\b)" - by blast - have "a\b \ a\Q & b\Q" - using X_is card_X events_X by force - hence "short_ch {a,b}" - using path_Q short_ch_def no_c by blast - thus ?thesis - by (simp add: X_is ch_by_ord_def ch_def) -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: "a\b \ a\c \ b\c" - 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: "a\Q \ b\Q \ c\Q" - 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: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and bY: "[[b a\<^sub>1 a\<^sub>n]]" - fixes g defines g_def: "g \ (\j::nat. if j\1 then f (j-1) else b)" - shows "[g[b .. a\<^sub>1 .. a\<^sub>n](insert b Y)]" -proof - - let ?X = "insert b Y" - have "b\Y" - by (metis abc_ac_neq abc_only_cba(1) bY ch_all_betw_f long_ch_Y) - have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" - using long_ch_Y by (simp add: fin_long_chain_def) - have fin_Y: "card Y \ 3" - using fin_long_chain_def long_ch_Y numeral_2_eq_2 - by (metis ch_by_ord_def long_ch_card_ge3) - hence num_ord: "0 \ (0::nat) \ 0<(1::nat) \ 1 < card Y - 1 \ card Y - 1 < card Y" - by linarith - hence "[[a\<^sub>1 (f 1) a\<^sub>n]]" - using order_finite_chain fin_long_chain_def long_ch_Y - by auto - text \Schutz has a step here that says \[b a\<^sub>1 a\<^sub>2 a\<^sub>n]\ is a chain (using Theorem 9). - We have no easy way of denoting an ordered 4-element chain, so we skip this step - using an ordering lemma from our script for 3.6, which Schutz doesn't list.\ - hence "[[b a\<^sub>1 (f 1)]]" - using bY abd_bcd_abc by blast - have "ordering2 g betw ?X" - proof - - { - fix n assume "finite ?X \ n ?X" - apply (cases "n\1") - prefer 2 apply (simp add: g_def) - proof - assume "1\n" "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 \b\Y\ card.insert diff_Suc_1 fin_long_chain_def long_ch_Y points_in_chain - by (metis \g n = f (n - 1)\) - next - case False - hence "n < card Y" - using points_in_chain \finite ?X \ n < card ?X\ \g n = f (n - 1)\ \g n \ Y\ \b\Y\ - by (metis card.insert fin_long_chain_def 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 unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (meson less_trans num_ord) - 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 n' n'' assume "(finite ?X \ n'' < card ?X)" "Suc n = n' \ Suc n' = n''" - hence "[[(g n) (g n') (g n'')]]" - using \b\Y\ \[[b a\<^sub>1 (f 1)]]\ g_def long_ch_Y ordering_ord_ijk - by (smt (verit, ccfv_threshold) fin_long_chain_def long_ch_by_ord_def - One_nat_def card.insert diff_Suc_Suc diff_diff_cancel diff_is_0_eq - finite_insert nat_less_le not_less not_less_eq_eq) - } moreover { - fix x assume "x\?X" "x=b" - have "(finite ?X \ 0 < card ?X) \ g 0 = x" - by (simp add: \b\Y\ \x = b\ g_def) - } moreover { - fix x assume "x\?X" "x\b" - hence "\n. (finite ?X \ n < card ?X) \ g n = x" - proof - - obtain n where "f n = x" "n < card Y" - using \x\?X\ \x\b\ - by (metis ordering_def fin_long_chain_def insert_iff long_ch_Y long_ch_by_ord_def) - have "(finite ?X \ n+1 < card ?X)" "g(n+1) = x" - apply (simp add: \b\Y\ \n < card Y\) - by (simp add: \f n = x\ g_def) - thus ?thesis by auto - qed - } - ultimately show ?thesis - unfolding ordering2_def - by smt (* sledgehammer proposes both meson and auto, neither of which work... *) - qed - hence "long_ch_by_ord2 g ?X" - unfolding long_ch_by_ord2_def - using points_in_chain fin_long_chain_def \b\Y\ - by (metis abc_abc_neq bY insert_iff long_ch_Y points_in_chain) - hence "long_ch_by_ord g ?X" - using ch_equiv fin_Y - by (meson fin_long_chain_def finite_insert long_ch_Y) - thus ?thesis - unfolding fin_long_chain_def - using bound_indices \b\Y\ g_def num_ord points_in_chain long_ch_Y fin_long_chain_def - by (metis card.insert diff_Suc_1 finite_insert insert_iff less_trans nat_less_le) -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: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and Yb: "[[a\<^sub>1 a\<^sub>n b]]" - fixes g defines g_def: "g \ (\j::nat. if j \ (card Y - 1) then f j else b)" - shows "[g[a\<^sub>1 .. a\<^sub>n .. b](insert b Y)]" -proof - - let ?X = "insert b Y" - have "b\Y" - by (metis Yb abc_abc_neq abc_only_cba(2) ch_all_betw_f long_ch_Y) - have fin_X: "finite ?X" - using fin_long_chain_def long_ch_Y by blast - have fin_Y: "card Y \ 3" - by (meson ch_by_ord_def fin_long_chain_def long_ch_Y long_ch_card_ge3) - have "a\<^sub>1\Y \ a\<^sub>n\Y \ a\Y" - using long_ch_Y points_in_chain by blast - have "a\<^sub>1\a \ a\ a\<^sub>n \ a\<^sub>1\a\<^sub>n" - using fin_long_chain_def long_ch_Y by auto - have "Suc (card Y) = card ?X" - using \b\Y\ fin_X fin_long_chain_def long_ch_Y by auto - obtain f2 where f2_def: "[f2[a\<^sub>n..a..a\<^sub>1]Y]" "f2=(\n. f (card Y - 1 - n))" - using chain_sym long_ch_Y by blast - obtain g2 where g2_def: "g2 = (\j::nat. if j\1 then f2 (j-1) else b)" - by simp - have "[[b a\<^sub>n a\<^sub>1]]" - using abc_sym Yb by blast - hence g2_ord_X: "[g2[b .. a\<^sub>n .. a\<^sub>1]?X]" - using chain_append_at_left_edge [where a\<^sub>1="a\<^sub>n" and a\<^sub>n="a\<^sub>1" and f="f2"] - fin_X \b\Y\ f2_def g2_def - by blast - then obtain g1 where g1_def: "[g1[a\<^sub>1..a\<^sub>n..b]?X]" "g1=(\n. g2 (card ?X - 1 - n))" - using chain_sym by blast - have sYX: "(card Y) = (card ?X) - 1" - using assms(2,3) fin_long_chain_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\ fin_long_chain_def long_ch_Y sYX - by (metis Suc_1 Suc_diff_1 Suc_diff_le card_gt_0_iff diff_Suc_eq_diff_pred diff_commute - diff_diff_cancel equals0D less_one nat.simps(3) not_less) - 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: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and S_def: "S = {k::nat. [[a\<^sub>1 (f k) b]] \ k < card Y}" - and k_def: "S\{}" "k = Max S" - and k'_def: "k'>0" "k' S" -proof - - have "k\S" using k_def Max_in S_def - by (metis finite_Collect_conjI finite_Collect_less_nat) - show "k' \ S" - proof (rule ccontr) - assume "\k'\S" - hence "[[a\<^sub>1 b (f k')]]" - using order_finite_chain S_def abc_acd_bcd abc_bcd_acd abc_sym long_ch_Y - by (smt fin_long_chain_def \0 < k'\ \k \ S\ \k' < k\ le_numeral_extra(3) - less_trans mem_Collect_eq) - have "[[a\<^sub>1 (f k) b]]" - using S_def \k \ S\ by blast - have "[[(f k) b (f k')]]" - using abc_acd_bcd \[[a\<^sub>1 b (f k')]]\ \[[a\<^sub>1 (f k) b]]\ by blast - have "k' < card Y" - using S_def \k \ S\ \k' < k\ less_trans by blast - thus False - using abc_bcd_abd order_finite_chain S_def abc_only_cba(2) long_ch_Y - \0 < k'\ \[[(f k) b (f k')]]\ \k \ S\ \k' < k\ - unfolding fin_long_chain_def - by (metis (mono_tags, lifting) le_numeral_extra(3) mem_Collect_eq) - qed -qed - - -lemma (*for 10*) smallest_k_ex: - assumes long_ch_Y: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and Y_def: "b\Y" - and Yb: "[[a\<^sub>1 b a\<^sub>n]]" - shows "\k>0. [[a\<^sub>1 b (f k)]] \ k < card Y \ \(\k'1 b (f k')]])" -proof - -(* the usual suspects first, they'll come in useful I'm sure *) - have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" - using fin_long_chain_def long_ch_Y by auto - have fin_Y: "finite Y" - using fin_long_chain_def long_ch_Y by blast - have card_Y: "card Y \ 3" - using fin_long_chain_def long_ch_Y points_in_chain - by (metis (no_types, lifting) One_nat_def antisym card2_either_elt1_or_elt2 diff_is_0_eq' - not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) - - text \We consider all indices of chain elements between \a\<^sub>1\ and \b\, and find the maximal one.\ - let ?S = "{k::nat. [[a\<^sub>1 (f k) b]] \ k < card Y}" - obtain S where S_def: "S=?S" - by simp (* just to check this Isabelle-exists *) - 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 \ [[a\<^sub>1 b (f 1)]] \ 1 < card Y \ \ (\k'::nat. k' < 1 \ [[a\<^sub>1 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 - (* using card_Y by linarith *) - show "\ (\k'::nat. k' < 1 \ [[a\<^sub>1 b (f k')]])" - using abc_abc_neq bound_indices - by blast - show "[[a\<^sub>1 b (f 1)]]" - proof - - have "f 1 \ Y" - by (metis ordering_def diff_0_eq_0 fin_long_chain_def less_one long_ch_Y long_ch_by_ord_def nat_neq_iff) - (* have "[[a\<^sub>1 b f 1]] \ [[a\<^sub>1 f 1 b]]" *) - hence "[[a\<^sub>1 (f 1) a\<^sub>n]]" - using bound_indices long_ch_Y - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (smt One_nat_def card.remove card_Diff1_less card_Diff_singleton diff_is_0_eq' - le_eq_less_or_eq less_SucE neq0_conv zero_less_diff zero_less_one) - hence "[[a\<^sub>1 b (f 1)]] \ [[a\<^sub>1 (f 1) b]] \ [[b a\<^sub>1 (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 "[[a\<^sub>1 b (f 1)]]" - (* by (smt S_def Yb \S = {}\ \[[a\<^sub>1 f 1 a\<^sub>n]]\ abc_bcd_abd abc_sym abd_bcd_abc bound_indices - diff_is_0_eq' empty_iff mem_Collect_eq nat_le_linear nat_less_le) *) - proof - - have "\n. \ ([[a\<^sub>1 (f n) b]] \ n < card Y)" - using S_def \S = {}\ - by blast - then have "[[a\<^sub>1 b (f 1)]] \ \ [[a\<^sub>n (f 1) b]] \ \ [[a\<^sub>1 (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 \[[a\<^sub>1 b (f 1)]] \ [[a\<^sub>1 (f 1) b]] \ [[b a\<^sub>1 (f 1)]]\ \[[a\<^sub>1 (f 1) a\<^sub>n]]\) - 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 "k\1" - proof (rule ccontr) - assume "\ 1 \ k" - hence "k=0" by simp - have "[[a\<^sub>1 (f k) b]]" - using \k\S\ 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 [[a\<^sub>1 b (f ?k)]] \ ?k < card Y \ \ (\k'::nat. k' < ?k \ [[a\<^sub>1 b (f k')]])" - proof (intro conjI) - show "(0::nat)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 "[[a\<^sub>1 b (f ?k)]]" - proof - - have "f ?k \ Y" - using \k + 1 < card Y\ - by (metis ordering_def fin_long_chain_def long_ch_Y long_ch_by_ord_def) - have "[[a\<^sub>1 (f ?k) a\<^sub>n]] \ f ?k = a\<^sub>n" - using bound_indices long_ch_Y \k + 1 < card Y\ - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis (no_types, lifting) Suc_lessI add.commute add_gr_0 card_Diff1_less - card_Diff_singleton less_diff_conv plus_1_eq_Suc zero_less_one) - thus "[[a\<^sub>1 b (f ?k)]]" - proof (rule disjE) - assume "[[a\<^sub>1 (f ?k) a\<^sub>n]]" - hence "f ?k \ a\<^sub>n" - by (simp add: abc_abc_neq) - hence "[[a\<^sub>1 b (f ?k)]] \ [[a\<^sub>1 (f ?k) b]] \ [[b a\<^sub>1 (f ?k)]]" - using abc_ex_path_unique some_betw abc_sym \[[a\<^sub>1 (f ?k) a\<^sub>n]]\ - \f ?k \ Y\ Yb abc_abc_neq assms(3) cross_once_notin - by (smt Y_def) - moreover have "\ [[a\<^sub>1 (f ?k) b]]" - proof - assume "[[a\<^sub>1 (f ?k) b]]" - hence "?k \ S" - using S_def \[[a\<^sub>1 (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 a\<^sub>1 (f ?k)]]" - using Yb \[[a\<^sub>1 (f ?k) a\<^sub>n]]\ abc_only_cba - by blast - ultimately show "[[a\<^sub>1 b (f ?k)]]" - by blast - next assume "f ?k = a\<^sub>n" - show ?thesis - using Yb \f (k + 1) = a\<^sub>n\ by blast - qed - qed - show "\(\k'::nat. k' < k + 1 \ [[a\<^sub>1 b (f k')]])" - proof - assume "\k'::nat. k' < k + 1 \ [[a\<^sub>1 b (f k')]]" - then obtain k' where k'_def: "k'>0" "k' < k + 1" "[[a\<^sub>1 b (f k')]]" - using abc_ac_neq bound_indices neq0_conv - by blast - hence "k'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: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and Y_def: "b\Y" - and Yb: "[[a\<^sub>1 b a\<^sub>n]]" - shows "\k. [[(f k) b a\<^sub>n]] \ k < card Y - 1 \ \(\k'k \ [[(f k') b a\<^sub>n]])" -proof - -(* the usual suspects first, they'll come in useful I'm sure *) - have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" - using fin_long_chain_def long_ch_Y by auto - have fin_Y: "finite Y" - using fin_long_chain_def long_ch_Y by blast - have card_Y: "card Y \ 3" - using fin_long_chain_def long_ch_Y points_in_chain - by (metis (no_types, lifting) One_nat_def antisym card2_either_elt1_or_elt2 diff_is_0_eq' - not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) - - text \Again we consider all indices of chain elements between \a\<^sub>1\ and \b\.\ - let ?S = "{k::nat. [[a\<^sub>n (f k) b]] \ k < card Y}" - obtain S where S_def: "S=?S" - by simp (* just to check this Isabelle-exists *) - 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 a\<^sub>n]] \ ?n < card Y - 1 \ \(\k'?n \ [[(f k') b a\<^sub>n]])" - proof (intro conjI) - show "?n < card Y - 1" - using Yb abc_ac_neq bound_indices not_le by fastforce - next show "\(\k'?n \ [[(f k') b a\<^sub>n]])" - 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 a\<^sub>n]]" - proof - - have "f ?n \ Y" - by (metis ordering_def diff_less fin_long_chain_def gr_implies_not0 long_ch_Y - long_ch_by_ord_def neq0_conv not_less_eq numeral_2_eq_2) - hence "[[a\<^sub>1 (f ?n) a\<^sub>n]]" - using bound_indices long_ch_Y - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - using card_Y by force - hence "[[a\<^sub>n b (f ?n)]] \ [[a\<^sub>n (f ?n) b]] \ [[b a\<^sub>n (f ?n)]]" - using abc_ex_path_unique some_betw abc_sym - by (smt Y_def Yb \f ?n \ Y\ abc_abc_neq cross_once_notin) - thus "[[(f ?n) b a\<^sub>n]]" - proof - - have "\n. \ ([[a\<^sub>n (f n) b]] \ n < card Y)" - using S_def \S = {}\ - by blast - then have "[[a\<^sub>n b (f ?n)]] \ \ [[a\<^sub>1 (f ?n) b]] \ \ [[a\<^sub>n (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 \[[a\<^sub>n b (f ?n)]] \ [[a\<^sub>n (f ?n) b]] \ [[b a\<^sub>n (f ?n)]]\ \[[a\<^sub>1 (f ?n) a\<^sub>n]]\) - qed - qed - qed - qed - next assume "\S={}" - obtain k where "k = Min S" - by simp - hence "k \ S" using Max_in - by (simp add: \S \ {}\ \finite S\) - - show ?thesis - proof - let ?k = "k-1" - show "[[(f ?k) b a\<^sub>n]] \ ?k < card Y - 1 \ \ (\k' [[(f k') b a\<^sub>n]])" - 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 a\<^sub>n]]" - proof - - have "f ?k \ Y" - using \k - 1 < card Y - 1\ long_ch_Y long_ch_by_ord_def ordering_def - by (metis diff_less fin_long_chain_def less_trans neq0_conv zero_less_one) - have "[[a\<^sub>1 (f ?k) a\<^sub>n]] \ f ?k = a\<^sub>1" - using bound_indices long_ch_Y \k - 1 < card Y - 1\ - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (smt S_def \k \ S\ add_diff_inverse_nat card_Diff1_less card_Diff_singleton - less_numeral_extra(4) less_trans mem_Collect_eq nat_add_left_cancel_less - neq0_conv zero_less_diff) - thus "[[(f ?k) b a\<^sub>n]]" - proof (rule disjE) - assume "[[a\<^sub>1 (f ?k) a\<^sub>n]]" - hence "f ?k \ a\<^sub>1" - using abc_abc_neq by blast - hence "[[a\<^sub>n b (f ?k)]] \ [[a\<^sub>n (f ?k) b]] \ [[b a\<^sub>n (f ?k)]]" - using abc_ex_path_unique some_betw abc_sym \[[a\<^sub>1 (f ?k) a\<^sub>n]]\ - \f ?k \ Y\ Yb abc_abc_neq assms(3) cross_once_notin - by (smt Y_def) - moreover have "\ [[a\<^sub>n (f ?k) b]]" - proof - assume "[[a\<^sub>n (f ?k) b]]" - hence "?k \ S" - using S_def \[[a\<^sub>n (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) \ a\<^sub>1\ fin_long_chain_def long_ch_Y - by auto - qed - moreover have "\ [[b a\<^sub>n (f ?k)]]" - using Yb \[[a\<^sub>1 (f ?k) a\<^sub>n]]\ abc_only_cba(2) abc_bcd_acd - by blast - ultimately show "[[(f ?k) b a\<^sub>n]]" - using abc_sym by auto - next assume "f ?k = a\<^sub>1" - show ?thesis - using Yb \f (k - 1) = a\<^sub>1\ by blast - qed - qed - show "\(\k' [[(f k') b a\<^sub>n]])" - proof - assume "\k' [[(f k') b a\<^sub>n]]" - then obtain k' where k'_def: "k' k - 1" "[[a\<^sub>n 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) - hence "k'\S" - using S_is_dense long_ch_Y S_def \\S={}\ \k = Min S\ \k' - by (smt Yb \k \ S\ abc_acd_bcd abc_only_cba(3) card_Diff1_less card_Diff_singleton - fin_long_chain_def k'_def(3) less_le mem_Collect_eq neq0_conv order_finite_chain) - 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: "[f[a\<^sub>0..a..a\<^sub>n]Y]" - and x_def: "x\Y" "[[a\<^sub>0 x a\<^sub>n]]" - obtains n\<^sub>b n\<^sub>c b c - where "b=f n\<^sub>b" "c=f n\<^sub>c" "[[b x c]]" "b\Y" "c\Y" "n\<^sub>b = n\<^sub>c - 1" "n\<^sub>cc>0" - "\(\k < card Y. [[(f k) x a\<^sub>n]] \ k>n\<^sub>b)" "\(\kc. [[a\<^sub>0 x (f k)]])" -proof - - have "\ n\<^sub>b n\<^sub>c b c. b=f n\<^sub>b \ c=f n\<^sub>c \ [[b x c]] \ b\Y \ c\Y \ n\<^sub>b = n\<^sub>c - 1 \ n\<^sub>c n\<^sub>c>0 - \ \(\k < card Y. [[(f k) x a\<^sub>n]] \ k>n\<^sub>b) \ \(\k < n\<^sub>c. [[a\<^sub>0 x (f k)]])" - proof - - have bound_indices: "f 0 = a\<^sub>0 \ f (card Y - 1) = a\<^sub>n" - using fin_long_chain_def long_ch_Y by auto - have "finite Y" - using fin_long_chain_def long_ch_Y by blast - obtain P where P_def: "P\\

" "Y\P" - using chain_on_path long_ch_Y - unfolding fin_long_chain_def ch_by_ord_def - by blast - hence "x\P" - using betw_b_in_path x_def(2) long_ch_Y points_in_chain - by (metis abc_abc_neq in_mono) - obtain n\<^sub>c where nc_def: "\(\k. [[a\<^sub>0 x (f k)]] \ kc)" "[[a\<^sub>0 x (f n\<^sub>c)]]" "n\<^sub>cc>0" - using smallest_k_ex [where a\<^sub>1=a\<^sub>0 and a=a and a\<^sub>n=a\<^sub>n and b=x and f=f and Y=Y] - long_ch_Y x_def - by blast - then obtain c where c_def: "c=f n\<^sub>c" "c\Y" - using long_ch_Y long_ch_by_ord_def fin_long_chain_def - by (metis ordering_def) - have c_goal: "c=f n\<^sub>c \ c\Y \ n\<^sub>c n\<^sub>c>0 \ \(\k < card Y. [[a\<^sub>0 x (f k)]] \ kc)" - using c_def nc_def(1,3,4) by blast - obtain n\<^sub>b where nb_def: "\(\k < card Y. [[(f k) x a\<^sub>n]] \ k>n\<^sub>b)" "[[(f n\<^sub>b) x a\<^sub>n]]" "n\<^sub>b1=a\<^sub>0 and a=a and a\<^sub>n=a\<^sub>n and b=x and f=f and Y=Y] - long_ch_Y x_def - by blast - hence "n\<^sub>bb" "b\Y" - using nb_def long_ch_Y long_ch_by_ord_def fin_long_chain_def ordering_def - by metis - have "[[b x c]]" - proof - - have "[[b x a\<^sub>n]]" - using b_def(1) nb_def(2) by blast - have "[[a\<^sub>0 x c]]" - using c_def(1) nc_def(2) by blast - moreover have "\a. [[a x b]] \ \ [[a a\<^sub>n x]]" - using \[[b x a\<^sub>n]]\ abc_bcd_acd - by (metis (full_types) abc_sym) - moreover have "\a. [[a x b]] \ \ [[a\<^sub>n a x]]" - using \[[b x a\<^sub>n]]\ by (meson abc_acd_bcd abc_sym) - moreover have "a\<^sub>n = c \ [[b x c]]" - using \[[b x a\<^sub>n]]\ by meson - ultimately show ?thesis - using abc_abd_bcdbdc abc_sym x_def(2) - by meson - qed - have "n\<^sub>bc" - using \[[b x c]]\ \n\<^sub>c \n\<^sub>b \c = f n\<^sub>c\ \b = f n\<^sub>b\ - by (smt (* TODO *) - \\thesis. (\n\<^sub>b. \\ (\kn]] \ n\<^sub>b < k); [[(f n\<^sub>b) x a\<^sub>n]]; n\<^sub>b < card Y - 1\ - \ thesis) \ thesis\ abc_abd_acdadc abc_ac_neq abc_only_cba diff_less - fin_long_chain_def le_antisym le_trans less_imp_le_nat less_numeral_extra(1) - linorder_neqE_nat long_ch_Y nb_def(2) nc_def(4) order_finite_chain) - have "n\<^sub>b = n\<^sub>c - 1" - proof (rule ccontr) - assume "n\<^sub>b \ n\<^sub>c - 1" - have "n\<^sub>bc-1" - using \n\<^sub>b \ n\<^sub>c - 1\ \n\<^sub>bc\ by linarith - hence "[[(f n\<^sub>b) (f(n\<^sub>c-1)) (f n\<^sub>c)]]" - using \n\<^sub>b \ n\<^sub>c - 1\ fin_long_chain_def long_ch_Y nc_def(3) order_finite_chain - by auto - have "\[[a\<^sub>0 x (f(n\<^sub>c-1))]]" - using nc_def(1,4) diff_less less_numeral_extra(1) - by blast - have "n\<^sub>c-1\0" - using \n\<^sub>b < n\<^sub>c\ \n\<^sub>b \ n\<^sub>c - 1\ by linarith - hence "f(n\<^sub>c-1)\a\<^sub>0 \ a\<^sub>0\x" - using bound_indices - by (metis \[[(f n\<^sub>b) (f (n\<^sub>c - 1)) (f n\<^sub>c)]]\ abc_abc_neq abd_bcd_abc b_def(1,2) ch_all_betw_f - long_ch_Y nb_def(2) nc_def(2)) - have "x\f(n\<^sub>c-1)" - using x_def(1) nc_def(3) long_ch_Y - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis less_imp_diff_less) - hence "[[a\<^sub>0 (f (n\<^sub>c-1)) x]]" - using some_betw P_def(1,2) abc_abc_neq abc_acd_bcd abc_bcd_acd abc_sym b_def(1,2) - c_def(1,2) ch_all_betw_f in_mono long_ch_Y nc_def(2) betw_b_in_path - by (smt \[[(f n\<^sub>b) (f (n\<^sub>c-1)) (f n\<^sub>c)]]\ \\ [[a\<^sub>0 x (f (n\<^sub>c-1))]]\ \x \ P\ \f(n\<^sub>c-1)\a\<^sub>0 \ a\<^sub>0\x\) - hence "[[(f(n\<^sub>c-1)) x a\<^sub>n]]" - using abc_acd_bcd x_def(2) by blast - thus False using nb_def(1) - using \n\<^sub>b < n\<^sub>c - 1\ less_imp_diff_less nc_def(3) - by blast - qed - have b_goal: "b=f n\<^sub>b \ b\Y \ n\<^sub>b=n\<^sub>c-1 \ \(\k < card Y. [[(f k) x a\<^sub>n]] \ k>n\<^sub>b)" - using b_def nb_def(1) nb_def(3) \n\<^sub>b=n\<^sub>c-1\ by blast - thus ?thesis - using \[[b x c]]\ c_goal - using \n\<^sub>b < 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: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - and Y_def: "b\Y" - and Yb: "[[a\<^sub>1 b a\<^sub>n]]" - and k_def: "[[a\<^sub>1 b (f k)]]" "k < card Y" "\(\k'. (0::nat) k' [[a\<^sub>1 b (f k')]])" - fixes g - defines g_def: "g \ (\j::nat. if (j\k-1) then f j else (if (j=k) then b else f (j-1)))" - shows "[g[a\<^sub>1 .. b .. a\<^sub>n]insert b Y]" -proof - - let ?X = "insert b Y" - have fin_X: "finite ?X" - by (meson fin_long_chain_def finite.insertI long_ch_Y) - have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" - using fin_long_chain_def long_ch_Y - by auto - have fin_Y: "finite Y" - using fin_long_chain_def long_ch_Y by blast - have f_def: "long_ch_by_ord f Y" - using fin_long_chain_def long_ch_Y by blast - have \a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n\ - 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 \[[a\<^sub>1 b (f k)]]\ \k = 1\ bound_indices by auto - next assume "k\1" show "[[(f (k-1)) b (f k)]]" - proof - - have "[[a\<^sub>1 (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 fin_long_chain_def - 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 {a\<^sub>1, (f (k-1)), b, (f k)}" using chain4 - using k_def(1) abc_ex_path_unique between_chain cross_once_notin - by (smt \[[a\<^sub>1 (f (k - 1)) (f k)]]\ abc_abc_neq insert_absorb2) -(*TODO: the proof below needs extra assumptions, but we should still try to get rid of smt above. *) - (* proof - - have "a\<^sub>1\Q \ a\<^sub>n\Q \ b\Q" - using Y_def long_ch_Y events_X points_in_chain by auto - have f1: "a\<^sub>1 \ b \ a\<^sub>1 \ f k \ b \ f k" - using abc_abc_neq k_def(1) by presburger - then have f2: "f k \ Q" - using betw_c_in_path k_def(1) path_Q - \a\<^sub>1 \ Q \ a\<^sub>n \ Q \ b \ Q\ - by blast - then have "f (k - 1) \ Q" - using betw_b_in_path path_Q - f1 \[[a\<^sub>1 f (k - 1) f k]]\ \a\<^sub>1 \ Q \ a\<^sub>n \ Q \ b \ Q\ - by presburger - then show ?thesis - using abc_abc_neq between_chain chain4 k_def(1) path_Q - by (metis (no_types) f2 \[[a\<^sub>1 f (k - 1) f k]]\ \a\<^sub>1 \ Q \ a\<^sub>n \ Q \ b \ Q\ - insertI1 insert_absorb) - qed *) - 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 ordering_def \[[a\<^sub>1 (f (k - 1)) (f k)]]\ less_imp_diff_less 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 chain_on_path ch_with_b some_betw Y_def unfolding ch_def - by (metis abc_sym insert_subset) - thus "[[(f (k-1)) b (f k)]]" - proof - - have "\ [[a\<^sub>1 (f k) b]]" - by (simp add: \[[a\<^sub>1 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 \[[a\<^sub>1 b (f k)]]\ \[[a\<^sub>1 (f (k - 1)) (f k)]]\) - 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)) (f k)]]" - using f_def k_def(2) \k-2 < k-1\ \k-1 < k\ unfolding long_ch_by_ord_def ordering_def - by blast - 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 "ordering2 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 "n\k-1" | "n\k+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) long_ch_by_ord_def 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 - - (* assume asm: "Suc k \ n" "?X = insert b Y" "b \ Y" "g n \ Y" *) - have "f n \ Y \ \(n < card Y)" for n - by (metis ordering_def f_def long_ch_by_ord_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 "x\Y" - show ?thesis - proof - - obtain ix where "f ix = x" "ix < card Y" - using \x \ Y\ f_def fin_Y - unfolding long_ch_by_ord_def ordering_def - by auto - have "ix\k-1 \ ix\k" - by linarith - thus ?thesis - proof - assume "ix\k-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 "ix\k" - 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 "x\Y" - 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 n\k-1 \ Suc n=k" if "n\k-1" - using \k \ 0\ that by linarith - have cases_ssn: "Suc(Suc n)\k-1 \ Suc(Suc n)=k" if "n\k-1" "Suc n\k-1" - using that(2) by linarith - - consider "n\k-1" | "n\k+1" | "n=k" - by linarith - then show "[[(g n) (g (Suc n)) (g (Suc (Suc n)))]]" - proof (cases) - assume "n\k-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 f_def g_def k_def(2) - by (metis (no_types, lifting) add_diff_inverse_nat lessI less_Suc_eq_le - less_imp_le_nat less_le_trans less_one 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 "n\k+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 nk + 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 long_ch_by_ord_def ordering_def - by auto - 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 "ordering2 g betw ?X" - unfolding ordering2_def - by presburger - qed - hence "long_ch_by_ord2 g ?X" - using Y_def f_def long_ch_by_ord2_def long_ch_by_ord_def - by auto - thus "[g[a\<^sub>1..b..a\<^sub>n]?X]" - unfolding fin_long_chain_def - using ch_equiv fin_X \a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n\ bound_indices k_def(2) Y_def g_def - by simp -qed - -(* jep *) -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 - -(* jep *) -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: "long_ch_by_ord f Y" - using ch_long_if_card_ge3 \4 \ card Y\ by fastforce - then obtain a\<^sub>1 a a\<^sub>n where long_ch_Y: "[f[a\<^sub>1..a..a\<^sub>n]Y]" - using \4 \ card Y\ get_fin_long_ch_bounds by fastforce - hence bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" - by (simp add: fin_long_chain_def) - have "a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n" - using \b \ Y\ abc_abc_neq fin_ch_betw long_ch_Y points_in_chain by blast - moreover have "a\<^sub>1 \ Q \ a\<^sub>n \ Q \ b \ Q" - using IH.prems(3) X_eq long_ch_Y points_in_chain by auto - ultimately consider "[[b a\<^sub>1 a\<^sub>n]]" | "[[a\<^sub>1 a\<^sub>n b]]" | "[[a\<^sub>n b a\<^sub>1]]" - using some_betw [of Q b a\<^sub>1 a\<^sub>n] \Q \ \

\ by blast - thus "ch X" - proof (cases) - (* case (i) *) - assume "[[b a\<^sub>1 a\<^sub>n]]" - have X_eq': "X = Y \ {b}" - using X_eq by auto - let ?g = "\j. if j \ 1 then f (j - 1) else b" - have "[?g[b..a\<^sub>1..a\<^sub>n]X]" - using chain_append_at_left_edge IH.prems(4) X_eq' \[[b a\<^sub>1 a\<^sub>n]]\ \b \ Y\ long_ch_Y X_eq - by presburger - thus "ch X" - using ch_by_ord_def ch_def fin_long_chain_def by auto - next - (* case (ii) *) - assume "[[a\<^sub>1 a\<^sub>n b]]" - let ?g = "\j. if j \ (card X - 2) then f j else b" - have "[?g[a\<^sub>1..a\<^sub>n..b]X]" - using chain_append_at_right_edge IH.prems(4) X_eq \[[a\<^sub>1 a\<^sub>n b]]\ \b \ Y\ long_ch_Y - by auto - thus "ch X" - unfolding ch_def ch_by_ord_def using fin_long_chain_def by auto - next - (* case (iii) *) - assume "[[a\<^sub>n b a\<^sub>1]]" (* jep: I have it this way so it matches some_betw, then switching it so it matches your old proof -- messy but easy to rectify, I'm just too tired to think too hard! *) - then have "[[a\<^sub>1 b a\<^sub>n]]" - by (simp add: abc_sym) - obtain k where - k_def: "[[a\<^sub>1 b (f k)]]" "k < card Y" "\ (\k'. 0 < k' \ k' < k \ [[a\<^sub>1 b (f k')]])" - using \[[a\<^sub>1 b a\<^sub>n]]\ \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 "[g[a\<^sub>1..b..a\<^sub>n]X]" - using chain_append_inside [of f a\<^sub>1 a a\<^sub>n Y b k] IH.prems(4) X_eq - \[[a\<^sub>1 b a\<^sub>n]]\ \b \ Y\ k_def long_ch_Y - by auto - thus "ch X" - using ch_by_ord_def ch_def fin_long_chain_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 "[f[a..b]X]" -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: "[f[a..b]X]" "a\X \ b\X" - using assms finX ch_X ch_some_betw get_fin_long_ch_bounds ch_long_if_card_ge3 - by (metis ch_by_ord_def ch_def fin_chain_def short_ch_def) - 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" "Q\P" "N=2" - and f_def: "[f[a..b]Q]" - 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) \ (\x\S. is_segment x) \ - P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" -proof - - have "a\Q \ b\Q \ a\b" - by (metis f_def fin_chain_def fin_long_chain_def points_in_chain) - hence "Q={a,b}" - using assms(3,5) - by (smt card_2_iff insert_absorb insert_commute insert_iff singleton_insert_inj_eq) - have "a\P \ b\P" - using \Q={a,b}\ assms(4) by auto - have "a\b" 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) \ - P1\P2={} \ (\x\{s}. (x\P1={} \ x\P2={} \ (\y\{s}. x\y \ x\y={})))" - proof (rule conjI) - { fix x assume "x\P" - have "[[a x b]] \ [[b a x]] \ [[a b x]] \ x=a \ x=b" - using \a\P \ b\P\ some_betw path_P \a\b\ - by (meson \x \ P\ abc_sym) - then have "x\s \ x\P1 \ x\P2 \ 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 "x\s \ x\P1 \ x\P2 \ 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 "x\P" - 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) \ P1\P2={} \ - (\x\{s}. (x\P1={} \ x\P2={} \ (\y\{s}. x\y \ x\y={})))" - 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 - { - fix x assume "x\P1" "x\P2" - show "x\{}" - using P1_def P2_def \x \ P1\ \x \ P2\ abc_only_cba pro_betw - by metis - } { - fix x assume "x\s" "x\P1" - show "x\{}" - using abc_only_cba seg_betw pro_betw P1_def \x \ P1\ \x \ s\ s_def - by (metis) - } { - fix x assume "x\s" "x\P2" - show "x\{}" - using abc_only_cba seg_betw pro_betw - by (metis P2_def \x \ P2\ \x \ s\ s_def) - } - qed - qed - thus ?thesis - by (simp add: S_def s_def) -qed - - - -lemma int_split_to_segs: - assumes f_def: "[f[a..b..c]Q]" - fixes S defines S_def: "S \ {segment (f i) (f(i+1)) | i. iS) \ Q" -proof - let ?N = "card Q" - have f_def_2: "a\Q \ b\Q \ c\Q" - using f_def points_in_chain by blast - hence "?N \ 3" - by (meson ch_by_ord_def f_def fin_long_chain_def long_ch_card_ge3) - have bound_indices: "f 0 = a \ f (card Q - 1) = c" - using f_def fin_long_chain_def 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 "p\Q" thus ?thesis by blast - next assume "p\Q" - hence "p\a \ p\c" - 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 n\<^sub>y n\<^sub>z y z - where yz_def: "y=f n\<^sub>y" "z=f n\<^sub>z" "[[y p z]]" "y\Q" "z\Q" "n\<^sub>y=n\<^sub>z-1" "n\<^sub>z(\k < card Q. [[(f k) p c]] \ k>n\<^sub>y)" "\(\kz. [[a p (f k)]])" - using get_closest_chain_events [where f=f and x=p and Y=Q and a\<^sub>n=c and a\<^sub>0=a and a=b] - f_def \p\Q\ - by metis - have "n\<^sub>y?s" - using \[[y p z]]\ abc_abc_neq seg_betw yz_def(1,2) - by blast - have "n\<^sub>z = n\<^sub>y + 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 "?s\S" - using S_def \n\<^sub>y assms(2) - by blast - hence "p\\S" - using \p \ ?s\ by blast - thus ?thesis by blast - qed - qed - show "?u\?i" - proof - fix p assume "p \ ?u" - hence "p\\S \ p\Q" by blast - thus "p\?i" - proof - assume "p\Q" - then consider "p=a"|"p=c"|"[[a p c]]" - using ch_all_betw_f f_def by blast - 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 "p\\S" - then obtain s where "p\s" "s\S" - by blast - then obtain y where "s = segment (f y) (f (y+1))" "yQ \ f (y+1) \ Q" - using f_def unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (meson add_lessD1) - have "[[a (f y) c]] \ y=0" - using \y assms(2) f_def fin_long_chain_def order_finite_chain by auto - moreover have "[[a (f (y+1)) c]] \ y = ?N-2" - using \y + 1 < card Q\ assms(2) f_def fin_long_chain_def order_finite_chain - by (smt One_nat_def Suc_diff_1 Suc_eq_plus1 diff_Suc_eq_diff_pred gr_implies_not0 - lessI less_Suc_eq_le linorder_neqE_nat not_le numeral_2_eq_2) - 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\ \?N\3\ - 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 \?N\3\ 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\ \?N\3\ - 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" "Q\P" "N\3" - and f_def: "a\Q \ b\Q \ c\Q" "[f[a..b..c]Q]" - 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: "a\P \ b\P \ c\P" - using assms(4) f_def by blast - have bound_indices: "f 0 = a \ f (card Q - 1) = c" - using f_def fin_long_chain_def by auto - have points_neq: "a\b \ b\c \ a\c" - using f_def fin_long_chain_def by auto - - text \The proof in two parts: subset inclusion one way, then the other.\ - { fix x assume "x\P" - 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 ch_all_betw_f f_def) - then have "(\s\S. x\s) \ x\P1 \ x\P2 \ x\Q" - 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 "x\Q \ x\\S" - 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 " x\P1 \ x\P2 \ x\Q" - 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 "(\s\S. x\s) \ x\P1 \ x\P2 \ x\Q" - by blast - hence "x\\S \ [[b a x]] \ [[b c x]] \ x\Q" - using S_def P1_def P2_def - unfolding segment_def prolongation_def - by auto - hence "x\P" - proof (cases) - assume "x\\S" - have "S = {segment (f i) (f(i+1)) | i. iinterval a c" - using int_split_to_segs [OF f_def(2)] assms \x\\S\ - 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 "x\\S" - hence "[[b a x]] \ [[b c x]] \ x\Q" - 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" "Q\P" "N\3" - and f_def: "a\Q \ b\Q \ c\Q" "[f[a..b..c]Q]" - and S_def: "S = {s. \i<(N-1). s = segment (f i) (f (i+1))}" - and x_def: "x\s" "s\S" - shows "[[a x c]]" -proof - - have inseg_neq_ac: "x\a \ x\c" if "x\s" "s\S" for x s - proof - show "x\a" - proof (rule notI) - assume "x=a" - obtain n where s_def: "s = segment (f n) (f (n+1))" "ns \ S\ by blast - have "f n \ Q" - using f_def \n < N - 1\ fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis assms(3) diff_diff_cancel less_imp_diff_less less_irrefl_nat not_less) - hence "[[a (f n) c]]" - using f_def fin_long_chain_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 ch_all_betw_f inside_not_bound(2) less_diff_conv) - moreover have "[[(f(n)) x (f(n+1))]]" - using \x\s\ seg_betw s_def(1) by simp - ultimately show False - using \x=a\ abc_only_cba(1) assms(3) f_def fin_long_chain_def s_def(2) order_finite_chain - by (metis le_numeral_extra(3) less_add_one less_diff_conv neq0_conv) - qed - - show "x\c" - proof (rule notI) - assume "x=c" - obtain n where s_def: "s = segment (f n) (f (n+1))" "ns \ S\ by blast - hence "n+1x\s\ seg_betw s_def(1) by simp - have "f (n) \ Q" - using f_def \n+1 < N\ fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis add_lessD1 assms(3)) - have "f (n+1) \ Q" - using f_def \n+1 < N\ fin_long_chain_def long_ch_by_ord_def 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 fin_long_chain_def assms(3) order_finite_chain seg_betw that(1) - abc_abc_neq abc_only_cba ch_all_betw_f - 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) fin_long_chain_def order_finite_chain - by (metis \f n \ Q\ abc_bcd_acd abc_only_cba(1,2) ch_all_betw_f) - qed - qed - - show "[[a x c]]" - proof - - have "x\interval a c" - using int_split_to_segs [OF f_def(2)] S_def assms(2,3,5) x_def - by blast - have "x\a \ x\c" 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" "Q\P" "N\3" - and f_def: "a\Q \ b\Q \ c\Q" "[f[a..b..c]Q]" - 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 "P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" -proof (rule conjI) - show "P1 \ P2 = {}" - proof (safe) - fix x assume "x\P1" "x\P2" - 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 "\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={}))" - proof (rule ballI) - fix s assume "s\S" - show "s \ P1 = {} \ s \ P2 = {} \ (\y\S. s \ y \ s \ y = {})" - proof (intro conjI ballI impI) - show "s\P1={}" - proof (safe) - fix x assume "x\s" "x\P1" - 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 "s\P2={}" - proof (safe) - fix x assume "x\s" "x\P2" - 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 "r\S" "s\r" - show "s\r={}" - 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))" - "n\m" "nr \ 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 \y\r\ \y\s\ rs_def(1,2) by simp - have False - proof (cases) - assume "nn < m\ assms(3) f_def fin_long_chain_def order_finite_chain rs_def(5) by auto - have "n+1[[(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 f_def assms(3) rs_def(5) - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis add_lessD1 less_add_one less_diff_conv) - 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 "\nm" using nat_neq_iff rs_def(3) by blast - have "[[(f m) (f n) (f(n+1))]]" - using \n > m\ assms(3) f_def fin_long_chain_def order_finite_chain rs_def(4) by auto - hence "m+1n > 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 f_def assms(3) rs_def(4) - unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis add_lessD1 less_add_one less_diff_conv) - 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" "Q\P" "N\3" - and f_def: "a\Q \ b\Q \ c\Q" "[f[a..b..c]Q]" - 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) \ - (\x\S. is_segment x) \ - P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" -proof - - have "P = ((\S) \ P1 \ P2 \ Q) \ - (\x\S. is_segment x) \ P1\P2={} \ - (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" - proof (intro disjoint_segmentation conjI) - show "P = ((\S) \ P1 \ P2 \ Q)" - using path_is_union assms - by blast - show "\x\S. is_segment x" - proof - fix s assume "s\S" - thus "is_segment s" using S_def by auto - qed - qed (use assms disjoint_segmentation in auto) - then show ?thesis by auto -qed - -text \ - We define \disjoint\ to be the same as in HOL-Library.DisjointSets. - This saves importing a lot of baggage we don't need. - The two lemmas below are just for safety. -\ - -abbreviation disjoint - where "disjoint A \ (\a\A. \b\A. a \ b \ a \ b = {})" - -lemma - fixes S:: "('a set) set" and P1:: "'a set" and P2:: "'a set" - assumes "\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={}))" "P1\P2={}" - shows "disjoint (S\{P1,P2})" -proof (rule ballI) - let ?U = "S\{P1,P2}" - fix a assume "a\?U" - then consider (aS)"a\S"|(a1)"a=P1"|(a2)"a=P2" - by fastforce - thus "\b\?U. a \ b \ a \ b = {}" - proof cases - case aS - { fix b assume "b\?U" "a\b" - then consider "b\S"|"b=P1"|"b=P2" - by fastforce - hence "a\b={}" - apply cases - apply (simp add: \a \ S\ \a \ b\ assms) - apply (meson \a \ S\ assms) - by (simp add: \a \ S\ assms) - } - thus ?thesis - by meson - next - case a1 - { fix b assume "b\?U" "a\b" - then consider "b\S"|"b=P2" - using a1 by fastforce - hence "a\b={}" - apply cases - apply (metis a1 assms(1) inf_commute) - by (simp add: a1 assms(2)) - } - thus ?thesis - by meson - next - case a2 - { fix b assume "b\?U" "a\b" - then consider "b\S"|"b=P1" - using a2 by fastforce - hence "a\b={}" - apply cases - apply (metis a2 assms(1) inf_commute) - by (simp add: a2 assms(2) inf_commute) - } - thus ?thesis - by meson - qed -qed -lemma - fixes S:: "('a set) set" and P1:: "'a set" and P2:: "'a set" - assumes "disjoint (S\{P1,P2})" "P1\S" "P2\S" "P1\P2" - shows "\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={}))" "P1\P2={}" -proof (rule ballI) - show "P1\P2={}" - using assms(1,4) by simp - fix x assume "x\S" - show "x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})" - proof (rule conjI, rule_tac[2] conjI, rule_tac[3] ballI, rule_tac[3] impI) - show "x\P1={}" - using \x \ S\ assms(1,2) by fastforce - show "x\P2={}" - using \x \ S\ assms(1,3) by fastforce - fix y assume "y\S" "x\y" - thus "x\y={}" - by (simp add: \x \ S\ assms(1)) - qed -qed - - -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: "Q\P" - and f_def: "[f[a..b]Q]" - fixes P1 defines P1_def: "P1 \ prolongation b a" - fixes P2 defines P2_def: "P2 \ prolongation a b" - fixes S defines S_def: "S \ if card Q=2 then {segment a b} - else {segment (f i) (f (i+1)) | i. iS) \ P1 \ P2 \ Q)" "(\x\S. is_segment x)" - "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" -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)) -(* Hooray for theorem 10! Without it, we couldn't so brazenly go from a set of events - to an ordered chain of events. The line below doesn't need f_def (which is needed for S_def)! *) - have ch_Q: "ch Q" - using Q_def card_Q path_P path_finsubset_chain [where X=Q and Q=P] - by blast - have f_def_2: "a\Q \ b\Q" - using f_def points_in_chain fin_chain_def by auto - have "a\b" - using f_def fin_chain_def fin_long_chain_def by auto - { - assume "card Q = 2" - hence "S={segment a b}" - by (simp add: S_def) - have "P = ((\S) \ P1 \ P2 \ Q)" "(\x\S. is_segment x)" "P1\P2={}" - "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" - using assms ch_Q \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: "[f[a..c..b]Q]" - using assms(3,5) \a\b\ - by (metis f_def fin_chain_def short_ch_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)" "(\x\S. is_segment x)" "P1\P2={}" - "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" - using f_def_2 assms ch_Q \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_chain \finite Q\ S_def2 by presburger+ - } - ultimately have old_thesis: "P = ((\S) \ P1 \ P2 \ Q)" "(\x\S. is_segment x)" "P1\P2={}" - "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" by meson+ - thus "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" - "P = ((\S) \ P1 \ P2 \ Q)" "(\x\S. is_segment x)" - 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 Q\2" "Q\P" - shows "\S P1 P2. P = ((\S) \ P1 \ P2 \ Q) \ - disjoint (S\{P1,P2}) \ P1\P2 \ P1\S \ P2\S \ - (\x\S. is_segment x) \ is_prolongation P1 \ is_prolongation P2" -proof - - let ?N = "card Q" - obtain f a b where f_def: "[f[a..b]Q]" - using path_finsubset_chain2[OF path_P Q_def(2,1)] - by metis - let ?S = "if ?N=2 then {segment a b} else {segment (f i) (f (i+1)) | i. i?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) \[f[a..b]Q]\] - by force+ - thus ?thesis - by blast -qed - - - -end (* context MinkowskiSpacetime *) - - -section "Chains are unique up to reversal" - -lemma (in MinkowskiSpacetime) chain_remove_at_right_edge: - assumes "[f[a..c]X]" "f (card X - 2) = p" "3 \ card X" "X = insert c Y" "c\Y" - shows "[f[a..p]Y]" -proof - - - have lch_X: "long_ch_by_ord f X" - using assms(1,3) fin_chain_def fin_long_chain_def ch_by_ord_def short_ch_card_2 - by fastforce - have "p\X" - by (metis ordering_def assms(2,3) card.empty card_gt_0_iff diff_less lch_X - 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 fin_chain_def fin_long_chain_def - by (metis (no_types, opaque_lifting) One_nat_def Suc_1 ch_by_ord_def diff_Suc_Suc - less_Suc_eq_le neq0_conv numeral_3_eq_3 short_ch_card_2 zero_less_diff) - - 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 assms(2) lch_X bound_ind \3 \ card X\ unfolding long_ch_by_ord_def ordering_def - by (metis One_nat_def diff_Suc_less less_le_trans zero_less_numeral) - qed - hence "p\c" - using abc_abc_neq by blast - hence "p\Y" - using \p \ X\ assms(4) by blast - - show ?thesis - proof (cases) - assume "3 = card X" - hence "2 = card Y" - by (metis assms(4,5) card.insert card.infinite diff_Suc_1 finite_insert nat.simps(3) - numeral_2_eq_2 numeral_3_eq_3) - have "a\p" - using \[[a p c]]\ abc_abc_neq by auto - moreover have "a\Y \ p\Y" - using \[[a p c]]\ \p \ Y\ abc_abc_neq assms(1,4) fin_chain_def points_in_chain - by fastforce - moreover have "short_ch Y" - proof - - obtain ap where "path ap a p" - using \[[a p c]]\ abc_ex_path_unique calculation(1) by blast - hence "\Q. path Q a p" - by blast - moreover have "\ (\z\Y. z \ a \ z \ p)" - using \2 = card Y\ \a \ Y \ p \ Y\ \a \ p\ - by (metis card_2_iff') - ultimately show ?thesis - unfolding short_ch_def using \a \ Y \ p \ Y\ - by blast - qed - ultimately show ?thesis unfolding fin_chain_def by blast - next - assume "3 \ card X" - hence "4 \ card X" - using assms(3) by auto - - obtain b where "b = f 1" by simp - have "\b. [f[a..b..p]Y]" - proof - have "[[a b p]]" - using bound_ind \b = f 1\ \3 \ card X\ assms(2,3) lch_X order_finite_chain - by fastforce - hence all_neq: "b\a \ b\p \ a\p" - using abc_abc_neq by blast - have "b\X" - using \b = f 1\ lch_X assms(3) unfolding long_ch_by_ord_def ordering_def - by force - hence "b\Y" - using \[[a b p]]\ \[[a p c]]\ abc_only_cba(2) assms(4) by blast - - have "ordering f betw Y" - unfolding ordering_def - proof (safe) - show "\n. infinite Y \ f n \ Y" - using assms(3) assms(4) by auto - show "\n. n < card Y \ f n \ Y" - using assms(3,4,5) bound_ind lch_X - unfolding long_ch_by_ord_def ordering_def - using get_fin_long_ch_bounds indices_neq_imp_events_neq - by (smt Suc_less_eq add_leD1 cancel_comm_monoid_add_class.diff_cancel card_Diff1_less - card_Diff_singleton card_eq_0_iff card_insert_disjoint gr_implies_not0 insert_iff lch_X - le_add_diff_inverse less_SucI numeral_3_eq_3 plus_1_eq_Suc zero_less_diff) - { - fix x assume "x\Y" - hence "x\X" - using assms(4) by blast - then obtain n where "n < card X" "f n = x" - using lch_X unfolding long_ch_by_ord_def ordering_def - using assms(3) by auto - show "\n. (finite Y \ n < card Y) \ f n = x" - proof - show "(finite Y \ n < card Y) \ f n = x" - using \f n = x\ \n < card X\ \x \ Y\ assms(4,5) bound_ind - by (metis Diff_insert_absorb card.remove card_Diff_singleton - finite.insertI insertI1 less_SucE) - qed - } - fix n n' n'' - assume "(n::nat)\n. infinite Y \ f n \ Y\ \infinite Y\ assms(5) bound_ind by blast - } { - assume "n'' < card Y" - show "[[(f n) (f n') (f n'')]]" - using \n < n'\ \n' < n''\ \n'' < card Y\ assms(4,5) lch_X order_finite_chain - using \infinite Y \ [[(f n) (f n') (f n'')]]\ by fastforce - } - qed - hence lch_Y: "long_ch_by_ord f Y" - using \[[a p c]]\ \b \ Y\ \p \ X\ abc_abc_neq all_neq assms(4) bound_ind - long_ch_by_ord_def zero_into_ordering - by fastforce - - show "[f[a..b..p]Y]" - using all_neq lch_Y bound_ind \b\Y\ assms(2,3,4,5) unfolding fin_long_chain_def - by (metis Diff_insert_absorb One_nat_def add_leD1 card.infinite finite_insert plus_1_eq_Suc - diff_diff_left card_Diff_singleton not_one_le_zero insertI1 numeral_2_eq_2 numeral_3_eq_3) - qed - - thus ?thesis unfolding fin_chain_def - using points_in_chain by blast - qed -qed - - -lemma (in MinkowskiChain) fin_long_ch_imp_fin_ch: - assumes "[f[a..b..c]X]" - shows "[f[a..c]X]" - using assms fin_chain_def points_in_chain by auto - - -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 (in MinkowskiSpacetime) chain_unique_induction_ax: - assumes "card X \ 3" - and "i < card X" - and "[f[a..c]X]" - and "[g[x..z]X]" - 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: "[f[a..b..c]X]" - by (metis Nil.prems(1,3) fin_chain_def short_ch_def three_in_set3) - obtain y where g_ch: "[g[x..y..z]X]" - using Nil.prems fin_chain_def short_ch_card_2 - by (metis Suc_n_not_le_n ch_by_ord_def numeral_2_eq_2 numeral_3_eq_3) - - 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_chain by blast - hence "(g i = a \ g i = c)" - using \g i \ b\ \card X = 3\ points_in_chain - by (smt f_ch card2_either_elt1_or_elt2 card_Diff_singleton diff_Suc_1 - fin_long_chain_def insert_Diff insert_iff numeral_2_eq_2 numeral_3_eq_3) - 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 fin_long_chain_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 - by (metis diff_Suc_1 fin_long_chain_def numeral_2_eq_2 numeral_3_eq_3) - qed -next - case IH: (Suc n) - have lch_fX: "long_ch_by_ord f X" - using ch_by_ord_def fin_chain_def fin_long_chain_def long_ch_card_ge3 IH(3,5) - by fastforce - have lch_gX: "long_ch_by_ord g X" - using IH(3,6) ch_by_ord_def fin_chain_def fin_long_chain_def 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: "[f[a..b..c]X]" - using \ch_by_ord f X\ IH(3,5) fin_chain_def short_ch_card_2 - by auto - obtain y where g_ch: "[g[x..y..z]X]" - using \ch_by_ord f X\ IH.prems(1,4) fin_chain_def short_ch_card_2 - by auto - - 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 unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by (metis card_Diff1_less card_Diff_singleton) - qed - hence "p\c \ p\a" - using abc_abc_neq by blast - - obtain Y where Y_def: "X = insert c Y" "c\Y" - using f_ch points_in_chain - by (meson mk_disjoint_insert) - hence fin_Y: "finite Y" - using f_ch fin_long_chain_def by auto - hence "n = card Y - 3" - using \Suc n = card X - 3\ \X = insert c Y\ \c\Y\ 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 "p\Y" - 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 fin_chain_def points_in_chain) - have "[f[a..p]Y]" - 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: "long_ch_by_ord f Y" - unfolding fin_chain_def - using card_Y ch_by_ord_def fin_Y fin_long_chain_def long_ch_card_ge3 - by force - - have p_closest: "\ (\q\X. [[p q c]])" - proof - assume "(\q\X. [[p q c]])" - then obtain q where "q\X" "[[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 \p\c \ p\a\ - by (metis ordering_def long_ch_by_ord_def) - have "j > card X - 2 \ j < card X - 1" - proof - - have "j > card X - 2 \ j < card X - 1 \ j < card X - 2 \ j > card X - 1" - using index_order3 [where b=j and a="card X - 2" and c="card X - 1"] - using \[[p q c]]\ \f j = q\ \j < card X\ f_ch p_def - by (metis (no_types, lifting) One_nat_def card_gt_0_iff diff_less empty_iff - fin_long_chain_def lessI zero_less_numeral) - 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\ \p\Y\ points_in_chain \p\c \ p\a\ - by (metis (no_types, opaque_lifting) chain_bounds_unique f_ch - fin_long_chain_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\ - unfolding long_ch_by_ord_def ordering_def - by (metis (mono_tags, lifting) One_nat_def card_Diff1_less card_Diff_singleton - diff_diff_left fin_long_chain_def g_ch numeral_2_eq_2 plus_1_eq_Suc) - qed - moreover have "d\X" - using lch_gX \d = g (card X - 2)\ unfolding long_ch_by_ord_def ordering_def - by auto - ultimately show False - using p_closest abc_sym IH.prems(5) chain_bounds_unique f_ch g_ch - by blast - qed - - hence ch_gY: "long_ch_by_ord g Y" - using IH.prems(1,4,5) g_ch f_ch ch_fY card_Y ch_by_ord_def chain_remove_at_right_edge fin_Y - by (metis Y_def chain_bounds_unique fin_chain_def fin_long_chain_def long_ch_card_ge3) - - have "f i \ Y \ f i = c" - by (metis ordering_def \X = insert c Y\ \i < card X\ lch_fX insert_iff 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\ \c\Y\ IH(3,4) f_ch fin_Y fin_long_chain_def not_less_less_Suc_eq - by fastforce - 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_unique2) - 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 - by (metis \card Y = card X - 1\ Y_def card_insert_disjoint fin_Y fin_long_chain_def lessI) - qed -qed - -text \I'm really impressed \sledgehammer\/\smt\ can solve this if I just tell them "Use symmetry!".\ - -lemma (in MinkowskiSpacetime) chain_unique_induction_cx: - assumes "card X \ 3" - and "i < card X" - and "[f[a..c]X]" - and "[g[x..z]X]" - and "c = x \ a = z" - shows "f i = g (card X - i - 1)" - using chain_sym chain_unique_induction_ax - by (smt (verit, best) assms diff_right_commute fin_chain_def fin_long_ch_imp_fin_ch) - - - -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) "ordering" of the chain. - This could be made generic over the ordering similar to \chain_sym\ relying on \ordering_sym\. -\ - -lemma (in MinkowskiSpacetime) chain_unique_upto_rev_cases: - assumes ch_f: "[f[a..c]X]" - and ch_g: "[g[x..z]X]" - 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 (in MinkowskiSpacetime) chain_unique_upto_rev: - assumes "[f[a..c]X]" "[g[x..z]X]" "card X \ 3" "i < card X" - shows "f i = g i \ f i = g (card X - i - 1)" "a=x\c=z \ c=x\a=z" -proof - - have "(a=x \ c=z) \ (a=z \ c=x)" - using chain_bounds_unique - by (metis assms(1,2) fin_chain_def points_in_chain short_ch_def) - 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=x\c=z) \ (c=x\a=z)" - by (meson assms(1-3) chain_bounds_unique2) -qed - - - -section "Subchains" -context MinkowskiSpacetime begin - - -lemma f_img_is_subset: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" - shows "Y\X" -proof - fix x assume "x\Y" - then obtain n where "n\{i..j}" "f n = x" - using assms(4) by blast - hence "f n \ X" - by (metis ordering_def assms(1) inf_chain_is_long long_ch_by_ord_def) - thus "x\X" - using \f n = x\ by blast -qed - - -lemma f_inj_on_index_subset: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" - shows "inj_on f {i..j}" - unfolding inj_on_def -proof (safe) - fix x y assume "x\{i..j}" "y\{i..j}" "f x = f y" - show "x=y" - proof (rule ccontr) - assume "x\y" - let ?P = "\r s. f r \ f s" - { - assume "x\y" - hence "xx \ y\ le_imp_less_or_eq by blast - obtain n where "n>y" by blast - hence "[[(f x)(f y)(f n)]]" - using assms(1) \x inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk by fastforce - hence "?P x y" - using abc_abc_neq by blast - } moreover { (* TODO: use a wlog theorem for sets instead? *) - assume "x>y" - obtain n where "n>x" by blast - hence "[[(f y)(f x)(f n)]]" - using assms(1) \x>y\ inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk by fastforce - hence "?P y x" - using abc_abc_neq by blast - } - ultimately show False - using not_le_imp_less \f x = f y\ by auto - qed -qed - - -lemma f_bij_on_index_subset: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" - shows "bij_betw f {i..j} Y" - using f_inj_on_index_subset - by (metis assms inj_on_imp_bij_betw) - - -lemma only_one_index: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" "f n \ Y" - shows "n\{i..j}" -proof - - obtain m where "m\{i..j}" "f m = f n" - using assms(4) assms(5) by auto - have "inj_on f {i..j}" - using assms(1,3) f_inj_on_index_subset by blast - have "m = n" - proof (rule ccontr) - assume "m\n" - obtain l where "f l \ X" "l\m" "l\n" - using assms(1) inf_chain_is_long - by (metis ordering_def le_eq_less_or_eq lessI long_ch_by_ord_def not_less_eq_eq) - hence "[[(f l)(f m)(f n)]] \ [[(f m)(f l)(f n)]] \ [[(f l)(f n)(f m)]]" - using \f m = f n\ \m\n\ - using abc_abc_neq assms(1) inf_chain_is_long inf_ordering_inj' long_ch_by_ord_def - by blast - thus False - using \f m = f n\ abc_abc_neq by auto - qed - thus ?thesis - using \m \ {i..j}\ by blast -qed - - -lemma f_one_to_one_on_index_subset: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" "y\Y" - shows "\!k\{i..j}. f k = y" "f k = y \ k\{i..j}" - using f_inj_on_index_subset only_one_index assms image_iff inj_on_eq_iff apply metis - using assms(1,3,4,5) only_one_index by blast - - -lemma card_of_subchain: - assumes "[f[(f 0) ..]X]" "i\0" "j>i" "Y=f`{i..j}" - shows "card Y = card {i..j}" "card Y = j-i+1" -proof - - show "card Y = card {i..j}" - by (metis assms bij_betw_same_card f_bij_on_index_subset) - thus "card Y = j-i+1" - using card_Collect_nat - by (simp add: assms(3)) -qed - - -lemma fin_long_subchain_of_semifin: - assumes "[f[(f 0) ..]X]" "i\0" "j>i+1" "Y=f`{i..j}" - "g = (\n. f(n+i))" - shows "[g[(f i)..(f j)]Y]" (* "j=i+1 \ short_ch Y" "j>i+1 \ fin_long_ch_by_ord g Y" *) -proof - - obtain k where "k=i+1" by simp - hence ind_ord: "i k f k \ f i \ f j \ f k \ f j" - proof - - have "[[(f i) (f k) (f j)]]" - using assms(1) ind_ord long_ch_by_ord_def ordering_ord_ijk semifin_chain_def - by fastforce - thus ?thesis - using abc_abc_neq by blast - qed - moreover have "finite Y" - proof - - have "inj f" - using inf_ordering_inj [where ord="betw"] abc_abc_neq - using assms(1) long_ch_by_ord_def semifin_chain_def by auto - hence "card Y \ card {i..j}" - using assms(4) inf_ordering_inj - using card_image_le by blast - have "finite {i..j}" - by simp - thus "finite Y" - by (simp add: assms(4)) - qed - moreover have "long_ch_by_ord g Y" - proof - - obtain x y z where "x=f i" "y=f k" "z=f j" - by auto - have "x\Y \ y\Y \ z\Y \ x \ y \ y \ z \ x \ z" - using \x = f i\ \y = f k\ \z = f j\ assms(4) calculation(1) ind_ord by auto - moreover have "ordering g betw Y" - unfolding ordering_def - proof (intro conjI) - show "\n. (finite Y \ n < card Y) \ g n \ Y" - apply (safe) apply (auto simp add: \finite Y\) - proof - - fix n assume "n{i..j}" - proof - - assume asm: "\n'. \n + i = n'; n' \ {i..j}\ \ thesis" - have "n < card {i..j}" - by (metis \n < card Y\ assms(4) card_image_le finite_atLeastAtMost less_le_trans) - thus ?thesis - using asm by simp - qed - show "g n \ Y" - using \n + i = n'\ \n' \ {i..j}\ assms(4,5) by blast - qed - show "\x\Y. \n. (finite Y \ n < card Y) \ g n = x" - proof (rule ballI) - fix x assume "x\Y" - hence "x\X" - using f_img_is_subset assms(1,4) - by (metis ordering_def imageE inf_chain_is_long long_ch_by_ord_def) - then obtain n where "f n = x" - using \x \ Y\ assms(4) by blast - have "n\{i..j}" using only_one_index - by (metis \f n = x\ \x \ Y\ assms(1,2,4) ind_ord less_trans) - show "\n. (finite Y \ n < card Y) \ g n = x" - proof (rule exI, rule conjI) - have "n-i\0" - by blast - have "g (n-i) = f (n-i+i)" - using assms(5) by blast - show "g (n-i) = x" - proof (cases) - assume "n-i>0" - thus ?thesis - by (simp add: \f n = x\ \g (n - i) = f (n - i + i)\) - next assume "\n-i>0" - hence "n-i=0" by blast - thus ?thesis - using \n\{i..j}\ \f n = x\ \g (n - i) = f (n - i + i)\ by auto - qed - show "finite Y \ (n-i) < card Y" - proof - assume "finite Y" - show "n-in \ {i..j}\ assms(1,4) ind_ord by auto - qed - qed - qed - show "\n n' n''. (finite Y \ n'' < card Y) \ n n' [[(g n)(g n')(g n'')]]" - apply (safe) using \finite Y\ apply blast - proof - - fix l m n - assume "lm < n\) - hence "[[(f(l+i))(f(m+i))(f(n+i))]]" - using assms(1) inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk by fastforce - thus "[[(g l)(g m)(g n)]]" - using assms(5) by blast - qed - qed - ultimately show ?thesis - using long_ch_by_ord_def by auto - qed - moreover have "g 0 = f i \ f k \ Y \ g (card Y - 1) = f j" - using card_of_subchain assms(1,4,5) ind_ord less_imp_le_nat - by force - ultimately show ?thesis - using fin_long_chain_def by blast - qed - thus ?thesis - using fin_long_ch_imp_fin_ch by blast -qed - -end (*Context MinkowskiSpacetime*) - - - -section "Extensions of results to infinite chains" -context MinkowskiSpacetime begin - -lemma i_neq_j_imp_events_neq_inf: - assumes "[f[(f 0)..]X]" "i\j" - shows "f i \ f j" -proof - - let ?P = "\ i j. i\j \ f i \ f j" - { - fix i j assume "(i::nat)\j" - have "?P i j" - proof (cases) - assume "ij" by blast - hence "[[(f i)(f j)(f k)]]" - using \i < j\ assms(1) inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk by fastforce - thus "?P i j" - using abc_abc_neq by blast - next - assume "\ii \ j\ by auto - show "?P i j" by (simp add: \i = j\) - qed - } moreover { - fix i j assume "?P j i" - hence "?P i j" by auto - } - ultimately show ?thesis - by (metis assms(2) leI less_imp_le_nat) -qed - - -lemma i_neq_j_imp_events_neq: - assumes "long_ch_by_ord f X" "i\j" "finite X \ (i j f j" - using i_neq_j_imp_events_neq_inf indices_neq_imp_events_neq - by (meson assms get_fin_long_ch_bounds semifin_chain_def) - - -lemma inf_chain_origin_unique: - assumes "[f[f 0..]X]" "[g[g 0..]X]" - shows "f 0 = g 0" -proof (rule ccontr) - assume "f 0 \ g 0" - obtain P where "P\\

" "X\P" - using assms(1) semifin_chain_on_path by blast - obtain x where "x = g 1" by simp - hence "x\g 0" - using assms(2) i_neq_j_imp_events_neq_inf zero_neq_one by blast - have "x\X" - by (metis ordering_def \x = g 1\ assms(2) inf_chain_is_long long_ch_by_ord_def) - have "x=f 0 \ x\f 0" by auto - thus False - proof (rule disjE) - assume "x=f 0" - hence "[[(g 0)(f 0)(g 2)]]" - using \x=g 1\ \x=f 0\ assms(2) inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk - by fastforce - then obtain m n where "f m = g 0" "f n = g 2" - by (metis ordering_def assms(1) assms(2) inf_chain_is_long long_ch_by_ord_def) - hence "[[(f m)(f 0)(f n)]]" - by (simp add: \[[(g 0)(f 0)(g 2)]]\) - hence "m\n" - using abc_abc_neq by blast - have "m>0 \ n>0" - using \[[(f m)(f 0)(f n)]]\ abc_abc_neq neq0_conv by blast - hence "(0 m (0 nm \ n\ by auto - thus False - using \[[(f m)(f 0)(f n)]]\ assms(1) index_order3 inf_chain_is_long by blast - next - assume "x\f 0" - - (*Help for Sledgehammer*) - have fn: "\n. f n \ X" - by (metis (no_types) ordering_def assms(1) inf_chain_is_long long_ch_by_ord_def) - have gn: "\n. g n \ X" - by (metis ordering_def assms(2) inf_chain_is_long long_ch_by_ord_def) - - have "[[(g 0)x(f 0)]]" - proof - - have "[[(f 0)(g 0)x]] \ [[(g 0)(f 0)x]] \ [[(g 0)x(f 0)]]" - using \f 0 \ g 0\ \x \ f 0\ \x \ g 0\ all_aligned_on_semifin_chain - by (metis ordering_def \x \ X\ assms inf_chain_is_long long_ch_by_ord_def) - moreover have "\[[(f 0)(g 0)x]]" - using abc_only_cba(1,3) all_aligned_on_semifin_chain assms(2) fn - by (metis \x\X\ \x\f 0\ \x\g 0\) - moreover have "\[[(g 0)(f 0)x]]" - using fn gn \x \ X\ \x \ g 0\ - by (metis (no_types) abc_only_cba(1,2,4) all_aligned_on_semifin_chain assms(1)) - ultimately show ?thesis by blast - qed - - obtain m m' where "g m' = f 0" "m = Suc m'" - using ordering_def assms inf_chain_is_long long_ch_by_ord_def by metis - hence "[[(g 0)(f 0)(g m)]]" - by (metis Suc_le_eq \f 0 \ g 0\ assms(2) inf_chain_is_long lessI linorder_neqE_nat - long_ch_by_ord_def not_le ordering_ord_ijk zero_less_Suc) - then obtain n p where "f n = g 0" "f p = g m" - by (metis abc_abc_neq abc_only_cba(1,4) all_aligned_on_semifin_chain assms(1) gn) - hence "m<0 \ n<0" - using all_aligned_on_semifin_chain assms(1) \[[(g 0)(f 0)(g m)]]\ - by (metis abc_abc_neq abc_only_cba(1,4) fn) - thus False by simp - qed -qed - - -lemma inf_chain_unique: - assumes "[f[f 0..]X]" "[g[g 0..]X]" - shows "\i::nat. f i = g i" -proof - - { - assume asm: "[f[f 0..]X]" "[g[f 0..]X]" - have "\i::nat. f i = g i" - proof - fix i::nat - show "f i = g i" - proof (induct i) - show "f 0 = g 0" - using asm(2) inf_chain_is_long by fastforce - fix i assume "f i = g i" - show "f (Suc i) = g (Suc i)" - proof (rule ccontr) - assume "f (Suc i) \ g (Suc i)" - let ?i = "Suc i" - have "f 0\X \ g?i\X \ f?i\X" - by (metis ordering_def assms(1) assms(2) inf_chain_is_long long_ch_by_ord_def) - hence "[[(f 0)(f ?i)(g ?i)]] \ [[(f 0)(g ?i)(f ?i)]] \ [[(f ?i)(f 0)(g ?i)]]" - using all_aligned_on_semifin_chain assms(1,2) i_neq_j_imp_events_neq_inf - by (metis \f?i \g?i\ \f 0 = g 0\) - hence "[[(f 0)(f ?i)(g ?i)]] \ [[(f 0)(g ?i)(f ?i)]]" - using all_aligned_on_semifin_chain asm(2) - by (metis \f 0 \ X \ g (Suc i) \ X \ f (Suc i) \ X\ abc_abc_neq) - have "([[(f 0)(f i)(f ?i)]] \ [[(f 0)(g i)(g ?i)]]) \ i=0" - using long_ch_by_ord_def ordering_ord_ijk asm(1,2) - by (metis Suc_inject Suc_lessI Suc_less_eq inf_chain_is_long lessI zero_less_Suc) - thus False - proof (rule disjE) - assume "i=0" - have "[[(g 0)(f 1)(g 1)]]" - proof - - obtain x where "x = g 1" by simp - hence "x\X" - using \f 0 \ X \ g (Suc i) \ X \ f (Suc i) \ X\ \i = 0\ by force - then obtain m where "f m = x" - by (metis ordering_def assms(1) inf_chain_is_long long_ch_by_ord_def) - hence "f m = g 1" - using \x = g 1\ by blast - have "m>1" - using assms(2) i_neq_j_imp_events_neq_inf \f?i \ g?i\ - by (metis One_nat_def Suc_lessI \f 0 = g 0\ \f m = x\ \i = 0\ \x = g 1\ neq0_conv) - thus "[[(g 0)(f 1)(g 1)]]" - using \[[(f 0)(f?i)(g?i)]] \ [[(f 0)(g?i)(f?i)]]\ \f 0 = g 0\ \f m = x\ \i=0\ \x = g 1\ - by (metis One_nat_def assms(1) gr_implies_not_zero index_order3 inf_chain_is_long order.asym) - qed - have "f 1 \ X" - using \f 0 \ X \ g (Suc i) \ X \ f (Suc i) \ X\ \i = 0\ by auto - then obtain m' where "g m' = f 1" - by (metis ordering_def assms(2) inf_chain_is_long long_ch_by_ord_def) - hence "[[(g 0)(g m')(g 1)]]" - using \[[(g 0)(f 1)(g 1)]]\ by auto - have "[[(g 0)(g 1)(g m')]]" - proof - - have "m' \ 1 \ m' \ 0" - using \[[(g 0)(g m')(g 1)]]\ by (meson abc_abc_neq) - hence "m'>1" by auto - thus "[[(g 0)(g 1)(g m')]]" - using \[[(g 0)(g m')(g 1)]]\ assms(2) index_order3 inf_chain_is_long by blast - qed - thus False - using \[[(g 0)(g m')(g 1)]]\ abc_only_cba(2) by blast - next - assume "[[(f 0)(f i)(f ?i)]] \ [[(f 0)(g i)(g ?i)]]" - have "[[(g 0)(f ?i)(g ?i)]]" - proof - - obtain x where "x = g ?i" by simp - hence "x\X" - by (simp add: \f 0 \ X \ g (Suc i) \ X \ f (Suc i) \ X\) - then obtain m where "f m = x" - by (metis ordering_def assms(1) inf_chain_is_long long_ch_by_ord_def) - hence "f m = g ?i" - using \x = g ?i\ by blast - have "m>?i" - using assms(2) i_neq_j_imp_events_neq_inf \f?i \ g?i\ - by (metis Suc_lessI \[[(f 0)(f i)(f ?i)]] \ [[(f 0)(g i)(g ?i)]]\ \f i = g i\ \f m = x\ - \x = g (Suc i)\ assms(1) index_order3 less_nat_zero_code semifin_chain_def) - thus "[[(g 0)(f ?i)(g ?i)]]" - using \[[(f 0)(f?i)(g?i)]] \ [[(f 0)(g?i)(f?i)]]\ \f 0 = g 0\ \f m = x\ \x = g ?i\ - by (metis assms(1) gr_implies_not_zero index_order3 inf_chain_is_long order.asym) - qed - obtain m where "g m = f ?i" - using \(f 0)\X \ g?i\X \ f?i\X\ assms(2) - by (metis ordering_def inf_chain_is_long long_ch_by_ord_def) - hence "[[(g i)(g m)(g ?i)]]" - using abc_acd_bcd \[[(f 0)(f i)(f?i)]] \ [[(f 0)(g i)(g ?i)]]\ \[[(g 0)(f ?i)(g ?i)]]\ - by (metis \f 0 = g 0\ \f i = g i\) - have "[[(g i)(g ?i)(g m)]]" - proof - - have "m>?i" - using \[[(g i)(g m)(g ?i)]]\ assms(2) index_order3 inf_chain_is_long by fastforce - thus ?thesis - using assms(2) inf_chain_is_long long_ch_by_ord_def ordering_ord_ijk by fastforce - qed - thus False - using \[[(g i)(g m)(g ?i)]]\ abc_only_cba by blast - qed - qed - qed - qed - } - moreover have "f 0 = g 0" using inf_chain_origin_unique assms by blast - ultimately show ?thesis using assms by auto -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 ordering - cf Schutz' monograph \cite[ p.~27]{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 "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - 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 "betw4 a b c d" - shows "\betw4 a b d c" "\betw4 a c b d" "\betw4 a c d b" "\betw4 a d b c" "\betw4 a d c b" - "\betw4 b a c d" "\betw4 b a d c" "\betw4 b c a d" "\betw4 b c d a" "\betw4 b d c a" "\betw4 b d a c" - "\betw4 c a b d" "\betw4 c a d b" "\betw4 c b a d" "\betw4 c b d a" "\betw4 c d a b" "\betw4 c d b a" - "\betw4 d a b c" "\betw4 d a c b" "\betw4 d b a c" "\betw4 d b c a" "\betw4 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\\

" "a\P" "b\P" "c\P" "d\P" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - and "\(betw4 a b c d \ betw4 a b d c \ betw4 a c b d \ betw4 a c d b \ betw4 a d b c \ betw4 a d c b)" - shows "betw4 b a c d \ betw4 b a d c \ betw4 b c a d \ betw4 b d a c \ betw4 c a b d \ betw4 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\\

" "a\P" "b\P" "c\P" "d\P" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - and "\(betw4 b a c d \ betw4 b a d c \ betw4 b c a d \ betw4 b d a c \ betw4 c a b d \ betw4 c b a d)" - shows "betw4 a b c d \ betw4 a b d c \ betw4 a c b d \ betw4 a c d b \ betw4 a d b c \ betw4 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 "b\c" - shows "betw4 a b c d \ betw4 a c b d" -proof - - obtain P where "P\\

" "a\P" "b\P" "d\P" - using abc_ex_path abd by blast - have "c\P" - 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; x\I; (x\a \ x\b)\ \ P x I" - shows any_element: "\x I. \x\I; (\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; betw4 a b c d\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; - betw4 b a c d \ betw4 a b d c \ betw4 b a d c \ betw4 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; betw4 a c b d\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; - betw4 b c a d \ betw4 a d b c \ betw4 b d a c \ betw4 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; betw4 a c d b\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; - betw4 a d c b \ betw4 b c d a \ betw4 b d c a \ betw4 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. \I\A; Q I a b\ \ b\A \ a\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; I\A; J\A; betw4 a b c d \ betw4 a c b d \ betw4 a c d b\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; I\A; J\A; - a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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" - "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - have endpoints_on_path: "a\A" "b\A" "c\A" "d\A" - using Q_implies_path asm by blast+ - show "P I J" - proof (cases) (* have to split like this, because the full some_betw4 is too large for Isabelle *) - assume "betw4 b a c d \ betw4 b a d c \ betw4 b c a d \ - betw4 b d a c \ betw4 c a b d \ betw4 c b a d" - then consider "betw4 b a c d"|"betw4 b a d c"|"betw4 b c a d"| - "betw4 b d a c"|"betw4 c a b d"|"betw4 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 "\(betw4 b a c d \ betw4 b a d c \ betw4 b c a d \ - betw4 b d a c \ betw4 c a b d \ betw4 c b a d)" - hence "betw4 a b c d \ betw4 a b d c \ betw4 a c b d \ - betw4 a c d b \ betw4 a d b c \ betw4 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 "betw4 a b c d"|"betw4 a b d c"|"betw4 a c b d"| - "betw4 a c d b"|"betw4 a d b c"|"betw4 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; I\A; J\A; betw4 a b c d \ betw4 a c b d \ betw4 a c d b\ \ P I J" - and "Q I a b" (* Is it better style to have these assumptions first, or last like this? *) - 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. \I\A; Q I a b\ \ b\A \ a\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; I\A; J\A; betw4 a b c d \ betw4 a c b d \ betw4 a c d b\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; I\A; J\A; - a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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 \ b\c \ 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 \ c\d) \ (a=b \ b\c \ c\d \ a\d)\ \ 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; - (a\b \ b\c \ c\d \ a=d) \ (a\b \ b=c \ c\d \ a=d)\ \ P I J" - shows "\I J a b c d. \Q I a b; Q J c d; \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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. \I\A; A\\

; Q I a b\ \ b\A \ a\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 A. \Q I a b; Q J c d; I\A; J\A; 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; I\A; J\A; 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; I\A; J\A; A\\

; - a\b \ b\c \ c\d \ a=d\ \ P I J" -proof - - have last_case: "\I J a b c d A. \Q I a b; Q J c d; I\A; J\A; 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; I\A; J\A; A\\

; - a\b \ b\c \ c\d \ 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. \I\A; Q I a b\ \ b\A \ a\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; 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)" - shows "\I J a b c d. \Q I a b; Q J c d; I\A; J\A; - \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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; I\A; J\A; - [[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; I\A; J\A; - [[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; I\A; J\A; - a\b \ b\c \ c\d \ 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; I\A; J\A; - \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ P I J" - proof - - - text \Fix the sets on the path, and obtain the assumptions of \wlog_endpoints_degenerate1\.\ - fix I J - assume asm1: "I\A" "J\A" - 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 \ b\c \ 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 \ c\d\ \ P I J" - "\ a b c d. \Q I a b; Q J c d; a=b \ b\c \ c\d \ a\d\ \ 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; a\b \ b\c \ c\d \ a=d\ \ P I J" - "\ a b c d. \Q I a b; Q J c d; a\b \ b=c \ c\d \ 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" "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" - 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 metis *) (* metis is proposed by solver "e", but is slow :-( *) - by smt - qed -qed - - -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\ - \ (betw4 a b c d \ P I J) \ (betw4 a c b d \ P I J) \ (betw4 a c d b \ P I J)" - shows "\I J Q a b c d. \I = interval a b; J = interval c d; I\Q; J\Q; Q\\

; - a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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" "I\A" "J\A" "A\\

" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - 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; betw4 a b c d \ betw4 a c b d \ betw4 a c d b\ - \ P I J" - by (meson assms(2)) - show "I = interval a b" "J = interval c d" "I\A" "J\A" "A\\

" - "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - 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; I\Q; J\Q; Q\\

\ - \ ((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)" - shows "\I J a b c d Q. \I = interval a b; J = interval c d; I\Q; J\Q; Q\\

; - \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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" "I\A" "J\A" "A\\

" "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" - 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" "I\A" "J\A" "A\\

" - "\ (a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" - 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" "I1\P" "I2\P" "P\\

" "I1\I2 \ {}" - and events_neq: "a\b" "a\c" "a\d" "b\c" "b\d" "c\d" - shows "is_interval (I1 \ I2)" -proof - - have on_path: "a\P \ b\P \ c\P \ d\P" - using assms(1-4) interval_def by auto - - let ?prop = "\ I J. is_interval (I\J) \ (I\J) = {}" (* 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" (* "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" *) - have "(betw4 a b c d \ ?prop I J)" - "(betw4 a c b d \ ?prop I J)" - "(betw4 a c d b \ ?prop I J)" - proof (rule_tac [!] impI) - assume "betw4 a b c d" - have "I\J = {}" - proof (rule ccontr) - assume "I\J\{}" - then obtain x where "x\I\J" - by blast - show False - proof (cases) - assume "x\a \ x\b \ x\c \ x\d" - hence "[[a x b]]" "[[c x d]]" - using \I=interval a b\ \x\I\J\ \J=interval c d\ \x\I\J\ - 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 "\(x\a \ x\b \ x\c \ x\d)" - 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 "betw4 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 "I\J = interval c b" - proof (safe) - fix x - assume "x \ interval c b" - { - assume "x=b \ x=c" - hence "x\I" - using \betw4 a c b d\ \I = interval a b\ interval_def seg_betw by auto - have "x\J" - using \x=b \ x=c\ - using \betw4 a c b d\ \J = interval c d\ interval_def seg_betw by auto - hence "x\I \ x\J" using \x \ I\ by blast - } moreover { - assume "\(x=b \ x=c)" - hence "[[c x b]]" - using \x\interval c b\ unfolding interval_def segment_def by simp - hence "[[a x b]]" - by (meson \betw4 a c b d\ abc_acd_abd abc_sym) - have "[[c x d]]" - using \betw4 a c b d\ \[[c x b]]\ abc_acd_abd by blast - have "x\I" "x\J" - using \I = interval a b\ \[[a x b]]\ \J = interval c d\ \[[c x d]]\ - interval_def seg_betw by auto - } - ultimately show "x\I" "x\J" by blast+ - next - fix x - assume "x\I" "x\J" - show "x \ interval c b" - proof (cases) - assume not_eq: "x\a \ x\b \ x\c \ x\d" - have "[[a x b]]" "[[c x d]]" - using \x\I\ \I = interval a b\ \x\J\ \J = interval c d\ - not_eq unfolding interval_def segment_def by blast+ - hence "[[c x b]]" - by (meson \betw4 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: "\(x\a \ x\b \ x\c \ x\d)" - { - assume "x=a" - have "\[[d a c]]" - using \betw4 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 \x\J\ \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 "d\interval 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 \x\I\ \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 "betw4 a c d b" - have "I\J = interval c d" - proof (safe) - fix x - assume "x \ interval c d" - { - assume "x\c \ x\d" - 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 "x\I \ x\J" by (simp add: \x \ J\) - } moreover { - assume "\ (x\c \ x\d)" - hence "x\I \ x\J" - by (metis \I = interval a b\ \J = interval c d\ \betw4 a c d b\ \x \ interval c d\ - abc_bcd_abd abc_bcd_acd insertI2 interval_def seg_betw) - } - ultimately show "x\I" "x\J" by blast+ - next - fix x - assume "x\I" "x\J" - 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 (I1\I2)" - 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" "I\J \ {}" "I\P" "J\P" "P\\

" - and events_deg: "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" - shows "is_interval (I \ J)" -proof - - - let ?p = "\ I J. (is_interval (I \ J) \ I\J = {})" - - 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; I\Q; J\Q; Q\\

\ - \ ((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)" - proof - - fix I J a b c d Q - assume "I = interval a b" "J = interval c d" "I\Q" "J\Q" "Q\\

" - 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)" - 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 "I\J = {}" thus ?thesis by simp - next - assume "I\J \ {}" - have "I\J = interval a b" - proof (safe) - fix x assume "x\I" "x\J" - thus "x\interval a b" - using \I = interval a b\ by blast - next - fix x assume "x\interval a b" - show "x\I" - 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 "x\J" - 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 "I\J = {}" thus ?thesis by simp - next - assume "I\J \ {}" - have "I\J = {a}" - proof (safe) - fix x assume "x\I" "x\J" "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 "betw4 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 "a\I" - by (simp add: \I = interval a b\ ends_in_int) - show "a\J" - 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" "I\P" "J\P" "P\\

" "I\J \ {}" - 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 "\x\S. (is_interval x \ x\P)" "P\\

" "\S \ {}" "finite S" "S\{}" - shows "is_interval (\S)" -proof - - obtain n where "n = card S" - by simp - consider "n=0"|"n=1"|"n\2" - 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 "2\n" - obtain m where "m+2=n" - using \2 \ n\ le_add_diff_inverse2 by blast - have ind: "\S. \\x\S. (is_interval x \ x\P); 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}" "I\J" - by (meson card_2_iff) - then have "I\S" "J\S" - by blast+ - then have "is_interval I" "is_interval J" "I\P" "J\P" - by (simp add: "0.prems"(1))+ - also have "I\J \ {}" - using \S={I,J}\ "0.prems"(3) by force - then have "is_interval(I\J)" - 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 "I\S" "S = insert I S'" "I\S'" - 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'\ \I\S'\ by auto - moreover have "\x\S'. 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: "a\Q" - shows "\R L. (is_ray_on R Q \ is_ray_on L Q - \ Q-{a} \ (R \ L) \<^cancel>\events of Q excl. a belong to two rays\ - \ (\r\R. \l\L. [[l a r]]) \<^cancel>\a is betw any 'a of one ray and any 'a of the other\ - \ (\x\R. \y\R. \ [[x a y]]) \<^cancel>\but a is not betw any two events \\ - \ (\x\L. \y\L. \ [[x a y]]))" \<^cancel>\\ of the same ray\ -proof - - text \Schutz here uses Theorem 6, but we don't need it.\ - obtain b where "b\\" and "b\Q" and "b\a" - 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: "\x\Q. [[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: "nonstrict_betw_right4 x a y b \ nonstrict_betw_right4 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" "w\x" "y\z" for x w y z - using path_finsubset_chain order_finite_chain2 (* 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! *) - - 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" "w\x" - by (metis \b \ Q\ \b \ a\ abc_sym event_a mem_Collect_eq path_Q prolong_betw3) - obtain z where "z\?R" "y\z" - 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 "e\ray 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 "e\ray a x" - proof (cases) - assume "e=x" - thus ?thesis - by (simp add: ray_def) - next - assume "e\x" - hence "[[e x a]] \ [[x e a]]" using wxyz_ord - by (meson \[[e a b]]\ \[[x a b]]\ abc_abd_bcdbdc abc_sym) - thus "e\ray 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 -(* - 1. Q \ \

- 2. \x. [[a x b]] \ x \ Q - 3. \x. [[a b x]] \ x \ Q - 4. \x. b \ Q - 5. is_ray {y. [[a y b]] \ [[a b y\} -*) - 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 "e\ray 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 "e\ray 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 "y\e" - hence "[[a e y]] \ [[a y e]]" - using aey_cases by auto - hence "e\ray 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 "e\ray 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 \ c\L \ (\r\R. \l\L-{c}. [[l c r]])" - - -lemma int_on_path: - assumes "l\L" "r\R" "Q\\

" - and partition: "L\Q" "L\{}" "R\Q" "R\{}" "L\R=Q" (*disjoint?*) - shows "interval l r \ Q" -proof - fix x assume "x\interval l r" - thus "x\Q" - 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\\

" "[f[(f 0)..]X]" "X\Q" "closest_bound c X" "is_bound_f b X f" "b\c" - assumes "is_bound_f x X f" - shows "x=b \ x=c \ [[c x b]] \ [[c b x]]" -proof - - have "x\Q" - 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 "x\b" "x\c" - 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\\

" "[f[(f 0)..]X]" "X\Q" "closest_bound_f c X f" "is_bound_f b X f" "b\c" - assumes "x=b \ x=c \ [[c x b]] \ [[c b x]]" - shows "is_bound_f x X f" -proof - - have "x\Q" - 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 "[f[f 0..]X]" - by (simp add: assms(2)) - assume "i [[(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 "[f[f 0..]X]" - by (simp add: assms(2)) - assume "i [[(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\\

" "[f[(f 0)..]X]" "X\Q" "closest_bound_f c X f" "is_bound_f b X f" "b\c" - 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 ray_of_bounds: - assumes "[f[(f 0)..]X]" "closest_bound_f c X f" "is_bound_f b X f" "b\c" - shows "all_bounds X = insert c (ray c b)" - using ray_of_bounds3 assms semifin_chain_on_path by blast - - -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 "c\ray a b" - using abc_abc_neq pro_betw ray_def by auto - have "c\interval 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 - - -lemma bound_any_f: - assumes "Q\\

" "[f[(f 0)..]X]" "X\Q" "is_bound c X" - shows "is_bound_f c X f" -proof - - obtain g where "is_bound_f c X g" "[g[g 0..]X]" - using assms(4) is_bound_def is_bound_f_def by blast - show ?thesis - unfolding is_bound_f_def - proof (safe) - fix i j::nat - show "[f[f 0 ..]X]" by (simp add: assms(2)) - assume "ii < j\ \is_bound_f c X g\ is_bound_f_def by blast - thus "[[(f i)(f j)c]]" - using inf_chain_unique \[g[g 0 ..]X]\ assms(2) by force - qed -qed - - -lemma closest_bound_any_f: - assumes "Q\\

" "[f[(f 0)..]X]" "X\Q" "closest_bound c X" - shows "closest_bound_f c X f" -proof (unfold closest_bound_f_def, safe) - show "is_bound_f c X f" - using bound_any_f assms closest_bound_def is_bound_def by blast -next - fix Q\<^sub>b' - assume "is_bound Q\<^sub>b' X" "Q\<^sub>b' \ c" - then show "[[(f 0) c Q\<^sub>b']]" - by (metis (full_types) assms(2,4) closest_bound_def inf_chain_unique is_bound_f_def) -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: "b\Q" "b\\" - and unreach: "Q\<^sub>x \ \ Q b" "Q\<^sub>z \ \ Q b" "Q\<^sub>x \ Q\<^sub>z" - and xyz: "[[Q\<^sub>x Q\<^sub>y Q\<^sub>z]]" - shows "Q\<^sub>y \ \ Q b" -proof - - - text \First we obtain the chain from I6.\ - have in_Q: "Q\<^sub>x\Q \ Q\<^sub>y\Q \ Q\<^sub>z\Q" - using betw_b_in_path path_Q unreach(1,2,3) unreach_on_path xyz by blast - hence event_y: "Q\<^sub>y\\" - using in_path_event path_Q by blast - obtain X f where X_def: "ch_by_ord f X" "f 0 = Q\<^sub>x" "f (card X - 1) = Q\<^sub>z" - "(\i\{1 .. card X - 1}. (f i) \ \ Q b \ (\Qy\\. [[(f (i - 1)) Qy (f i)]] \ Qy \ \ Q b))" - "short_ch X \ Q\<^sub>x \ X \ Q\<^sub>z \ X \ (\Q\<^sub>y\\. [[Q\<^sub>x Q\<^sub>y Q\<^sub>z]] \ Q\<^sub>y \ \ Q b)" - using I6 [OF assms(1-6)] by blast - hence fin_X: "finite X" - using unreach(3) not_less by fastforce - obtain N where "N=card X" "N\2" - using X_def(2,3) unreach(3) by fastforce - - text \ - Then we have to manually show the bounds, defined via indices only, are in the obtained chain. - This step made me add the two-element-chain-case to I6 in \Minkowski.thy\; - this case is referenced here as \X_def(5)\. -\ - 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 unreach(3) by blast+ - } - hence "[f[Q\<^sub>x..Q\<^sub>z]X]" - unfolding fin_chain_def - by (metis X_def(1-3,5) ch_by_ord_def fin_X fin_long_chain_def get_fin_long_ch_bounds unreach(3)) - - text \ - Further on, we split the proof into two cases, namely the split Schutz absorbs into his - non-strict ordering. Just below is the statement we use \disjE\ with.\ - have y_cases: "Q\<^sub>y\X \ Q\<^sub>y\X" by blast - have y_int: "Q\<^sub>y\interval Q\<^sub>x Q\<^sub>z" - using interval_def seg_betw xyz by auto - have X_in_Q: "X\Q" - using chain_on_path_I6 [where Q=Q and X=X] X_def event_b path_Q unreach - by blast - - show ?thesis - proof (cases) - text \As usual, we treat short chains separately, and they have their own clause in I6.\ - 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 \int_split_to_segs\. \?S\ is defined like in \segmentation\.\ - assume "N\2" - hence "N\3" using \2 \ N\ by auto - have "2\card X" using \2 \ N\ \N = card X\ by blast - show ?thesis using y_cases - proof (rule disjE) - assume "Q\<^sub>y\X" - then obtain i where i_def: "iy = f i" - using X_def(1) - unfolding ch_by_ord_def long_ch_by_ord_def ordering_def - by (metis X_def(5) abc_abc_neq fin_X short_ch_def xyz) - have "i\0 \ i\card 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 "Q\<^sub>y\X" - - let ?S = "if card X = 2 then {segment ?a ?d} else {segment (f i) (f(i+1)) | i. iy\\?S" - proof - - obtain c where "[f[Q\<^sub>x..c..Q\<^sub>z]X]" - using X_def(1) \N = card X\ \N\2\ \[f[Q\<^sub>x..Q\<^sub>z]X]\ fin_chain_def short_ch_card_2 by auto - have "interval Q\<^sub>x Q\<^sub>z = \?S \ X" - using int_split_to_segs [OF \[f[Q\<^sub>x..c..Q\<^sub>z]X]\] by auto - thus ?thesis - using \Q\<^sub>y\X\ y_int by blast - qed - then obtain s where "s\?S" "Q\<^sub>y\s" by blast - - have "\i. i\{1..(card X)-1} \ [[(f(i-1)) Q\<^sub>y (f i)]]" - proof - - obtain i' where i'_def: "i' < N-1" "s = segment (f i') (f (i' + 1))" - using \Q\<^sub>y\s\ \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)) Q\<^sub>y (f(i'+1))]]" - using i'_def(2) \Q\<^sub>y\s\ seg_betw by simp - qed - qed - then obtain i where i_def: "i\{1..(card X)-1}" "[[(f(i-1)) Q\<^sub>y (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 "\x\A. [[a x b]]" "\x\B. [[c x d]]" "A\Q" "B\Q" "Q\\

" - "card A > 1 \ infinite A" "card B > 1 \ infinite B" - shows "\l\Q. \u\Q. \x\A\B. [[l x u]]" -proof - - let ?P = "\ A B. \l\Q. \u\Q. \x\A\B. [[l x u]]" - let ?I = "\ A a b. (card A > 1 \ infinite A) \ (\x\A. [[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 "A\Q" "?I A a b" - show "b\Q\a\Q" - 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 "x\A" "y\A" "x\y" - 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 "x\Q" "y\Q" - using \A \ Q\ \x \ A\ \y \ A\ by auto - have "[[a x b]]" "[[a y b]]" - by (simp add: \(1 < card A \ infinite A) \ (\x\A. [[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 "a\Q \ b\Q" - using \Q\\

\ \x\Q\ \x\y\ \x\Q\ \y\Q\ betw_a_in_path betw_c_in_path by blast - thus ?thesis by simp - qed - qed - - show ?thesis - proof (cases) - assume "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - 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" "A\Q" "B\Q" "Q\\

" - using assms by simp+ - show "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" - using \a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ 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" "I\Q" "J\Q" - and "betw4 a b c d \ betw4 a c b d \ betw4 a c d b" - for I J a b c d - proof - - consider "betw4 a b c d"|"betw4 a c b d"|"betw4 a c d b" - using \betw4 a b c d \ betw4 a c b d \ betw4 a c d b\ by fastforce - thus ?thesis - proof (cases) - assume asm: "betw4 a b c d" show "?P I J" - proof - - have "\x\ I\J. [[a x d]]" - by (metis Un_iff asm betw4_strong betw4_weak that(1) that(2)) - moreover have "a\Q" "d\Q" - using assms(5) on_path that(1-4) by blast+ - ultimately show ?thesis by blast - qed - next - assume "betw4 a c b d" show "?P I J" - proof - - have "\x\ I\J. [[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 "a\Q" "d\Q" - using assms(5) on_path that(1-4) by blast+ - ultimately show ?thesis by blast - qed - next - assume "betw4 a c d b" show "?P I J" - proof - - have "\x\ I\J. [[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 "a\Q" "b\Q" - using assms(5) on_path that(1-4) by blast+ - ultimately show ?thesis by blast - qed - qed - qed - qed - next - assume "\(a\b \ a\c \ a\d \ b\c \ b\d \ c\d)" - - 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" "A\Q" "B\Q" "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 "\l\Q. \u\Q. \x\I \ 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 "\l\Q. \u\Q. \x\I \ 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\ I\J. [[c x d]]" - using abc_abc_neq that(1,2) by fastforce - moreover have "c\Q" "d\Q" - using on_path \a = b \ b = c \ c \ d\ that(1,3) abc_abc_neq by metis+ - ultimately show "\l\Q. \u\Q. \x\I \ J. [[l x u]]" by blast - next - assume "a = b \ b \ c \ c \ d \ a \ d" - hence "\x\ I\J. [[c x d]]" - using abc_abc_neq that(1,2) by fastforce - moreover have "c\Q" "d\Q" - using on_path \a = b \ b \ c \ c \ d \ a \ d\ that(1,3) abc_abc_neq by metis+ - ultimately show "\l\Q. \u\Q. \x\I \ J. [[l x u]]" by blast - next - assume "a \ b \ b = c \ c \ d \ a = d" - hence "\x\ I\J. [[c x d]]" - using abc_sym that(1,2) by auto - moreover have "c\Q" "d\Q" - using on_path \a \ b \ b = c \ c \ d \ a = d\ that(1,3) abc_abc_neq by metis+ - ultimately show "\l\Q. \u\Q. \x\I \ J. [[l x u]]" by blast - next - assume "[[a b c]] \ a = d" - hence "\x\ I\J. [[c x d]]" - by (metis UnE abc_acd_abd abc_sym that(1,2)) - moreover have "c\Q" "d\Q" - using on_path that(2,4) by blast+ - ultimately show "\l\Q. \u\Q. \x\I \ J. [[l x u]]" by blast - next - assume "[[b a c]] \ a = d" - hence "\x\ I\J. [[c x b]]" - using abc_sym abd_bcd_abc betw4_strong that(1,2) by (metis Un_iff) - moreover have "c\Q" "b\Q" - using on_path that by blast+ - ultimately show "\l\Q. \u\Q. \x\I \ J. [[l x u]]" by blast - qed - qed - qed -qed - - -lemma (*for 14i*) union_of_bounded_sets_is_bounded2: - assumes "\x\A. [[a x b]]" "\x\B. [[c x d]]" "A\Q" "B\Q" "Q\\

" - "1 infinite A" "1 infinite B" - shows "\l\Q-(A\B). \u\Q-(A\B). \x\A\B. [[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\\ Q 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: "a\Q" "b\Q" - and reachable: "path_ex a q1" "path_ex b q2" "q1\Q" "q2\Q" (* "\P\\

. \q\Q. path P a q" *)(* "\P\\

. \q\Q. path P b q" *) - shows "\y\Q. \z\Q. (\x\\ Q a. [[y x z]]) \ (\x\\ Q 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 "\q\Q. q\(\ Q a)" "\q\Q. q\(\ Q 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: "\la\Q. \ua\Q. la\\ Q a \ ua\\ Q a \ (\x\\ Q a. [[la x ua]])" - if asm: "a\Q" "path_ex a q" "q\Q" - for a q - proof - - obtain Qy where "Qy\\ Q a" - using asm(2) \a \ Q\ in_path_event path_Q two_in_unreach by blast - then obtain la where "la \ Q - \ Q a" - using asm(2,3) cross_in_reachable by blast - then obtain ua where "ua \ Q - \ Q a" "[[la Qy ua]]" "la \ ua" - using unreachable_set_bounded [where Q=Q and b=a and Qx=la and Qy=Qy] - using \Qy \ \ Q a\ asm in_path_event path_Q by blast - have "la \ \ Q a \ ua \ \ Q a \ (\x\\ Q a. (x\la \ x\ua) \ [[la x ua]])" - proof (intro conjI) - show "la \ \ Q a" - using \la \ Q - \ Q a\ by force - next - show "ua \ \ Q a" - using \ua \ Q - \ Q a\ by force - next show "\x\\ Q a. x \ la \ x \ ua \ [[la x ua]]" - proof (safe) - fix x assume "x\\ Q a" "x\la" "x\ua" - { - assume "x=Qy" hence "[[la x ua]]" by (simp add: \[[la Qy ua]]\) - } moreover { - assume "x\Qy" - have "[[Qy x la]] \ [[la Qy x]]" - proof - - { assume "[[x la Qy]]" - hence "la\\ Q a" - using unreach_connected \Qy\\ Q a\\x\\ Q a\\x\Qy\ in_path_event path_Q that by blast - hence False - using \la \ Q - \ Q 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 \ \ Q a\ \la \ Q - \ Q a\ \x \ \ Q a\ \x \ Qy\ \x \ la\ by force - qed - hence "[[la x ua]]" - (* by (smt DiffD1 DiffD2 \Qy \ \ Q a\ \[[la Qy ua]]\ \ua \ Q - \ Q a\ \x \ \ Q a\ \x \ la\ abc_abd_acdadc abc_acd_abd abc_only_cba abc_sym in_path_event path_Q some_betw that(1) that(2) unreach_connected unreach_on_path) *) - 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 \ \ Q a\ \[[la Qy ua]]\ \ua \ Q - \ Q a\ \x \ \ Q 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 - \ Q a\ \ua \ Q - \ Q a\ by force - qed - - have "\y\Q. \z\Q. (\x\(\ Q a)\(\ Q b). [[y x z]])" - proof - - obtain la ua where "\x\\ Q a. [[la x ua]]" - using events(1) get_bds reachable(1,3) by blast - obtain lb ub where "\x\\ Q b. [[lb x ub]]" - using events(2) get_bds reachable(2,4) by blast - have "\ Q a \ Q" "\ Q b \ Q" - by (simp add: subsetI unreach_on_path)+ - moreover have "1 < card (\ Q a) \ infinite (\ Q 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 (\ Q b) \ infinite (\ Q 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="\ Q a" and B="\ Q b"] - using get_bds assms \\x\\ Q a. [[la x ua]]\ \\x\\ Q b. [[lb x ub]]\ - by blast - qed - - then obtain y z where "y\Q" "z\Q" "(\x\(\ Q a)\(\ Q b). [[y x z]])" - by blast - show ?thesis - proof (rule bexI)+ - show "y\Q" by (simp add: \y \ Q\) - show "z\Q" by (simp add: \z \ Q\) - show "(\x\\ Q a. [[z x y]]) \ (\x\\ Q b. [[z x y]])" - by (simp add: \\x\\ Q a \ \ Q b. [[y x z]]\ abc_sym) - qed -qed - - -theorem (*14*) second_existence_thm_2: - assumes path_Q: "Q\\

" - and events: "a\Q" "b\Q" "c\Q" "d\Q" "c\d" - and reachable: "\P\\

. \q\Q. path P a q" "\P\\

. \q\Q. path P b q" - shows "\e\Q. \ae\\

. \be\\

. path ae a e \ path be b e \ [[c d e]]" -proof - - obtain y z where bounds_yz: "(\x\\ Q a. [[z x y]]) \ (\x\\ Q b. [[z x y]])" - and yz_inQ: "y\Q" "z\Q" - 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\(\ Q a)\(\ Q b)" "z\(\ Q a)\(\ Q b)" - by (meson Un_iff \(\x\\ Q a. [[z x y]]) \ (\x\\ Q b. [[z x y]])\ abc_abc_neq)+ - let ?P = "\e ae be. (e\Q \ path ae a e \ path be b e \ [[c d e]])" - - have exist_ay: "\ay. path ay a y" - if "a\Q" "\P\\

. \q\Q. path P a q" "y\(\ Q a)" "y\Q" - 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\(\ Q a)" "y\(\ Q b)" - using \y \ \ Q a \ \ Q b\ by blast+ - then obtain ay yb where "path ay a y" "path yb b y" - using \y\Q\ 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\(\ Q a)" "z\(\ Q b)" - using \z \ \ Q a \ \ Q b\ by blast+ - then obtain az bz where "path az a z" "path bz b z" - using \z\Q\ 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 "e\Q" - using \[[c d e]]\ betw_c_in_path events(3-5) path_Q by blast - have "e\ \ Q a" "e\ \ Q 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\, \\ Q a\ is empty, so the third conjunct of the conclusion is meaningless. -\ - -theorem (*14*) second_existence_thm_3: - assumes paths: "Q\\

" "R\\

" "Q\R" - and events: "x\Q" "x\R" "a\R" "a\x" "b\Q" - and reachable: "\P\\

. \q\Q. path P b q" - shows "\e\\. \ae\\

. \be\\

. path ae a e \ path be b e \ (\y\\ Q a. [[x y e]])" -proof - - have "a\Q" - using events(1-4) paths eq_paths by blast - hence "\ Q a \ {}" - by (metis events(3) ex_in_conv in_path_event paths(1,2) two_in_unreach) - - then obtain d where "d\ \ Q a" (*as in Schutz*) - by blast - have "x\d" - using \d \ \ Q a\ cross_in_reachable events(1) events(2) events(3) paths(2) by auto - have "d\Q" - using \d \ \ Q a\ unreach_on_path by blast - - have "\e\Q. \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\(\ Q a). [[x y e]]" - proof - fix y assume "y\(\ Q a)" - hence "y\Q" - 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 \d\Q\ \y\Q\ 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 \ \ Q 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 \y\Q\ assms(1) conds empty_iff same_empty_unreach unreach_equiv \y \ \ Q a\) - next - assume "[[y x e]]" - hence "[[y x d]]" - using abd_bcd_abc conds by blast - hence "x\(\ Q a)" - using unreach_connected [where Q=Q and Q\<^sub>x=y and Q\<^sub>y=x and Q\<^sub>z=d and b=a] - using \\[[x y e]]\ \a\Q\ \d\\ Q a\ \y\\ Q 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\(\ Q a)" - using unreach_connected [where Q=Q and Q\<^sub>x=y and Q\<^sub>y=e and Q\<^sub>z=d and b=a] - using \a \ Q\ \d \ \ Q a\ \y \ \ Q 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 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: "Q\P" - and f_def: "[f[a..b..c]Q]" - 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" - by (meson ch_by_ord_def f_def fin_long_chain_def long_ch_card_ge3) - 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)) \ iai\{(0::nat)..?N-2}\ \?N\3\ - 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 ?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 "i\j" - 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 "[[(f i) (f j) (f (j+1))]]" - using asm(2) assms fin_long_chain_def order_finite_chain \?N\3\ - by (metis (no_types, lifting) One_nat_def Suc_diff_Suc Suc_less_eq add.commute - add_leD2 atLeastAtMost_iff card.remove card_Diff_singleton less_Suc_eq_le - less_add_one numeral_2_eq_2 numeral_3_eq_3 plus_1_eq_Suc) - 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 "j\i+1" - have "i < card Q \ j < card Q \ (i+1) < card Q" - using add_mono_thms_linordered_field(3) asm(1,2) assms \?N\3\ by auto - hence "f i \ Q \ f j \ Q \ f (i+1) \ Q" - using f_def unfolding fin_long_chain_def long_ch_by_ord_def ordering_def - by blast - 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 \?N\3\ - 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: "Q\P" - and f_def: "[f[a..b]Q]" (* 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 \ (if card Q=2 then {segment a b} else {segment (f i) (f (i+1)) | i. iS) \ P1 \ P2 \ Q)" - (* The union of these segments and prolongations with the separating points is the path. *) - "card S = (card Q-1) \ (\x\S. is_segment x)" - (* There are N-1 segments. *) - (* There are two prolongations. *) - (* "P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" *) - "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" - (* 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)" "(\x\S. is_segment x)" - "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" - 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})" "P1\P2" "P1\S" "P2\S" - 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 "?N\2" - hence "?N\3" - using \2 \ card Q\ by linarith - then obtain c where "[f[a..c..b]Q]" - using assms ch_by_ord_def fin_chain_def short_ch_card_2 \2 \ card Q\ \card Q \ 2\ - by force - show ?thesis - using number_of_segments [OF assms(1,2) \[f[a..c..b]Q]\] - 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 *) - +(* 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]{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 \<^term>\local_ordering\ is void (respectively, \<^term>\False \ P\ for any \<^term>\P\). + 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 ?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 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 "ii 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 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: "0\i \ i jk. 0 k [f (k-1); f k; f l]" + using thm2_ind2b chX finiteX ordered_nats + by blast + have "j j=l-1" + using ordered_nats by linarith + thus ?thesis + proof + assume "jj < 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: "[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 - + 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 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 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: "[f\X]" + and finiteX: "finite X" + and indices: "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.. {0.. X" + using assms unfolding chain_defs local_ordering_def by auto + } moreover { + fix x assume "x \ X" + then have "\n \ {0..X]" + and finiteX: "finite X" + shows "bij_betw f {0.. {0.. 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..Q|x..z]" + shows "x\z" + 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 "[f\X|a..b..c]" + shows "\n. 0 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<(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: "[f\X|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: "\[f\X|a..c]; card X > 2\ \ \b. [f\X|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: "\[f\X|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]: "[f\Q|x..y..z] \ [f\Q|x..z] \ [x;y;z] \ y\Q" +proof + { + assume "[f\Q|x .. z] \ [x;y;z] \ y\Q" + thus "[f\Q|x..y..z]" + using abc_abc_neq finite_long_chain_with_def by blast + } { + assume asm: "[f\Q|x..y..z]" + show "[f\Q|x .. z] \ [x;y;z] \ y\Q" + using asm fin_ch_betw finite_long_chain_with_def by blast + } +qed + + +lemma finite_long_chain_with_card: "[f\Q|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] \ y\Q" + shows "[f\Q|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" "y\x \ y\z \ y\Q" + shows "[f\Q|x..y..z]" + using assms finite_chain_alt finite_chain_with_def finite_long_chain_with_def by auto + + +lemma chain_sym_obtain: + assumes "[f\X|a..b..c]" + obtains g where "[g\X|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 "[f\X|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 "[f\X|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 "nn. if n = 0 then f 1 else f 0) = (\n. f (card X - Suc n))" if "nn. 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. [f\X|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 "[f\X|a..c]" + obtains g where "[g\X|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 "\f\de\ab. [a;f;b]" +proof - + have "\f\ab \ de. \X ord. [ord\X|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 \ (\f\de\ab. [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 "\f\ab \ de. \X g. [g\X|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 "(\f\de\(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. [ord\X|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\\

. \a\Q. 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 "[f\X]" "x\X" "finite X" + obtains i where "f i = x" "iiiiX]" "x\X" "infinite X" + obtains i where "f i = x" + using assms unfolding chain_defs local_ordering_def by blast + + +lemma fin_chain_on_path2: + assumes "[f\X]" "finite X" + shows "\P\\

. X\P" + 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 "x\X" + then obtain i where i: "f i = x" "ii=1"|"i>1" by linarith + hence "x\P" + 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 "[f\X]" "finite X" + shows "\!P\\

. X\P" +proof - + obtain P where P: "P\\

" "X\P" + using fin_chain_on_path2[OF assms] by auto + obtain a b where ab: "a\X" "b\X" "a\b" + 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 "[f\X]" "finite X" "a\X" "b\X" "a\b" + shows "X \ path_of a b" +proof - + let ?ab = "path_of a b" + obtain P where P: "P\\

" "X\P" 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 "a\Q \ a\Q" by auto + thus ?thesis + proof + assume "a\Q" + thus ?thesis by blast + next + assume "a\Q" + then obtain b where "b\unreach-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 "\x\unreach-on Q from b. \y\unreach-on Q from b. x \ y" + shows "\x\Q. \y\Q. 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: "b\Q" "b\\" + and unreach: "Q\<^sub>x \ unreach-on Q from b" "Q\<^sub>z \ unreach-on Q from b" "Q\<^sub>x \ Q\<^sub>z" + and X_def: "[f\X|Q\<^sub>x..Q\<^sub>z]" + "(\i\{1 .. card X - 1}. (f i) \ unreach-on Q from b \ (\Q\<^sub>y\\. [(f(i-1)); Q\<^sub>y; f i] \ Q\<^sub>y \ unreach-on Q from b))" + shows "X\Q" +proof - + have 1: "path Q Q\<^sub>x Q\<^sub>z" using unreachable_subset_def unreach path_Q by simp + then have 2: "Q = path_of Q\<^sub>x Q\<^sub>z" using path_of_ex[of Q\<^sub>x Q\<^sub>z] by (meson eq_paths) + have "X\path_of Q\<^sub>x Q\<^sub>z" + proof (rule fin_chain_on_path3[of f]) + from unreach(3) show "Q\<^sub>x \ Q\<^sub>z" by simp + from X_def chain_defs show "[f\X]" "finite X" by metis+ + from assms(7) points_in_chain show "Q\<^sub>x \ X" "Q\<^sub>z \ 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 \ R\Q\ \ a \ b" +by blast + +lemma nocross_events_neq: + "\Q \ \

; R \ \

; a \ Q; b \ R; R\Q = {}\ \ 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 "\Qz\Q - 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 "R\Q" + using R_neq_Q path_R path_sub_events by auto + obtain c where "c\R" "c\Q" + 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 "a\Q" "b\Q" + 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 "\b\Q. 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 "\a\Q. \b\Q. 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 "\c\Q. [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 "\c\Q. \d\Q. [a;b;c] \ [a;b;d] \ c\d" + 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 "\a\X. \b\X. a \ b \ (\c\X. 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: "\c\Y. (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: "Q\X" + obtains a b where "a\b" and "a\Q" and "b\Q" and "\c\Q. (a\c \ b\c) \ [a;c;b]" +proof - + obtain n where "n\0" and "card Q = n+3" + using card_Q nat_le_iff_add + by auto + then obtain a b where "a\b" and "a\Q" and "b\Q" and "\c\Q. (a\c \ b\c) \ [a;c;b]" + using finite_path_has_ends assms \n\0\ + 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 "n\1" + using \card Q \ 0\ by linarith + then consider (n1) "n=1" | (n2) "n=2" | (n3) "n\3" + 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 "a\b" and "a\Q" and "b\Q" + using ge2_events_lax assms(1) by blast + then obtain c where "c\Q" and "c\a" and "c\b" + 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 "Q\Q" by simp + then obtain a b where "a\Q" and "b\Q" and "a\b" + and acb: "\c\Q. (c\a \ c\b) \ [a;c;b]" + using obtain_fin_path_ends card_Q fin_Q assms(1) + by metis + then obtain x where "[a;b;x]" and "x\Q" + 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 "dc\ce" + 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 "fe\ea" + 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 "\x\a'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 \ e\R \ d'\e)" + by blast + hence "\(\R\\

. e\R \ d'\R)" + using d'_def(3) by blast + moreover have "ab\\

\ e\\ \ e\ab" + 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 "e\ab" + using S_neq_ab a_inS e_inS e_neq_a eq_paths path_S path_ab by auto + have "ab\\

\ e\ab" + 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 \e\ab\ + 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" "b\d'" "e\d'" + using unreachable_set_bounded \b \ ab - unreach-on ab from e\ \d \ unreach-on ab from e\ e_event \e\ab\ path_ab + by (metis DiffE) + then obtain d'e where "path d'e d' e" + using unreachable_bounded_path_only e_event \e\ab\ 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] \ d\ab" + 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 "c\d" 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]" "f\c'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\\ \ b\c'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'" "b\f'" + 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 "f\c'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 \f\c'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\\ \ b\c'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'" "b\f'" + 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 "\e\S. 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: "f\c'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: "b\c" + shows "[a;b;c] \ [a;c;b]" +proof - + obtain P where P_def: "P\\

" "a\P" "b\P" "d\P" + using abd abc_ex_path by blast + hence "c\P" + 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 + \<^term>\local_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 "[f\X|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. [f\X|x..]" + using assms infinite_chain_with_def chain_alts by auto + +lemma endpoint_in_semifin: + assumes "[f\X|x..]" + shows "x\X" + 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 "[f\X]" "finite X" + and x: "x\X" and y: "y\X" and z:"z\X" and xy: "x\y" and xz: "x\z" and yz: "y\z" + 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 j j k ij \ 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 "x\X" and "y\X" and "x\y" + and "z\X" and "z\x" + 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 "\x\X. \y\X. \z\X. [f\X|x..y..z]" +proof (rule bexI)+ + show 1:"[f\X|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 n\<^sub>x n\<^sub>y n\<^sub>z + where "x\X" "y\X" "z\X" "[f\X|x..y..z]" "f n\<^sub>x = x" "f n\<^sub>y = y" "f n\<^sub>z = 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 "[f\X|a..c]" and "b\X" + 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 "[f\X|a..c]" "[g\X|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 "[f\X|a..c]" + and "j0 \ f j \ a" "j 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 j k jX|a..c]" + and "f j = b" "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 "b\X" + 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<(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 \b\X\ 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 \b\X\ 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: "[f\X|a..b..c]" "f i = b" "card X > i" + shows "bij_betw f {0<..X. [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.. {0.. X. [a;e;b]}" + proof + show "f ` {0<.. {e \ X. [a;e;b]}" + proof (auto simp add: image_subset_iff conjI) + fix j assume asm: "j>0" "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<.. X" "[a;e;b]" + obtain j where "f j = e" "j f ` {0<..j ij<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{0<..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" "a\A" "b\B" + 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 "a\X" shows "insert a {x\X. P x} = {x\X. P x \ x=a}" + using insert_iff assms by blast + + +lemma index_bij_betw_subset2: + assumes chX: "[f\X|a..b..c]" "f i = b" "card X > i" + shows "bij_betw f {0..i} {e\X. [a;e;b]\a=e\b=e}" +proof - + have "bij_betw f {0<..X. [a;e;b]}" using index_bij_betw_subset[OF assms] . + moreover have "0\{0<..{0<..{e\X. [a;e;b]}" "b\{e\X. [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 {e\X. [a;e;b]})) = {e\X. [a;e;b]\a=e\b=e}" + proof - + have 1: "(insert a {e\X. [a;e;b]}) = {e\X. [a;e;b]\a=e}" + using insert_iff2[OF points_in_long_chain(1)[OF chX(1)]] by auto + have "b\{e\X. [a;e;b]\a=e}" + using abc_abc_neq chX(1) fin_ch_betw by fastforce + thus "(insert b (insert a {e\X. [a;e;b]})) = {e\X. [a;e;b]\a=e\b=e}" + using 1 insert_iff2 points_in_long_chain(2)[OF chX(1)] by auto + qed + moreover have "(insert i (insert 0 {0<..X|a..b..c]" + shows "[f \ {e\X. [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 = "{e\X. [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 "a\b" "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: "iThe path \P\ on which \X\ lies. If \?Y\ has two arguments, \P\ makes it a short chain.\ + obtain P where P_def: "P\\

" "X\P" "\Q. Q\\

\ X\Q \ Q=P" + using fin_chain_on_path[of f X] assms unfolding chain_defs by force + have "a\P" "b\P" 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: "3\card ?Y" using gt_1 cardY i_eq by linarith + { + fix n assume "n < card ?Y" + hence "ni 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 assms chX chain_defs local_ordering_def by metis + consider "0 nn 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 "x\X" by simp + obtain i\<^sub>x where i\<^sub>x: "i\<^sub>x < card X" "f i\<^sub>x = x" + using assms obtain_index_fin_chain chain_defs \x\X\ by metis + have "i\<^sub>x < card ?Y" + proof - + consider "[a;x;b]"|"x=a"|"x=b" using \x\?Y\ by auto + hence "(i\<^sub>x i\<^sub>x<0) \ i\<^sub>x=0 \ i\<^sub>x=i" + apply cases + apply (metis \f 0=a\ chX finX i i\<^sub>x less_nat_zero_code neq0_conv order_finite_chain_indices) + using \f 0 = a\ chX finX i\<^sub>x index_injective apply blast + by (metis chX finX i(2) i\<^sub>x index_injective linorder_neqE_nat) + thus ?thesis using gt_1 i_eq by linarith + qed + hence "\n. n < card ?Y \ f n = x" using i\<^sub>x(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 "[f\X|a..f i..c]" "j\i" "j j = card X - 1 \ j = i" +proof - + consider (inter)"j>i \ jX|(f 0) ..]" "i\0" "j>i" "Y=f`{i..j}" + shows "Y\X" +proof + fix x assume "x\Y" + 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 "x\X" + using \f n = x\ by blast +qed + + +lemma i_le_j_events_neq: + assumes "[f\X|a..b..c]" + and "i f j" + using chain_defs by (meson assms index_injective2) + +lemma indices_neq_imp_events_neq: + assumes "[f\X|a..b..c]" + and "i\j" "j f j" + by (metis assms i_le_j_events_neq less_linear) + +end (* context MinkowskiChain *) + +context MinkowskiSpacetime begin + +lemma bound_on_path: + assumes "Q\\

" "[f\X|(f 0)..]" "X\Q" "is_bound_f b X f" + shows "b\Q" +proof - + obtain a c where "a\X" "c\X" "[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 "x\segment a b \ x\segment b c" for x + proof + fix x assume "x\segment 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 "x\segment 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}. c\a \ c\b)" + by blast + have "a\b \ a\Q & b\Q" + 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: "a\b \ a\c \ b\c" + 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: "a\Q \ b\Q \ c\Q" + 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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and bY: "[b; a\<^sub>1; a\<^sub>n]" + fixes g defines g_def: "g \ (\j::nat. if j\1 then f (j-1) else b)" + shows "[g\(insert b Y)|b .. a\<^sub>1 .. a\<^sub>n]" +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 "b\Y" + using abc_ac_neq abc_only_cba(1) assms by (metis fin_ch_betw2 finite_long_chain_with_alt) + have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" + 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 "[a\<^sub>1; f 1; a\<^sub>n]" + 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 \<^term>\local_ordering\ lemma from our script for 3.6, which Schutz doesn't list.\ + hence "[b; a\<^sub>1; f 1]" + using bY abd_bcd_abc by blast + have "local_ordering g betw ?X" + proof - + { + fix n assume "finite ?X \ n ?X" + apply (cases "n\1") + prefer 2 apply (simp add: g_def) + proof + assume "1\n" "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 \b\Y\ 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\ \b\Y\ 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 "n\1") + using \b\Y\ \[b; a\<^sub>1; 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;a\<^sub>1;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: \b\Y\ \x = b\ g_def) + } moreover { + fix x assume "x\?X" "x\b" + hence "\n. (finite ?X \ n < card ?X) \ g n = x" + proof - + obtain n where "f n = x" "n < card Y" + using \x\?X\ \x\b\ 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: \b\Y\ \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 \b\Y\ + 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;a\<^sub>1;a\<^sub>n] \ a\<^sub>1 \ ?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 \b\Y\ eval_nat_numeral by (metis card.infinite finite.insertI not_numeral_le_zero) + show "g (card ?X - 1) = a\<^sub>n" + using g_def \b\Y\ 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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and Yb: "[a\<^sub>1; a\<^sub>n; b]" + fixes g defines g_def: "g \ (\j::nat. if j \ (card Y - 1) then f j else b)" + shows "[g\(insert b Y)|a\<^sub>1 .. a\<^sub>n .. b]" +proof - + let ?X = "insert b Y" + have "b\Y" + 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 "a\<^sub>1\Y \ a\<^sub>n\Y \ a\Y" + using long_ch_Y points_in_long_chain by meson + have "a\<^sub>1\a \ a\ a\<^sub>n \ a\<^sub>1\a\<^sub>n" + using Yb abc_abc_neq finite_long_chain_with_def long_ch_Y by auto + have "Suc (card Y) = card ?X" + using \b\Y\ fin_X finite_long_chain_with_def long_ch_Y by auto + obtain f2 where f2_def: "[f2\Y|a\<^sub>n..a..a\<^sub>1]" "f2=(\n. f (card Y - 1 - n))" + using chain_sym long_ch_Y by blast + obtain g2 where g2_def: "g2 = (\j::nat. if j\1 then f2 (j-1) else b)" + by simp + have "[b; a\<^sub>n; a\<^sub>1]" + using abc_sym Yb by blast + hence g2_ord_X: "[g2\?X|b .. a\<^sub>n .. a\<^sub>1]" + using chain_append_at_left_edge [where a\<^sub>1="a\<^sub>n" and a\<^sub>n="a\<^sub>1" and f="f2"] + fin_X \b\Y\ f2_def g2_def + by blast + then obtain g1 where g1_def: "[g1\?X|a\<^sub>1..a\<^sub>n..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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and S_def: "S = {k::nat. [a\<^sub>1; f k; b] \ k < card Y}" + and k_def: "S\{}" "k = Max S" + and k'_def: "k'>0" "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 "[f\Y]" 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 "k\S" using k_def Max_in S_def by (metis finite_Collect_conjI finite_Collect_less_nat) + hence "kk\S\ by auto + show "k' \ S" + proof (rule ccontr) + assume asm: "\k'\S" + have 1: "[f 0;f k;f k']" + proof - + have "[a\<^sub>1; 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 "[a\<^sub>1; f k; b]" + using S_def \k \ S\ by blast + have "[f k; b; f k']" + using abc_acd_bcd \[a\<^sub>1; b; f k']\ \[a\<^sub>1; f k; b]\ by blast + thus ?thesis + using \[a\<^sub>1;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 \[f\Y]\ \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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and Y_def: "b\Y" + and Yb: "[a\<^sub>1; b; a\<^sub>n]" + shows "\k>0. [a\<^sub>1; b; f k] \ k < card Y \ \(\k'1; b; f k'])" +proof - +(* the usual suspects first, they'll come in useful I'm sure *) + have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" + 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 \a\<^sub>1\ and \b\, and find the maximal one.\ + let ?S = "{k::nat. [a\<^sub>1; 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 \ [a\<^sub>1; b; f 1] \ 1 < card Y \ \ (\k'::nat. k' < 1 \ [a\<^sub>1; 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 \ [a\<^sub>1; b; f k'])" + using abc_abc_neq bound_indices + by blast + show "[a\<^sub>1; 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 "[a\<^sub>1; f 1; a\<^sub>n]" + 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 "[a\<^sub>1; b; f 1] \ [a\<^sub>1; f 1; b] \ [b; a\<^sub>1; 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 "[a\<^sub>1; b; f 1]" + proof - + have "\n. \ ([a\<^sub>1; f n; b] \ n < card Y)" + using S_def \S = {}\ + by blast + then have "[a\<^sub>1; b; f 1] \ \ [a\<^sub>n; f 1; b] \ \ [a\<^sub>1; 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 \[a\<^sub>1; b; f 1] \ [a\<^sub>1; f 1; b] \ [b; a\<^sub>1; f 1]\ \[a\<^sub>1; f 1; a\<^sub>n]\) + 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 "k\1" + proof (rule ccontr) + assume "\ 1 \ k" + hence "k=0" by simp + have "[a\<^sub>1; f k; b]" + using \k\S\ 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 [a\<^sub>1; b; f ?k] \ ?k < card Y \ \ (\k'::nat. k' < ?k \ [a\<^sub>1; b; f k'])" + proof (intro conjI) + show "(0::nat)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 "[a\<^sub>1; 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 "[a\<^sub>1; f ?k; a\<^sub>n] \ f ?k = a\<^sub>n" + 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 "[a\<^sub>1; b; f ?k]" + proof (rule disjE) + assume "[a\<^sub>1; f ?k; a\<^sub>n]" + hence "f ?k \ a\<^sub>n" + by (simp add: abc_abc_neq) + hence "[a\<^sub>1; b; f ?k] \ [a\<^sub>1; f ?k; b] \ [b; a\<^sub>1; f ?k]" + using abc_ex_path_unique some_betw abc_sym \[a\<^sub>1; f ?k; a\<^sub>n]\ + \f ?k \ Y\ Yb abc_abc_neq assms(3) cross_once_notin + by (smt Y_def) + moreover have "\ [a\<^sub>1; f ?k; b]" + proof + assume "[a\<^sub>1; f ?k; b]" + hence "?k \ S" + using S_def \[a\<^sub>1; 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; a\<^sub>1; f ?k]" + using Yb \[a\<^sub>1; f ?k; a\<^sub>n]\ abc_only_cba + by blast + ultimately show "[a\<^sub>1; b; f ?k]" + by blast + next assume "f ?k = a\<^sub>n" + show ?thesis + using Yb \f (k + 1) = a\<^sub>n\ by blast + qed + qed + show "\(\k'::nat. k' < k + 1 \ [a\<^sub>1; b; f k'])" + proof + assume "\k'::nat. k' < k + 1 \ [a\<^sub>1; b; f k']" + then obtain k' where k'_def: "k'>0" "k' < k + 1" "[a\<^sub>1; b; f k']" + using abc_ac_neq bound_indices neq0_conv + by blast + hence "k'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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and Y_def: "b\Y" + and Yb: "[a\<^sub>1; b; a\<^sub>n]" + shows "\k. [f k; b; a\<^sub>n] \ k < card Y - 1 \ \(\k'k \ [f k'; b; a\<^sub>n])" +proof - + have bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" + 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 \a\<^sub>1\ and \b\.\ + let ?S = "{k::nat. [a\<^sub>n; 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; a\<^sub>n] \ ?n < card Y - 1 \ \(\k'?n \ [f k'; b; a\<^sub>n])" + proof (intro conjI) + show "?n < card Y - 1" + using Yb abc_ac_neq bound_indices not_le by fastforce + next show "\(\k'?n \ [f k'; b; a\<^sub>n])" + 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; a\<^sub>n]" + 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 "[a\<^sub>1; f ?n; a\<^sub>n]" + 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 "[a\<^sub>n; b; f ?n] \ [a\<^sub>n; f ?n; b] \ [b; a\<^sub>n; f ?n]" + using abc_ex_path_unique some_betw abc_sym \[a\<^sub>1; f ?n; a\<^sub>n]\ + by (smt Y_def Yb \f ?n \ Y\ abc_abc_neq cross_once_notin) + thus "[f ?n; b; a\<^sub>n]" + proof - + have "\n. \ ([a\<^sub>n; f n; b] \ n < card Y)" + using S_def \S = {}\ + by blast + then have "[a\<^sub>n; b; f ?n] \ \ [a\<^sub>1; f ?n; b] \ \ [a\<^sub>n; 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 \[a\<^sub>n; b; f ?n] \ [a\<^sub>n; f ?n; b] \ [b; a\<^sub>n; f ?n]\ \[a\<^sub>1; f ?n; a\<^sub>n]\) + 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; a\<^sub>n] \ ?k < card Y - 1 \ \ (\k' [f k'; b; a\<^sub>n])" + 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; a\<^sub>n]" + 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 "[a\<^sub>1; f ?k; a\<^sub>n] \ f ?k = a\<^sub>1" + 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; a\<^sub>n]" + proof (rule disjE) + assume "[a\<^sub>1; f ?k; a\<^sub>n]" + hence "f ?k \ a\<^sub>1" + using abc_abc_neq by blast + hence "[a\<^sub>n; b; f ?k] \ [a\<^sub>n; f ?k; b] \ [b; a\<^sub>n; f ?k]" + using abc_ex_path_unique some_betw abc_sym \[a\<^sub>1; f ?k; a\<^sub>n]\ + \f ?k \ Y\ Yb abc_abc_neq assms(3) cross_once_notin + by (smt Y_def) + moreover have "\ [a\<^sub>n; f ?k; b]" + proof + assume "[a\<^sub>n; f ?k; b]" + hence "?k \ S" + using S_def \[a\<^sub>n; 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) \ a\<^sub>1\ chain_defs long_ch_Y + by auto + qed + moreover have "\ [b; a\<^sub>n; f ?k]" + using Yb \[a\<^sub>1; f ?k; a\<^sub>n]\ abc_only_cba(2) abc_bcd_acd + by blast + ultimately show "[f ?k; b; a\<^sub>n]" + using abc_sym by auto + next assume "f ?k = a\<^sub>1" + show ?thesis + using Yb \f (k - 1) = a\<^sub>1\ by blast + qed + qed + show "\(\k' [f k'; b; a\<^sub>n])" + proof + assume "\k' [f k'; b; a\<^sub>n]" + then obtain k' where k'_def: "k' k - 1" "[a\<^sub>n; 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 a\<^sub>1 a a\<^sub>n _ 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: "[f\Y|a\<^sub>0..a..a\<^sub>n]" + and x_def: "x\Y" "[a\<^sub>0; x; a\<^sub>n]" + obtains n\<^sub>b n\<^sub>c b c + where "b=f n\<^sub>b" "c=f n\<^sub>c" "[b;x;c]" "b\Y" "c\Y" "n\<^sub>b = n\<^sub>c - 1" "n\<^sub>cc>0" + "\(\k < card Y. [f k; x; a\<^sub>n] \ k>n\<^sub>b)" "\(\kc. [a\<^sub>0; x; f k])" +proof - + have "\ n\<^sub>b n\<^sub>c b c. b=f n\<^sub>b \ c=f n\<^sub>c \ [b;x;c] \ b\Y \ c\Y \ n\<^sub>b = n\<^sub>c - 1 \ n\<^sub>c n\<^sub>c>0 + \ \(\k < card Y. [f k; x; a\<^sub>n] \ k>n\<^sub>b) \ \(\k < n\<^sub>c. [a\<^sub>0; x; f k])" + proof - + have bound_indices: "f 0 = a\<^sub>0 \ f (card Y - 1) = a\<^sub>n" + 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\\

" "Y\P" + using fin_chain_on_path long_ch_Y fin_Y chain_defs by meson + hence "x\P" + using betw_b_in_path x_def(2) long_ch_Y points_in_long_chain + by (metis abc_abc_neq in_mono) + obtain n\<^sub>c where nc_def: "\(\k. [a\<^sub>0; x; f k] \ kc)" "[a\<^sub>0; x; f n\<^sub>c]" "n\<^sub>cc>0" + using smallest_k_ex [where a\<^sub>1=a\<^sub>0 and a=a and a\<^sub>n=a\<^sub>n and b=x and f=f and Y=Y] + long_ch_Y x_def + by blast + then obtain c where c_def: "c=f n\<^sub>c" "c\Y" + using chain_defs local_ordering_def by (metis chY2) + have c_goal: "c=f n\<^sub>c \ c\Y \ n\<^sub>c n\<^sub>c>0 \ \(\k < card Y. [a\<^sub>0; x; f k] \ kc)" + using c_def nc_def(1,3,4) by blast + obtain n\<^sub>b where nb_def: "\(\k < card Y. [f k; x; a\<^sub>n] \ k>n\<^sub>b)" "[f n\<^sub>b; x; a\<^sub>n]" "n\<^sub>b1=a\<^sub>0 and a=a and a\<^sub>n=a\<^sub>n and b=x and f=f and Y=Y] + long_ch_Y x_def + by blast + hence "n\<^sub>bb" "b\Y" + using nb_def chY2 local_ordering_def by (metis local_long_ch_by_ord_alt) + have "[b;x;c]" + proof - + have "[b; x; a\<^sub>n]" + using b_def(1) nb_def(2) by blast + have "[a\<^sub>0; x; c]" + using c_def(1) nc_def(2) by blast + moreover have "\a. [a;x;b] \ \ [a; a\<^sub>n; x]" + using \[b; x; a\<^sub>n]\ abc_bcd_acd + by (metis (full_types) abc_sym) + moreover have "\a. [a;x;b] \ \ [a\<^sub>n; a; x]" + using \[b; x; a\<^sub>n]\ by (meson abc_acd_bcd abc_sym) + moreover have "a\<^sub>n = c \ [b;x;c]" + using \[b; x; a\<^sub>n]\ by meson + ultimately show ?thesis + using abc_abd_bcdbdc abc_sym x_def(2) + by meson + qed + have "n\<^sub>bc" + using \[b;x;c]\ \n\<^sub>c \n\<^sub>b \c = f n\<^sub>c\ \b = f n\<^sub>b\ + 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 "n\<^sub>b = n\<^sub>c - 1" + proof (rule ccontr) + assume "n\<^sub>b \ n\<^sub>c - 1" + have "n\<^sub>bc-1" + using \n\<^sub>b \ n\<^sub>c - 1\ \n\<^sub>bc\ by linarith + hence "[f n\<^sub>b; (f(n\<^sub>c-1)); f n\<^sub>c]" + using \n\<^sub>b \ n\<^sub>c - 1\ long_ch_Y nc_def(3) order_finite_chain chain_defs + by auto + have "\[a\<^sub>0; x; (f(n\<^sub>c-1))]" + using nc_def(1,4) diff_less less_numeral_extra(1) + by blast + have "n\<^sub>c-1\0" + using \n\<^sub>b < n\<^sub>c\ \n\<^sub>b \ n\<^sub>c - 1\ by linarith + hence "f(n\<^sub>c-1)\a\<^sub>0 \ a\<^sub>0\x" + using bound_indices \n\<^sub>b < n\<^sub>c - 1\ abc_abc_neq less_imp_diff_less nb_def(1) nc_def(3) x_def(2) + by blast + have "x\f(n\<^sub>c-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 "[a\<^sub>0; f (n\<^sub>c-1); x]" + using long_ch_Y nc_def c_def chain_defs + by (metis \[f n\<^sub>b;f (n\<^sub>c - 1);f n\<^sub>c]\ \\ [a\<^sub>0;x;f (n\<^sub>c - 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(n\<^sub>c-1)); x; a\<^sub>n]" + using abc_acd_bcd x_def(2) by blast + thus False using nb_def(1) + using \n\<^sub>b < n\<^sub>c - 1\ less_imp_diff_less nc_def(3) + by blast + qed + have b_goal: "b=f n\<^sub>b \ b\Y \ n\<^sub>b=n\<^sub>c-1 \ \(\k < card Y. [f k; x; a\<^sub>n] \ k>n\<^sub>b)" + using b_def nb_def(1) nb_def(3) \n\<^sub>b=n\<^sub>c-1\ by blast + thus ?thesis + using \[b;x;c]\ c_goal + using \n\<^sub>b < 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: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + and Y_def: "b\Y" + and Yb: "[a\<^sub>1; b; a\<^sub>n]" + and k_def: "[a\<^sub>1; b; f k]" "k < card Y" "\(\k'. (0::nat) k' [a\<^sub>1; b; f k'])" + fixes g + defines g_def: "g \ (\j::nat. if (j\k-1) then f j else (if (j=k) then b else f (j-1)))" + shows "[g\insert b Y|a\<^sub>1 .. b .. a\<^sub>n]" +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 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" + 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 \a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n\ + 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 \[a\<^sub>1; b; f k]\ \k = 1\ bound_indices by auto + next assume "k\1" show "[f(k-1); b; f k]" + proof - + have a1k: "[a\<^sub>1; 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 {a\<^sub>1, (f (k-1)), b, (f k)}" using chain4 + using k_def(1) abc_ex_path_unique between_chain cross_once_notin + by (smt \[a\<^sub>1; 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 \[a\<^sub>1; 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 "\ [a\<^sub>1; f k; b]" + by (simp add: \[a\<^sub>1; 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 \[a\<^sub>1; 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 "n\k-1" | "n\k+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 "x\Y" + 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 "ix\k-1 \ ix\k" + by linarith + thus ?thesis + proof + assume "ix\k-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 "ix\k" + 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 "x\Y" + 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 n\k-1 \ Suc n=k" if "n\k-1" + using \k \ 0\ that by linarith + have cases_ssn: "Suc(Suc n)\k-1 \ Suc(Suc n)=k" if "n\k-1" "Suc n\k-1" + using that(2) by linarith + + consider "n\k-1" | "n\k+1" | "n=k" + by linarith + then show "[g n; g (Suc n); g (Suc (Suc n))]" + proof (cases) + assume "n\k-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 "n\k+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 nk + 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|a\<^sub>1..b..a\<^sub>n]" + using fin_X \a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n\ 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 a\<^sub>1 a a\<^sub>n where long_ch_Y: "[f\Y|a\<^sub>1..a..a\<^sub>n]" + using \4 \ card Y\ get_fin_long_ch_bounds by fastforce + hence bound_indices: "f 0 = a\<^sub>1 \ f (card Y - 1) = a\<^sub>n" + by (simp add: chain_defs) + have "a\<^sub>1 \ a\<^sub>n \ a\<^sub>1 \ b \ b \ a\<^sub>n" + using \b \ Y\ abc_abc_neq fin_ch_betw long_ch_Y points_in_long_chain by metis + moreover have "a\<^sub>1 \ Q \ a\<^sub>n \ Q \ b \ Q" + using IH.prems(3) X_eq long_ch_Y points_in_long_chain by auto + ultimately consider "[b; a\<^sub>1; a\<^sub>n]" | "[a\<^sub>1; a\<^sub>n; b]" | "[a\<^sub>n; b; a\<^sub>1]" + using some_betw [of Q b a\<^sub>1 a\<^sub>n] \Q \ \

\ by blast + thus "ch X" + proof (cases) + (* case (i) *) + assume "[b; a\<^sub>1; a\<^sub>n]" + have X_eq': "X = Y \ {b}" + using X_eq by auto + let ?g = "\j. if j \ 1 then f (j - 1) else b" + have "[?g\X|b..a\<^sub>1..a\<^sub>n]" + using chain_append_at_left_edge IH.prems(4) X_eq' \[b; a\<^sub>1; a\<^sub>n]\ \b \ Y\ long_ch_Y X_eq + by presburger + thus "ch X" + using chain_defs by auto + next + (* case (ii) *) + assume "[a\<^sub>1; a\<^sub>n; b]" + let ?g = "\j. if j \ (card X - 2) then f j else b" + have "[?g\X|a\<^sub>1..a\<^sub>n..b]" + using chain_append_at_right_edge IH.prems(4) X_eq \[a\<^sub>1; a\<^sub>n; b]\ \b \ Y\ long_ch_Y + by auto + thus "ch X" using chain_defs by (meson ch_def) + next + (* case (iii) *) + assume "[a\<^sub>n; b; a\<^sub>1]" + then have "[a\<^sub>1; b; a\<^sub>n]" + by (simp add: abc_sym) + obtain k where + k_def: "[a\<^sub>1; b; f k]" "k < card Y" "\ (\k'. 0 < k' \ k' < k \ [a\<^sub>1; b; f k'])" + using \[a\<^sub>1; b; a\<^sub>n]\ \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 "[g\X|a\<^sub>1..b..a\<^sub>n]" + using chain_append_inside [of f Y a\<^sub>1 a a\<^sub>n b k] IH.prems(4) X_eq + \[a\<^sub>1; b; a\<^sub>n]\ \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 "[f\X|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: "[f\X|a..b]" "a\X \ b\X" + 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" "Q\P" "N=2" + and f_def: "[f\Q|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) \ (\x\S. is_segment x) \ + P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" +proof - + have "a\Q \ b\Q \ a\b" + 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 "a\P \ b\P" + using \Q={a,b}\ assms(4) by auto + have "a\b" 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) \ + P1\P2={} \ (\x\{s}. (x\P1={} \ x\P2={} \ (\y\{s}. x\y \ x\y={})))" + proof (rule conjI) + { fix x assume "x\P" + have "[a;x;b] \ [b;a;x] \ [a;b;x] \ x=a \ x=b" + using \a\P \ b\P\ some_betw path_P \a\b\ + by (meson \x \ P\ abc_sym) + then have "x\s \ x\P1 \ x\P2 \ 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 "x\s \ x\P1 \ x\P2 \ 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 "x\P" + 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) \ P1\P2={} \ + (\x\{s}. (x\P1={} \ x\P2={} \ (\y\{s}. x\y \ x\y={})))" + 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 "x\P1" "x\P2" + 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 "xa\s" "xa\P1" + 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 "xa\s" "xa\P2" + 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: "[f\Q|a..b..c]" + fixes S defines S_def: "S \ {segment (f i) (f(i+1)) | i. iS) \ Q" +proof + let ?N = "card Q" + have f_def_2: "a\Q \ b\Q \ c\Q" + 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 "p\Q" thus ?thesis by blast + next assume "p\Q" + hence "p\a \ p\c" + 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 n\<^sub>y n\<^sub>z y z + where yz_def: "y=f n\<^sub>y" "z=f n\<^sub>z" "[y;p;z]" "y\Q" "z\Q" "n\<^sub>y=n\<^sub>z-1" "n\<^sub>z(\k < card Q. [f k; p; c] \ k>n\<^sub>y)" "\(\kz. [a; p; f k])" + using get_closest_chain_events [where f=f and x=p and Y=Q and a\<^sub>n=c and a\<^sub>0=a and a=b] + f_def \p\Q\ + by metis + have "n\<^sub>y?s" + using \[y;p;z]\ abc_abc_neq seg_betw yz_def(1,2) + by blast + have "n\<^sub>z = n\<^sub>y + 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 "?s\S" + using S_def \n\<^sub>y assms(2) + by blast + hence "p\\S" + using \p \ ?s\ by blast + thus ?thesis by blast + qed + qed + show "?u\?i" + proof + fix p assume "p \ ?u" + hence "p\\S \ p\Q" by blast + thus "p\?i" + proof + assume "p\Q" + 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 "p\\S" + then obtain s where "p\s" "s\S" + by blast + then obtain y where "s = segment (f y) (f (y+1))" "yQ \ 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 \3\card 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 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\ \?N\3\ + 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 \?N\3\ 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\ \?N\3\ + 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" "Q\P" "N\3" + and f_def: "a\Q \ b\Q \ c\Q" "[f\Q|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: "a\P \ b\P \ c\P" + 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: "a\b \ b\c \ a\c" + 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 "x\P" + 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 "(\s\S. x\s) \ x\P1 \ x\P2 \ x\Q" + 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 "x\Q \ x\\S" + 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 " x\P1 \ x\P2 \ x\Q" + 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 "(\s\S. x\s) \ x\P1 \ x\P2 \ x\Q" + by blast + hence "x\\S \ [b;a;x] \ [b;c;x] \ x\Q" + using S_def P1_def P2_def + unfolding segment_def prolongation_def + by auto + hence "x\P" + proof (cases) + assume "x\\S" + have "S = {segment (f i) (f(i+1)) | i. iinterval a c" + using int_split_to_segs [OF f_def(2)] assms \x\\S\ + 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 "x\\S" + hence "[b;a;x] \ [b;c;x] \ x\Q" + 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" "Q\P" "N\3" + and f_def: "a\Q \ b\Q \ c\Q" "[f\Q|a..b..c]" + and S_def: "S = {s. \i<(N-1). s = segment (f i) (f (i+1))}" + and x_def: "x\s" "s\S" + 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: "x\a \ x\c" if "x\s" "s\S" for x s + proof + show "x\a" + proof (rule notI) + assume "x=a" + obtain n where s_def: "s = segment (f n) (f (n+1))" "ns \ S\ by blast + hence "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 \x\s\ 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 "x\c" + proof (rule notI) + assume "x=c" + obtain n where s_def: "s = segment (f n) (f (n+1))" "ns \ S\ by blast + hence "n+1x\s\ 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 "x\interval a c" + using int_split_to_segs [OF f_def(2)] S_def assms(2,3,5) x_def + by blast + have "x\a \ x\c" 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" "Q\P" "N\3" + and f_def: "a\Q \ b\Q \ c\Q" "[f\Q|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 "P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" +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 "x\P1" "x\P2" + 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 "\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={}))" + proof (rule ballI) + fix s assume "s\S" + show "s \ P1 = {} \ s \ P2 = {} \ (\y\S. s \ y \ s \ y = {})" + proof (intro conjI ballI impI) + show "s\P1={}" + proof (safe) + fix x assume "x\s" "x\P1" + 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 "s\P2={}" + proof (safe) + fix x assume "x\s" "x\P2" + 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 "r\S" "s\r" + show "s\r={}" + 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))" + "n\m" "nr \ 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 \y\r\ \y\s\ rs_def(1,2) by simp + have False + proof (cases) + assume "nn < m\ assms(3) fQ chain_defs order_finite_chain rs_def(5) by (metis assms(2) thm2_ind1) + have "n+1[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 "\nm" 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+1n > 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" "Q\P" "N\3" + and f_def: "a\Q \ b\Q \ c\Q" "[f\Q|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) \ + (\x\S. is_segment x) \ + P1\P2={} \ (\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" +proof (intro disjoint_segmentation conjI) + show "P = ((\S) \ P1 \ P2 \ Q)" + using path_is_union assms + by blast + show "\x\S. is_segment x" + proof + fix s assume "s\S" + 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: "[f\Q|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: "Q\P" + and f_def: "[f\Q|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. iS) \ P1 \ P2 \ Q)" "(\x\S. is_segment x)" + "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" +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: "a\Q \ b\Q" + using f_def points_in_chain finite_chain_with_def by auto + have "a\b" + 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)" "(\x\S. is_segment x)" "P1\P2={}" + "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" + 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: "[f\Q|a..c..b]" + using assms(3,5) \a\b\ 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)" "(\x\S. is_segment x)" "P1\P2={}" + "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" + 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)" "(\x\S. is_segment x)" "P1\P2={}" + "(\x\S. (x\P1={} \ x\P2={} \ (\y\S. x\y \ x\y={})))" by meson+ + thus "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" + "P = ((\S) \ P1 \ P2 \ Q)" "(\x\S. 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 Q\2" "Q\P" + shows "\S P1 P2. P = ((\S) \ P1 \ P2 \ Q) \ + disjoint (S\{P1,P2}) \ P1\P2 \ P1\S \ P2\S \ + (\x\S. 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: "[f\Q|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?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) \[f\Q|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 "[f\X|a..c]" "f (card X - 2) = p" "3 \ card X" "X = insert c Y" "c\Y" + shows "[f\Y|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 "p\X" + 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 "[f\X|a..p..c]" + unfolding finite_long_chain_with_alt by (simp add: assms(1) \[a;p;c]\ \p\X\) + + have 1: "x = a" if "x \ Y" "\ [a;x;p]" "x \ p" for x + proof - + have "x\X" + 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" "ix\X\ 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 \[f\X|a..p..c]\] by simp +qed + + +lemma (in MinkowskiChain) fin_long_ch_imp_fin_ch: + assumes "[f\X|a..b..c]" + shows "[f\X|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 "[f\X|a..c]" + and "[g\X|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: "[f\X|a..b..c]" + using chain_defs by (metis Nil.prems(1,3) three_in_set3) + obtain y where g_ch: "[g\X|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: "[f\X|a..b..c]" + using IH(3,5) chain_defs by (metis three_in_set3) + obtain y where g_ch: "[g\X|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 \[f\X]\ order_finite_chain2 by force + qed + hence "p\c \ p\a" + using abc_abc_neq by blast + + obtain Y where Y_def: "X = insert c Y" "c\Y" + 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\ \c\Y\ 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 "p\Y" + 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 "[f\Y|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: "\ (\q\X. [p;q;c])" + proof + assume "(\q\X. [p;q;c])" + then obtain q where "q\X" "[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 \p\c \ p\a\ 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\ \p\Y\ points_in_chain \p\c \ p\a\ + 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 "d\X" + 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\ \c\Y\ 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 "[f\X|a..c]" + and "[g\X|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) \<^term>\local_ordering\ of the chain. + This could be made generic over the \<^term>\local_ordering\ + similar to @{thm chain_sym} relying on @{thm ordering_sym_loc}. +\ + +lemma chain_unique_upto_rev_cases: + assumes ch_f: "[f\X|a..c]" + and ch_g: "[g\X|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 "[f\X|a..c]" "[g\X|x..z]" "card X \ 3" "i < card X" + shows "f i = g i \ f i = g (card X - i - 1)" "a=x\c=z \ c=x\a=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=x\c=z) \ (c=x\a=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 \<^term>\local_ordering\ - + cf Schutz' monograph \cite[ p.~27]{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 "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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\\

" "a\P" "b\P" "c\P" "d\P" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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\\

" "a\P" "b\P" "c\P" "d\P" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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 "b\c" + shows "[a;b;c;d] \ [a;c;b;d]" +proof - + obtain P where "P\\

" "a\P" "b\P" "d\P" + using abc_ex_path abd by blast + have "c\P" + 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; x\I; (x\a \ x\b)\ \ P x I" + shows any_element: "\x I. \x\I; (\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. \I\A; Q I a b\ \ b\A \ a\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; I\A; J\A; [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; I\A; J\A; + a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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" + "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + have endpoints_on_path: "a\A" "b\A" "c\A" "d\A" + 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; I\A; J\A; 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. \I\A; Q I a b\ \ b\A \ a\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; I\A; J\A; [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; I\A; J\A; + a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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 \ b\c \ 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 \ c\d) \ (a=b \ b\c \ c\d \ a\d)\ \ 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; + (a\b \ b\c \ c\d \ a=d) \ (a\b \ b=c \ c\d \ a=d)\ \ P I J" + shows "\I J a b c d. \Q I a b; Q J c d; \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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. \I\A; A\\

; Q I a b\ \ b\A \ a\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 A. \Q I a b; Q J c d; I\A; J\A; 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; I\A; J\A; 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; I\A; J\A; A\\

; + a\b \ b\c \ c\d \ a=d\ \ P I J" +proof - + have last_case: "\I J a b c d A. \Q I a b; Q J c d; I\A; J\A; 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; I\A; J\A; A\\

; + a\b \ b\c \ c\d \ 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. \I\A; Q I a b\ \ b\A \ a\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; 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)" + shows "\I J a b c d. \Q I a b; Q J c d; I\A; J\A; + \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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; I\A; J\A; + [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; I\A; J\A; + [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; I\A; J\A; + a\b \ b\c \ c\d \ 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; I\A; J\A; + \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ P I J" + proof - + + text \Fix the sets on the path, and obtain the assumptions of \wlog_endpoints_degenerate1\.\ + fix I J + assume asm1: "I\A" "J\A" + 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 \ b\c \ 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 \ c\d\ \ P I J" + "\ a b c d. \Q I a b; Q J c d; a=b \ b\c \ c\d \ a\d\ \ 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; a\b \ b\c \ c\d \ a=d\ \ P I J" + "\ a b c d. \Q I a b; Q J c d; a\b \ b=c \ c\d \ 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" "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" + 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. \I\A; Q I a b\ \ b\A \ a\A" + 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; 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) + \ ([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" "I\A" "J\A" + 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; I\Q; J\Q; Q\\

; + a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ \ 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" "I\A" "J\A" "A\\

" "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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" "I\A" "J\A" "A\\

" + "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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; I\Q; J\Q; Q\\

\ + \ ((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)" + shows "\I J a b c d Q. \I = interval a b; J = interval c d; I\Q; J\Q; Q\\

; + \(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)\ \ 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" "I\A" "J\A" "A\\

" "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" + 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" "I\A" "J\A" "A\\

" + "\ (a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" + 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" "I1\P" "I2\P" "P\\

" "I1\I2 \ {}" + and events_neq: "a\b" "a\c" "a\d" "b\c" "b\d" "c\d" + shows "is_interval (I1 \ I2)" +proof - + have on_path: "a\P \ b\P \ c\P \ d\P" + using assms(1-4) interval_def by auto + + let ?prop = "\ I J. is_interval (I\J) \ (I\J) = {}" (* 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 "I\J = {}" + proof (rule ccontr) + assume "I\J\{}" + then obtain x where "x\I\J" + by blast + show False + proof (cases) + assume "x\a \ x\b \ x\c \ x\d" + hence "[a;x;b]" "[c;x;d]" + using \I=interval a b\ \x\I\J\ \J=interval c d\ \x\I\J\ + 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 "\(x\a \ x\b \ x\c \ x\d)" + 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 "I\J = interval c b" + proof (safe) + fix x + assume "x \ interval c b" + { + assume "x=b \ x=c" + hence "x\I" + using \[a;c;b;d]\ \I = interval a b\ interval_def seg_betw by auto + have "x\J" + using \x=b \ x=c\ + using \[a;c;b;d]\ \J = interval c d\ interval_def seg_betw by auto + hence "x\I \ x\J" using \x \ I\ by blast + } moreover { + assume "\(x=b \ x=c)" + hence "[c;x;b]" + using \x\interval 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 "x\I" "x\J" + using \I = interval a b\ \[a;x;b]\ \J = interval c d\ \[c;x;d]\ + interval_def seg_betw by auto + } + ultimately show "x\I" "x\J" by blast+ + next + fix x + assume "x\I" "x\J" + show "x \ interval c b" + proof (cases) + assume not_eq: "x\a \ x\b \ x\c \ x\d" + have "[a;x;b]" "[c;x;d]" + using \x\I\ \I = interval a b\ \x\J\ \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: "\(x\a \ x\b \ x\c \ x\d)" + { + 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 \x\J\ \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 "d\interval 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 \x\I\ \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 "I\J = interval c d" + proof (safe) + fix x + assume "x \ interval c d" + { + assume "x\c \ x\d" + 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 "x\I \ x\J" by (simp add: \x \ J\) + } moreover { + assume "\ (x\c \ x\d)" + hence "x\I \ x\J" + 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 "x\I" "x\J" by blast+ + next + fix x + assume "x\I" "x\J" + 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 (I1\I2)" + 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" "I\J \ {}" "I\P" "J\P" "P\\

" + and events_deg: "\(a\b \ b\c \ c\d \ a\d \ a\c \ b\d)" + shows "is_interval (I \ J)" +proof - + + let ?p = "\ I J. (is_interval (I \ J) \ I\J = {})" + + 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; I\Q; J\Q; Q\\

\ + \ ((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)" + proof - + fix I J a b c d Q + assume "I = interval a b" "J = interval c d" "I\Q" "J\Q" "Q\\

" + 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)" + 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 "I\J = {}" thus ?thesis by simp + next + assume "I\J \ {}" + have "I\J = interval a b" + proof (safe) + fix x assume "x\I" "x\J" + thus "x\interval a b" + using \I = interval a b\ by blast + next + fix x assume "x\interval a b" + show "x\I" + 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 "x\J" + 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 "I\J = {}" thus ?thesis by simp + next + assume "I\J \ {}" + have "I\J = {a}" + proof (safe) + fix x assume "x\I" "x\J" "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 "a\I" + by (simp add: \I = interval a b\ ends_in_int) + show "a\J" + 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" "I\P" "J\P" "P\\

" "I\J \ {}" + 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 "\x\S. (is_interval x \ x\P)" "P\\

" "\S \ {}" "finite S" "S\{}" + shows "is_interval (\S)" +proof - + obtain n where "n = card S" + by simp + consider "n=0"|"n=1"|"n\2" + 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 "2\n" + obtain m where "m+2=n" + using \2 \ n\ le_add_diff_inverse2 by blast + have ind: "\S. \\x\S. (is_interval x \ x\P); 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}" "I\J" + by (meson card_2_iff) + then have "I\S" "J\S" + by blast+ + then have "is_interval I" "is_interval J" "I\P" "J\P" + by (simp add: "0.prems"(1))+ + also have "I\J \ {}" + using \S={I,J}\ "0.prems"(3) by force + then have "is_interval(I\J)" + 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 "I\S" "S = insert I S'" "I\S'" + 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'\ \I\S'\ by auto + moreover have "\x\S'. 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: "a\Q" + shows "\R L. (is_ray_on R Q \ is_ray_on L Q + \ Q-{a} \ (R \ L) \<^cancel>\events of Q excl. a belong to two rays\ + \ (\r\R. \l\L. [l;a;r]) \<^cancel>\a is betw any 'a of one ray and any 'a of the other\ + \ (\x\R. \y\R. \ [x;a;y]) \<^cancel>\but a is not betw any two events \\ + \ (\x\L. \y\L. \ [x;a;y]))" \<^cancel>\\ of the same ray\ +proof - + text \Schutz here uses Theorem 6, but we don't need it.\ + obtain b where "b\\" and "b\Q" and "b\a" + 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: "\x\Q. [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" "w\x" "y\z" 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" "w\x" + by (metis \b \ Q\ \b \ a\ abc_sym event_a mem_Collect_eq path_Q prolong_betw3) + obtain z where "z\?R" "y\z" + 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 "e\ray 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 "e\ray a x" + proof (cases) + assume "e=x" + thus ?thesis + by (simp add: ray_def) + next + assume "e\x" + hence "[e;x;a] \ [x;e;a]" using wxyz_ord + by (meson \[e;a;b]\ \[x;a;b]\ abc_abd_bcdbdc abc_sym) + thus "e\ray 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 "e\ray 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 "e\ray 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 "y\e" + hence "[a;e;y] \ [a;y;e]" + using aey_cases by auto + hence "e\ray 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 "e\ray 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 \ c\L \ (\r\R. \l\L-{c}. [l;c;r])" + + +lemma int_on_path: + assumes "l\L" "r\R" "Q\\

" + and partition: "L\Q" "L\{}" "R\Q" "R\{}" "L\R=Q" + shows "interval l r \ Q" +proof + fix x assume "x\interval l r" + thus "x\Q" + 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\\

" "[f\X|(f 0)..]" "X\Q" "closest_bound c X" "is_bound_f b X f" "b\c" + assumes "is_bound_f x X f" + shows "x=b \ x=c \ [c;x;b] \ [c;b;x]" +proof - + have "x\Q" + 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 "x\b" "x\c" + 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\\

" "[f\X|(f 0)..]" "X\Q" "closest_bound_f c X f" "is_bound_f b X f" "b\c" + assumes "x=b \ x=c \ [c;x;b] \ [c;b;x]" + shows "is_bound_f x X f" +proof - + have "x\Q" + 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 "[f\X|f 0..]" + by (simp add: assms(2)) + assume "i [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 "[f\X|f 0..]" + by (simp add: assms(2)) + assume "i [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\\

" "[f\X|(f 0)..]" "X\Q" "closest_bound_f c X f" "is_bound_f b X f" "b\c" + 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 "c\ray a b" + using abc_abc_neq pro_betw ray_def by auto + have "c\interval 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: "b\Q" "b\\" + and unreach: "Q\<^sub>x \ unreach-on Q from b" "Q\<^sub>z \ unreach-on Q from b" + and xyz: "[Q\<^sub>x; Q\<^sub>y; Q\<^sub>z]" + shows "Q\<^sub>y \ unreach-on Q from b" +proof - + have xz: "Q\<^sub>x \ Q\<^sub>z" using abc_ac_neq xyz by blast + + text \First we obtain the chain from @{thm I6}.\ + have in_Q: "Q\<^sub>x\Q \ Q\<^sub>y\Q \ Q\<^sub>z\Q" + using betw_b_in_path path_Q unreach(1,2) xz unreach_on_path xyz by blast + hence event_y: "Q\<^sub>y\\" + 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 = Q\<^sub>x" "f (card X - 1) = Q\<^sub>z" + "(\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 \ Q\<^sub>x \ X \ Q\<^sub>z \ X \ (\Q\<^sub>y\\. [Q\<^sub>x; Q\<^sub>y; Q\<^sub>z] \ Q\<^sub>y \ 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" "N\2" + 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 "[f\X|Q\<^sub>x..Q\<^sub>z]" + 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 \<^term>\local_ordering\. Just below is the statement we use @{thm disjE} with.\ + have y_cases: "Q\<^sub>y\X \ Q\<^sub>y\X" by blast + have y_int: "Q\<^sub>y\interval Q\<^sub>x Q\<^sub>z" + using interval_def seg_betw xyz by auto + have X_in_Q: "X\Q" + using chain_on_path_I6 [where Q=Q and X=X] X_def event_b path_Q unreach xz \[f\X|Q\<^sub>x .. Q\<^sub>z]\ 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 "N\2" + hence "N\3" using \2 \ N\ by auto + have "2\card X" using \2 \ N\ \N = card X\ by blast + show ?thesis using y_cases + proof (rule disjE) + assume "Q\<^sub>y\X" + then obtain i where i_def: "iy = f i" + using X_def(1) by (metis fin_X obtain_index_fin_chain) + have "i\0 \ i\card 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 "Q\<^sub>y\X" + + let ?S = "if card X = 2 then {segment ?a ?d} else {segment (f i) (f(i+1)) | i. iy\\?S" + proof - + obtain c where "[f\X|Q\<^sub>x..c..Q\<^sub>z]" + using X_def(1) \N = card X\ \N\2\ \[f\X|Q\<^sub>x..Q\<^sub>z]\ short_ch_card_2 + by (metis \2 \ N\ le_neq_implies_less long_chain_2_imp_3) + have "interval Q\<^sub>x Q\<^sub>z = \?S \ X" + using int_split_to_segs [OF \[f\X|Q\<^sub>x..c..Q\<^sub>z]\] by auto + thus ?thesis + using \Q\<^sub>y\X\ y_int by blast + qed + then obtain s where "s\?S" "Q\<^sub>y\s" by blast + + have "\i. i\{1..(card X)-1} \ [(f(i-1)); Q\<^sub>y; f i]" + proof - + obtain i' where i'_def: "i' < N-1" "s = segment (f i') (f (i' + 1))" + using \Q\<^sub>y\s\ \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); Q\<^sub>y; f(i'+1)]" + using i'_def(2) \Q\<^sub>y\s\ seg_betw by simp + qed + qed + then obtain i where i_def: "i\{1..(card X)-1}" "[(f(i-1)); Q\<^sub>y; 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 "\x\A. [a;x;b]" "\x\B. [c;x;d]" "A\Q" "B\Q" "Q\\

" + "card A > 1 \ infinite A" "card B > 1 \ infinite B" + shows "\l\Q. \u\Q. \x\A\B. [l;x;u]" +proof - + let ?P = "\ A B. \l\Q. \u\Q. \x\A\B. [l;x;u]" + let ?I = "\ A a b. (card A > 1 \ infinite A) \ (\x\A. [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 "A\Q" "?I A a b" + show "b\Q\a\Q" + 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 "x\A" "y\A" "x\y" + 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 "x\Q" "y\Q" + using \A \ Q\ \x \ A\ \y \ A\ by auto + have "[a;x;b]" "[a;y;b]" + by (simp add: \(1 < card A \ infinite A) \ (\x\A. [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 "a\Q \ b\Q" + using \Q\\

\ \x\Q\ \x\y\ \x\Q\ \y\Q\ betw_a_in_path betw_c_in_path by blast + thus ?thesis by simp + qed + qed + + show ?thesis + proof (cases) + assume "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + 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" "A\Q" "B\Q" "Q\\

" + using assms by simp+ + show "a\b \ a\c \ a\d \ b\c \ b\d \ c\d" + using \a\b \ a\c \ a\d \ b\c \ b\d \ c\d\ 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" "I\Q" "J\Q" + 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\ I\J. [a;x;d]" + by (metis Un_iff asm betw4_strong betw4_weak that(1) that(2)) + moreover have "a\Q" "d\Q" + 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\ I\J. [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 "a\Q" "d\Q" + 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\ I\J. [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 "a\Q" "b\Q" + using assms(5) on_path that(1-4) by blast+ + ultimately show ?thesis by blast + qed + qed + qed + qed + next + assume "\(a\b \ a\c \ a\d \ b\c \ b\d \ c\d)" + + 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" "A\Q" "B\Q" "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 "\l\Q. \u\Q. \x\I \ 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 "\l\Q. \u\Q. \x\I \ 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\ I\J. [c;x;d]" + using abc_abc_neq that(1,2) by fastforce + moreover have "c\Q" "d\Q" + using on_path \a = b \ b = c \ c \ d\ that(1,3) abc_abc_neq by metis+ + ultimately show "\l\Q. \u\Q. \x\I \ J. [l;x;u]" by blast + next + assume "a = b \ b \ c \ c \ d \ a \ d" + hence "\x\ I\J. [c;x;d]" + using abc_abc_neq that(1,2) by fastforce + moreover have "c\Q" "d\Q" + using on_path \a = b \ b \ c \ c \ d \ a \ d\ that(1,3) abc_abc_neq by metis+ + ultimately show "\l\Q. \u\Q. \x\I \ J. [l;x;u]" by blast + next + assume "a \ b \ b = c \ c \ d \ a = d" + hence "\x\ I\J. [c;x;d]" + using abc_sym that(1,2) by auto + moreover have "c\Q" "d\Q" + using on_path \a \ b \ b = c \ c \ d \ a = d\ that(1,3) abc_abc_neq by metis+ + ultimately show "\l\Q. \u\Q. \x\I \ J. [l;x;u]" by blast + next + assume "[a;b;c] \ a = d" + hence "\x\ I\J. [c;x;d]" + by (metis UnE abc_acd_abd abc_sym that(1,2)) + moreover have "c\Q" "d\Q" + using on_path that(2,4) by blast+ + ultimately show "\l\Q. \u\Q. \x\I \ J. [l;x;u]" by blast + next + assume "[b;a;c] \ a = d" + hence "\x\ I\J. [c;x;b]" + using abc_sym abd_bcd_abc betw4_strong that(1,2) by (metis Un_iff) + moreover have "c\Q" "b\Q" + using on_path that by blast+ + ultimately show "\l\Q. \u\Q. \x\I \ J. [l;x;u]" by blast + qed + qed + qed +qed + + +lemma (*for 14i*) union_of_bounded_sets_is_bounded2: + assumes "\x\A. [a;x;b]" "\x\B. [c;x;d]" "A\Q" "B\Q" "Q\\

" + "1 infinite A" "1 infinite B" + shows "\l\Q-(A\B). \u\Q-(A\B). \x\A\B. [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: "a\Q" "b\Q" + and reachable: "path_ex a q1" "path_ex b q2" "q1\Q" "q2\Q" + shows "\y\Q. \z\Q. (\x\unreach-on Q from a. [y;x;z]) \ (\x\unreach-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 "\q\Q. q\(unreach-on Q from a)" "\q\Q. 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: "\la\Q. \ua\Q. la\unreach-on Q from a \ ua\unreach-on Q from a \ (\x\unreach-on Q from a. [la;x;ua])" + if asm: "a\Q" "path_ex a q" "q\Q" + for a q + proof - + obtain Qy where "Qy\unreach-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 \ (\x\unreach-on Q from a. (x\la \ x\ua) \ [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 "\x\unreach-on Q from a. x \ la \ x \ ua \ [la;x;ua]" + proof (safe) + fix x assume "x\unreach-on Q from a" "x\la" "x\ua" + { + assume "x=Qy" hence "[la;x;ua]" by (simp add: \[la;Qy;ua]\) + } moreover { + assume "x\Qy" + have "[Qy;x;la] \ [la;Qy;x]" + proof - + { assume "[x;la;Qy]" + hence "la\unreach-on Q from a" + using unreach_connected \Qy\unreach-on Q from a\\x\unreach-on Q from a\\x\Qy\ 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 "\y\Q. \z\Q. (\x\(unreach-on Q from a)\(unreach-on Q from b). [y;x;z])" + proof - + obtain la ua where "\x\unreach-on Q from a. [la;x;ua]" + using events(1) get_bds reachable(1,3) by blast + obtain lb ub where "\x\unreach-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 \\x\unreach-on Q from a. [la;x;ua]\ \\x\unreach-on Q from b. [lb;x;ub]\ + by blast + qed + + then obtain y z where "y\Q" "z\Q" "(\x\(unreach-on Q from a)\(unreach-on Q from b). [y;x;z])" + by blast + show ?thesis + proof (rule bexI)+ + show "y\Q" by (simp add: \y \ Q\) + show "z\Q" by (simp add: \z \ Q\) + show "(\x\unreach-on Q from a. [z;x;y]) \ (\x\unreach-on Q from b. [z;x;y])" + by (simp add: \\x\unreach-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: "a\Q" "b\Q" "c\Q" "d\Q" "c\d" + and reachable: "\P\\

. \q\Q. path P a q" "\P\\

. \q\Q. path P b q" + shows "\e\Q. \ae\\

. \be\\

. path ae a e \ path be b e \ [c;d;e]" +proof - + obtain y z where bounds_yz: "(\x\unreach-on Q from a. [z;x;y]) \ (\x\unreach-on Q from b. [z;x;y])" + and yz_inQ: "y\Q" "z\Q" + 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 \(\x\unreach-on Q from a. [z;x;y]) \ (\x\unreach-on Q from b. [z;x;y])\ abc_abc_neq)+ + let ?P = "\e ae be. (e\Q \ path ae a e \ path be b e \ [c;d;e])" + + have exist_ay: "\ay. path ay a y" + if "a\Q" "\P\\

. \q\Q. path P a q" "y\(unreach-on Q from a)" "y\Q" + 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 \y\Q\ 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 \z\Q\ 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 "e\Q" + 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\\

" "Q\R" + and events: "x\Q" "x\R" "a\R" "a\x" "b\Q" + and reachable: "\P\\

. \q\Q. path P b q" + shows "\e\\. \ae\\

. \be\\

. path ae a e \ path be b e \ (\y\unreach-on Q from a. [x;y;e])" +proof - + have "a\Q" + 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 "x\d" + using \d \ unreach-on Q from a\ cross_in_reachable events(1) events(2) events(3) paths(2) by auto + have "d\Q" + using \d \ unreach-on Q from a\ unreach_on_path by blast + + have "\e\Q. \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 "y\Q" + 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 \d\Q\ \y\Q\ 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 \y\Q\ 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 Q\<^sub>x=y and Q\<^sub>y=x and Q\<^sub>z=d and b=a] + using \\[x;y;e]\ \a\Q\ \d\unreach-on Q from a\ \y\unreach-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 Q\<^sub>x=y and Q\<^sub>y=e and Q\<^sub>z=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 \<^term>\local_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: "Q\P" + and f_def: "[f\Q|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)) \ iai\{(0::nat)..?N-2}\ \?N\3\ + 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 ?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 "i\j" + 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 \?N\3\ 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 "j\i+1" + have "i < card Q \ j < card Q \ (i+1) < card Q" + using add_mono_thms_linordered_field(3) asm(1,2) assms \?N\3\ 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 \?N\3\ + 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: "Q\P" + and f_def: "[f\Q|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. iS) \ P1 \ P2 \ Q)" + (* The union of these segments and prolongations with the separating points is the path. *) + "card S = (card Q-1) \ (\x\S. is_segment x)" + (* There are N-1 segments. *) + (* There are two prolongations. *) + "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" + (* 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)" "(\x\S. is_segment x)" + "disjoint (S\{P1,P2})" "P1\P2" "P1\S" "P2\S" + 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})" "P1\P2" "P1\S" "P2\S" + 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 "?N\2" + hence "?N\3" + using \2 \ card Q\ by linarith + then obtain c where "[f\Q|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) \[f\Q|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 \ No newline at end of file diff --git a/thys/Schutz_Spacetime/TernaryOrdering.thy b/thys/Schutz_Spacetime/TernaryOrdering.thy --- a/thys/Schutz_Spacetime/TernaryOrdering.thy +++ b/thys/Schutz_Spacetime/TernaryOrdering.thy @@ -1,342 +1,460 @@ -(* Title: Schutz_Spacetime/TernaryOrdering.thy - Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot - University of Edinburgh, 2021 -*) -theory TernaryOrdering - imports Util - -begin - -text \ - Definition of chains using an ordering on sets of events - based on natural numbers, plus some proofs. -\ - -section \Totally ordered chains\ - -text \ - Based on page 110 of Phil Scott's thesis and the following HOL Light definition: - \begin{verbatim} - let ORDERING = new_definition - `ORDERING f X <=> (!n. (FINITE X ==> n < CARD X) ==> f n IN X) - /\ (!x. x IN X ==> ?n. (FINITE X ==> n < CARD X) - /\ f n = x) - /\ !n n' n''. (FINITE X ==> n'' < CARD X) - /\ n < n' /\ n' < n'' - ==> between (f n) (f n') (f n'')`;; - \end{verbatim} - I've made it strict for simplicity, and because that's how Schutz's ordering is. It could be - made more generic by taking in the function corresponding to $<$ as a paramater. - Main difference to Schutz: he has local order, not total (cf Theorem 2 and \ordering2\). -\ - -definition ordering :: "(nat \ 'a) \ ('a \ 'a \ 'a \ bool) \ 'a set \ bool" where - "ordering f ord X \ (\n. (finite X \ n < card X) \ f n \ X) - \ (\x\X. (\n. (finite X \ n < card X) \ f n = x)) - \ (\n n' n''. (finite X \ n'' < card X) \ n < n' \ n' < n'' - \ ord (f n) (f n') (f n''))" - -lemma ordering_ord_ijk: - assumes "ordering f ord X" - and "i < j \ j < k \ (finite X \ k < card X)" - shows "ord (f i) (f j) (f k)" -by (metis ordering_def assms) - -lemma empty_ordering [simp]: "\f. ordering f ord {}" -by (simp add: ordering_def) - -lemma singleton_ordering [simp]: "\f. ordering f ord {a}" -apply (rule_tac x = "\n. a" in exI) -by (simp add: ordering_def) - -lemma two_ordering [simp]: "\f. ordering f ord {a, b}" -proof cases - assume "a = b" - thus ?thesis using singleton_ordering by simp -next - assume a_neq_b: "a \ b" - let ?f = "\n. if n = 0 then a else b" - have ordering1: "(\n. (finite {a,b} \ n < card {a,b}) \ ?f n \ {a,b})" by simp - have ordering2: "(\x\{a,b}. \n. (finite {a,b} \ n < card {a,b}) \ ?f n = x)" - using a_neq_b all_not_in_conv card_Suc_eq card_0_eq card_gt_0_iff insert_iff lessI by auto - have ordering3: "(\n n' n''. (finite {a,b} \ n'' < card {a,b}) \ n < n' \ n' < n'' - \ ord (?f n) (?f n') (?f n''))" using a_neq_b by auto - have "ordering ?f ord {a, b}" using ordering_def ordering1 ordering2 ordering3 by blast - thus ?thesis by auto -qed - -lemma card_le2_ordering: - assumes finiteX: "finite X" - and card_le2: "card X \ 2" - shows "\f. ordering f ord X" -proof - - have card012: "card X = 0 \ card X = 1 \ card X = 2" using card_le2 by auto - have card0: "card X = 0 \ ?thesis" using finiteX by simp - have card1: "card X = 1 \ ?thesis" using card_eq_SucD by fastforce - have card2: "card X = 2 \ ?thesis" by (metis two_ordering card_eq_SucD numeral_2_eq_2) - thus ?thesis using card012 card0 card1 card2 by auto -qed - -lemma ord_ordered: - assumes abc: "ord a b c" - and abc_neq: "a \ b \ a \ c \ b \ c" - shows "\f. ordering f ord {a,b,c}" -apply (rule_tac x = "\n. if n = 0 then a else if n = 1 then b else c" in exI) -apply (unfold ordering_def) -using abc abc_neq by auto - -lemma overlap_ordering: - assumes abc: "ord a b c" - and bcd: "ord b c d" - and abd: "ord a b d" - and acd: "ord a c d" - and abc_neq: "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" - shows "\f. ordering f ord {a,b,c,d}" -proof - - let ?X = "{a,b,c,d}" - let ?f = "\n. if n = 0 then a else if n = 1 then b else if n = 2 then c else d" - have card4: "card ?X = 4" using abc bcd abd abc_neq by simp - have ordering1: "\n. (finite ?X \ n < card ?X) \ ?f n \ ?X" by simp - have ordering2: "\x\?X. \n. (finite ?X \ n < card ?X) \ ?f n = x" - by (metis card4 One_nat_def Suc_1 Suc_lessI empty_iff insertE numeral_3_eq_3 numeral_eq_iff - numeral_eq_one_iff rel_simps(51) semiring_norm(85) semiring_norm(86) semiring_norm(87) - semiring_norm(89) zero_neq_numeral) - have ordering3: "(\n n' n''. (finite ?X \ n'' < card ?X) \ n < n' \ n' < n'' - \ ord (?f n) (?f n') (?f n''))" - using card4 abc bcd abd acd card_0_eq card_insert_if finite.emptyI finite_insert less_antisym - less_one less_trans_Suc not_less_eq not_one_less_zero numeral_2_eq_2 by auto - have "ordering ?f ord ?X" using ordering1 ordering2 ordering3 ordering_def by blast - thus ?thesis by auto -qed - -lemma overlap_ordering_alt1: - assumes abc: "ord a b c" - and bcd: "ord b c d" - and abc_bcd_abd: "\ a b c d. ord a b c \ ord b c d \ ord a b d" - and abc_bcd_acd: "\ a b c d. ord a b c \ ord b c d \ ord a c d" - and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - shows "\f. ordering f ord {a,b,c,d}" -by (metis (full_types) assms overlap_ordering) - -lemma overlap_ordering_alt2: - assumes abc: "ord a b c" - and bcd: "ord b c d" - and abd: "ord a b d" - and acd: "ord a c d" - and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - shows "\f. ordering f ord {a,b,c,d}" -by (metis assms overlap_ordering) - -lemma overlap_ordering_alt: - assumes abc: "ord a b c" - and bcd: "ord b c d" - and abc_bcd_abd: "\ a b c d. ord a b c \ ord b c d \ ord a b d" - and abc_bcd_acd: "\ a b c d. ord a b c \ ord b c d \ ord a c d" - and abc_neq: "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" - shows "\f. ordering f ord {a,b,c,d}" -by (meson assms overlap_ordering) - -text \ - The lemmas below are easy to prove for \X = {}\, and if I included that case then I would have - to write a conditional definition in place of \{0..|X| - 1}\. -\ - -lemma finite_ordering_img: "\X \ {}; finite X; ordering f ord X\ \ f ` {0..card X - 1} = X" -by (force simp add: ordering_def image_def) - -lemma inf_ordering_img: "\infinite X; ordering f ord X\ \ f ` {0..} = X" -by (auto simp add: ordering_def image_def) - -lemma finite_ordering_inv_img: "\X \ {}; finite X; ordering f ord X\ \ f -` X = {0..card X - 1}" -apply (auto simp add: ordering_def) -oops - -lemma inf_ordering_inv_img: "\infinite X; ordering f ord X\ \ f -` X = {0..}" -by (auto simp add: ordering_def image_def) - -lemma inf_ordering_img_inv_img: "\infinite X; ordering f ord X\ \ f ` f -` X = X" -using inf_ordering_img by auto - -lemma finite_ordering_inj_on: "\finite X; ordering f ord X\ \ inj_on f {0..card X - 1}" -by (metis finite_ordering_img Suc_diff_1 atLeastAtMost_iff card_atLeastAtMost card_eq_0_iff - diff_0_eq_0 diff_zero eq_card_imp_inj_on gr0I inj_onI le_0_eq) - -lemma finite_ordering_bij: - assumes orderingX: "ordering f ord X" - and finiteX: "finite X" - and non_empty: "X \ {}" - shows "bij_betw f {0..card X - 1} X" -proof - - have f_image: "f ` {0..card X - 1} = X" by (metis orderingX finiteX finite_ordering_img non_empty) - thus ?thesis by (metis inj_on_imp_bij_betw orderingX finiteX finite_ordering_inj_on) -qed - -(* I think there might be a way of proving this without ord_distinct (?) *) -lemma inf_ordering_inj': - assumes infX: "infinite X" - and f_ord: "ordering f ord X" - and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - and f_eq: "f m = f n" - shows "m = n" -(* If m \ n and f m = f n then it wouldn't be an ordering, and this part: - \n n' n''. n < n' \ n' < n'' \ [[f n f n' f n'']] - would fail because two of f n f n' f n'' would be equal, and that violates ord_distinct. *) -proof (rule ccontr) - assume m_not_n: "m \ n" - have betw_3n: "\n n' n''. n < n' \ n' < n'' \ ord (f n) (f n') (f n'')" - using f_ord by (simp add: ordering_def infX) - thus False - proof cases - assume m_less_n: "m < n" - then obtain k where "n < k" by auto - then have "ord (f m) (f n) (f k)" using m_less_n betw_3n by simp - then have "f m \ f n" using ord_distinct by simp - thus ?thesis using f_eq by simp - next - assume "\ m < n" - then have n_less_m: "n < m" using m_not_n by simp - then obtain k where "m < k" by auto - then have "ord (f n) (f m) (f k)" using n_less_m betw_3n by simp - then have "f n \ f m" using ord_distinct by simp - thus ?thesis using f_eq by simp - qed -qed - -(* f is actually injective when X is infinite. *) -lemma inf_ordering_inj: - assumes "infinite X" - and "ordering f ord X" - and "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - shows "inj f" -using inf_ordering_inj' assms by (metis injI) - -text \ - The finite case is a little more difficult as I can't just choose some other natural number - to form the third part of the betweenness relation and the initial simplification isn't as nice. - Note that I cannot prove \inj f\ (over the whole type that \f\ is defined on, i.e. natural numbers), - because I need to capture the \m\ and \n\ that obey specific requirements for the finite case. - In order to prove \inj f\, I would have to extend the definition for ordering to include \m\ and \n\ - beyond \card X\, such that it is still injective. That would probably not be very useful. -\ - -lemma finite_ordering_inj: - assumes finiteX: "finite X" - and f_ord: "ordering f ord X" - and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - and m_less_card: "m < card X" - and n_less_card: "n < card X" - and f_eq: "f m = f n" - shows "m = n" -proof (rule ccontr) - assume m_not_n: "m \ n" - have surj_f: "\x\X. \nn n' n''. n'' < card X \ n < n' \ n' < n'' \ ord (f n) (f n') (f n'')" - using f_ord by (simp add: ordering_def) - show False - proof cases - assume card_le2: "card X \ 2" - have card0: "card X = 0 \ False" using m_less_card by simp - have card1: "card X = 1 \ False" using m_less_card n_less_card m_not_n by simp - have card2: "card X = 2 \ False" - proof (rule impI) - assume card_is_2: "card X = 2" - then have mn01: "m = 0 \ n = 1 \ n = 0 \ m = 1" using m_less_card n_less_card m_not_n by auto - then have "f m \ f n" using card_is_2 surj_f One_nat_def card_eq_SucD insertCI - less_2_cases numeral_2_eq_2 by (metis (no_types, lifting)) - thus False using f_eq by simp - qed - show False using card0 card1 card2 card_le2 by simp - next - assume "\ card X \ 2" - then have card_ge3: "card X \ 3" by simp - thus False - proof cases - assume m_less_n: "m < n" - then obtain k where k_pos: "k < m \ (m < k \ k < n) \ (n < k \ k < card X)" - using is_free_nat m_less_n n_less_card card_ge3 by blast - have k1: "k < m \ord (f k) (f m) (f n)" using m_less_n n_less_card betw_3n by simp - have k2: "m < k \ k < n \ ord (f m) (f k) (f n)" using m_less_n n_less_card betw_3n by simp - have k3: "n < k \ k < card X \ ord (f m) (f n) (f k)" using m_less_n betw_3n by simp - have "f m \ f n" using k1 k2 k3 k_pos ord_distinct by auto - thus False using f_eq by simp - next (* Should work on making these two cases into one; this is quite boilerplatery. *) - assume "\ m < n" - then have n_less_m: "n < m" using m_not_n by simp - then obtain k where k_pos: "k < n \ (n < k \ k < m) \ (m < k \ k < card X)" - using is_free_nat n_less_m m_less_card card_ge3 by blast - have k1: "k < n \ord (f k) (f n) (f m)" using n_less_m m_less_card betw_3n by simp - have k2: "n < k \ k < m \ ord (f n) (f k) (f m)" using n_less_m m_less_card betw_3n by simp - have k3: "m < k \ k < card X \ ord (f n) (f m) (f k)" using n_less_m betw_3n by simp - have "f n \ f m" using k1 k2 k3 k_pos ord_distinct by auto - thus False using f_eq by simp - qed - qed -qed - -lemma ordering_inj: - assumes "ordering f ord X" - and "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" - and "finite X \ m < card X" - and "finite X \ n < card X" - and "f m = f n" - shows "m = n" - using inf_ordering_inj' finite_ordering_inj assms by blast - -lemma ordering_sym: - assumes ord_sym: "\a b c. ord a b c \ ord c b a" - and "finite X" - and "ordering f ord X" - shows "ordering (\n. f (card X - 1 - n)) ord X" -unfolding ordering_def using assms(2) - apply auto - apply (metis ordering_def assms(3) card_0_eq card_gt_0_iff diff_Suc_less gr_implies_not0) -proof - - fix x - assume "finite X" - assume "x \ X" - obtain n where "finite X \ n < card X" and "f n = x" - by (metis ordering_def \x \ X\ assms(3)) - have "f (card X - ((card X - 1 - n) + 1)) = x" - by (simp add: Suc_leI \f n = x\ \finite X \ n < card X\ assms(2)) - thus "\nx \ X\ add.commute assms(2) card_Diff_singleton card_Suc_Diff1 diff_less_Suc plus_1_eq_Suc) -next - fix n n' n'' - assume "finite X" - assume "n'' < card X" "n < n'" "n' < n''" - have "ord (f (card X - Suc n'')) (f (card X - Suc n')) (f (card X - Suc n))" - using assms(3) unfolding ordering_def - using \n < n'\ \n' < n''\ \n'' < card X\ diff_less_mono2 by auto - thus " ord (f (card X - Suc n)) (f (card X - Suc n')) (f (card X - Suc n''))" - using ord_sym by blast -qed - -lemma zero_into_ordering: - assumes "ordering f betw X" - and "X \ {}" - shows "(f 0) \ X" - using ordering_def - by (metis assms card_eq_0_iff gr_implies_not0 linorder_neqE_nat) - - -section "Locally ordered chains" -text \Definitions for Schutz-like chains, with local order only.\ - -definition ordering2 :: "(nat \ 'a) \ ('a \ 'a \ 'a \ bool) \ 'a set \ bool" where - "ordering2 f ord X \ (\n. (finite X \ n < card X) \ f n \ X) - \ (\x\X. (\n. (finite X \ n < card X) \ f n = x)) - \ (\n n' n''. (finite X \ n'' < card X) \ Suc n = n' \ Suc n' = n'' - \ ord (f n) (f n') (f n''))" - - -text \Analogue to \ordering_ord_ijk\, which is quicker to use in sledgehammer than the definition.\ - -lemma ordering2_ord_ijk: - assumes "ordering2 f ord X" - and "Suc i = j \ Suc j = k \ (finite X \ k < card X)" - shows "ord (f i) (f j) (f k)" - by (metis ordering2_def assms) - - +(* Title: Schutz_Spacetime/TernaryOrdering.thy + Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot + University of Edinburgh, 2021 +*) +theory TernaryOrdering +imports Util + +begin + +text \ + Definition of chains using an ordering on sets of events + based on natural numbers, plus some proofs. +\ + +section \Totally ordered chains\ + +text \ + Based on page 110 of Phil Scott's thesis and the following HOL Light definition: + \begin{verbatim} + let ORDERING = new_definition + `ORDERING f X <=> (!n. (FINITE X ==> n < CARD X) ==> f n IN X) + /\ (!x. x IN X ==> ?n. (FINITE X ==> n < CARD X) + /\ f n = x) + /\ !n n' n''. (FINITE X ==> n'' < CARD X) + /\ n < n' /\ n' < n'' + ==> between (f n) (f n') (f n'')`;; + \end{verbatim} + I've made it strict for simplicity, and because that's how Schutz's ordering is. It could be + made more generic by taking in the function corresponding to $<$ as a paramater. + Main difference to Schutz: he has local order, not total (cf Theorem 2 and \local_ordering\). +\ + +definition ordering :: "(nat \ 'a) \ ('a \ 'a \ 'a \ bool) \ 'a set \ bool" where + "ordering f ord X \ (\n. (finite X \ n < card X) \ f n \ X) + \ (\x\X. (\n. (finite X \ n < card X) \ f n = x)) + \ (\n n' n''. (finite X \ n'' < card X) \ n < n' \ n' < n'' + \ ord (f n) (f n') (f n''))" + +lemma finite_ordering_intro: + assumes "finite X" + and "\n < card X. f n \ X" + and "\x \ X. \n < card X. f n = x" + and "\n n' n''. n < n' \ n' < n'' \ n'' < card X \ ord (f n) (f n') (f n'')" + shows "ordering f ord X" + unfolding ordering_def by (simp add: assms) + +lemma infinite_ordering_intro: + assumes "infinite X" + and "\n::nat. f n \ X" + and "\x \ X. \n::nat. f n = x" + and "\n n' n''. n < n' \ n' < n'' \ ord (f n) (f n') (f n'')" + shows "ordering f ord X" + unfolding ordering_def by (simp add: assms) + +lemma ordering_ord_ijk: + assumes "ordering f ord X" + and "i < j \ j < k \ (finite X \ k < card X)" + shows "ord (f i) (f j) (f k)" + by (metis ordering_def assms) + +lemma empty_ordering [simp]: "\f. ordering f ord {}" + by (simp add: ordering_def) + +lemma singleton_ordering [simp]: "\f. ordering f ord {a}" + apply (rule_tac x = "\n. a" in exI) + by (simp add: ordering_def) + +lemma two_ordering [simp]: "\f. ordering f ord {a, b}" +proof cases + assume "a = b" + thus ?thesis using singleton_ordering by simp +next + assume a_neq_b: "a \ b" + let ?f = "\n. if n = 0 then a else b" + have ordering1: "(\n. (finite {a,b} \ n < card {a,b}) \ ?f n \ {a,b})" by simp + have local_ordering: "(\x\{a,b}. \n. (finite {a,b} \ n < card {a,b}) \ ?f n = x)" + using a_neq_b all_not_in_conv card_Suc_eq card_0_eq card_gt_0_iff insert_iff lessI by auto + have ordering3: "(\n n' n''. (finite {a,b} \ n'' < card {a,b}) \ n < n' \ n' < n'' + \ ord (?f n) (?f n') (?f n''))" using a_neq_b by auto + have "ordering ?f ord {a, b}" using ordering_def ordering1 local_ordering ordering3 by blast + thus ?thesis by auto +qed + +lemma card_le2_ordering: + assumes finiteX: "finite X" + and card_le2: "card X \ 2" + shows "\f. ordering f ord X" +proof - + have card012: "card X = 0 \ card X = 1 \ card X = 2" using card_le2 by auto + have card0: "card X = 0 \ ?thesis" using finiteX by simp + have card1: "card X = 1 \ ?thesis" using card_eq_SucD by fastforce + have card2: "card X = 2 \ ?thesis" by (metis two_ordering card_eq_SucD numeral_2_eq_2) + thus ?thesis using card012 card0 card1 card2 by auto +qed + +lemma ord_ordered: + assumes abc: "ord a b c" + and abc_neq: "a \ b \ a \ c \ b \ c" + shows "\f. ordering f ord {a,b,c}" + apply (rule_tac x = "\n. if n = 0 then a else if n = 1 then b else c" in exI) + apply (unfold ordering_def) + using abc abc_neq by auto + +lemma overlap_ordering: + assumes abc: "ord a b c" + and bcd: "ord b c d" + and abd: "ord a b d" + and acd: "ord a c d" + and abc_neq: "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" + shows "\f. ordering f ord {a,b,c,d}" +proof - + let ?X = "{a,b,c,d}" + let ?f = "\n. if n = 0 then a else if n = 1 then b else if n = 2 then c else d" + have card4: "card ?X = 4" using abc bcd abd abc_neq by simp + have ordering1: "\n. (finite ?X \ n < card ?X) \ ?f n \ ?X" by simp + have local_ordering: "\x\?X. \n. (finite ?X \ n < card ?X) \ ?f n = x" + by (metis card4 One_nat_def Suc_1 Suc_lessI empty_iff insertE numeral_3_eq_3 numeral_eq_iff + numeral_eq_one_iff rel_simps(51) semiring_norm(85) semiring_norm(86) semiring_norm(87) + semiring_norm(89) zero_neq_numeral) + have ordering3: "(\n n' n''. (finite ?X \ n'' < card ?X) \ n < n' \ n' < n'' + \ ord (?f n) (?f n') (?f n''))" + using card4 abc bcd abd acd card_0_eq card_insert_if finite.emptyI finite_insert less_antisym + less_one less_trans_Suc not_less_eq not_one_less_zero numeral_2_eq_2 by auto + have "ordering ?f ord ?X" using ordering1 local_ordering ordering3 ordering_def by blast + thus ?thesis by auto +qed + +lemma overlap_ordering_alt1: + assumes abc: "ord a b c" + and bcd: "ord b c d" + and abc_bcd_abd: "\ a b c d. ord a b c \ ord b c d \ ord a b d" + and abc_bcd_acd: "\ a b c d. ord a b c \ ord b c d \ ord a c d" + and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + shows "\f. ordering f ord {a,b,c,d}" + by (metis (full_types) assms overlap_ordering) + +lemma overlap_ordering_alt2: + assumes abc: "ord a b c" + and bcd: "ord b c d" + and abd: "ord a b d" + and acd: "ord a c d" + and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + shows "\f. ordering f ord {a,b,c,d}" + by (metis assms overlap_ordering) + +lemma overlap_ordering_alt: + assumes abc: "ord a b c" + and bcd: "ord b c d" + and abc_bcd_abd: "\ a b c d. ord a b c \ ord b c d \ ord a b d" + and abc_bcd_acd: "\ a b c d. ord a b c \ ord b c d \ ord a c d" + and abc_neq: "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" + shows "\f. ordering f ord {a,b,c,d}" + by (meson assms overlap_ordering) + +text \ + The lemmas below are easy to prove for \X = {}\, and if I included that case then I would have + to write a conditional definition in place of \{0..|X| - 1}\. +\ + +lemma finite_ordering_img: "\X \ {}; finite X; ordering f ord X\ \ f ` {0..card X - 1} = X" + by (force simp add: ordering_def image_def) + +lemma inf_ordering_img: "\infinite X; ordering f ord X\ \ f ` {0..} = X" + by (auto simp add: ordering_def image_def) + +lemma inf_ordering_inv_img: "\infinite X; ordering f ord X\ \ f -` X = {0..}" + by (auto simp add: ordering_def image_def) + +lemma inf_ordering_img_inv_img: "\infinite X; ordering f ord X\ \ f ` f -` X = X" + using inf_ordering_img by auto + +lemma finite_ordering_inj_on: "\finite X; ordering f ord X\ \ inj_on f {0..card X - 1}" + by (metis finite_ordering_img Suc_diff_1 atLeastAtMost_iff card_atLeastAtMost card_eq_0_iff + diff_0_eq_0 diff_zero eq_card_imp_inj_on gr0I inj_onI le_0_eq) + +lemma finite_ordering_bij: + assumes orderingX: "ordering f ord X" + and finiteX: "finite X" + and non_empty: "X \ {}" + shows "bij_betw f {0..card X - 1} X" +proof - + have f_image: "f ` {0..card X - 1} = X" by (metis orderingX finiteX finite_ordering_img non_empty) + thus ?thesis by (metis inj_on_imp_bij_betw orderingX finiteX finite_ordering_inj_on) +qed + +lemma inf_ordering_inj': + assumes infX: "infinite X" + and f_ord: "ordering f ord X" + and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + and f_eq: "f m = f n" + shows "m = n" +proof (rule ccontr) + assume m_not_n: "m \ n" + have betw_3n: "\n n' n''. n < n' \ n' < n'' \ ord (f n) (f n') (f n'')" + using f_ord by (simp add: ordering_def infX) + thus False + proof cases + assume m_less_n: "m < n" + then obtain k where "n < k" by auto + then have "ord (f m) (f n) (f k)" using m_less_n betw_3n by simp + then have "f m \ f n" using ord_distinct by simp + thus ?thesis using f_eq by simp + next + assume "\ m < n" + then have n_less_m: "n < m" using m_not_n by simp + then obtain k where "m < k" by auto + then have "ord (f n) (f m) (f k)" using n_less_m betw_3n by simp + then have "f n \ f m" using ord_distinct by simp + thus ?thesis using f_eq by simp + qed +qed + +lemma inf_ordering_inj: + assumes "infinite X" + and "ordering f ord X" + and "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + shows "inj f" + using inf_ordering_inj' assms by (metis injI) + +text \ + The finite case is a little more difficult as I can't just choose some other natural number + to form the third part of the betweenness relation and the initial simplification isn't as nice. + Note that I cannot prove \inj f\ (over the whole type that \f\ is defined on, i.e. natural numbers), + because I need to capture the \m\ and \n\ that obey specific requirements for the finite case. + In order to prove \inj f\, I would have to extend the definition for ordering to include \m\ and \n\ + beyond \card X\, such that it is still injective. That would probably not be very useful. +\ + +lemma finite_ordering_inj: + assumes finiteX: "finite X" + and f_ord: "ordering f ord X" + and ord_distinct: "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + and m_less_card: "m < card X" + and n_less_card: "n < card X" + and f_eq: "f m = f n" + shows "m = n" +proof (rule ccontr) + assume m_not_n: "m \ n" + have surj_f: "\x\X. \nn n' n''. n'' < card X \ n < n' \ n' < n'' \ ord (f n) (f n') (f n'')" + using f_ord by (simp add: ordering_def) + show False + proof cases + assume card_le2: "card X \ 2" + have card0: "card X = 0 \ False" using m_less_card by simp + have card1: "card X = 1 \ False" using m_less_card n_less_card m_not_n by simp + have card2: "card X = 2 \ False" + proof (rule impI) + assume card_is_2: "card X = 2" + then have mn01: "m = 0 \ n = 1 \ n = 0 \ m = 1" using m_less_card n_less_card m_not_n by auto + then have "f m \ f n" using card_is_2 surj_f One_nat_def card_eq_SucD insertCI + less_2_cases numeral_2_eq_2 by (metis (no_types, lifting)) + thus False using f_eq by simp + qed + show False using card0 card1 card2 card_le2 by simp + next + assume "\ card X \ 2" + then have card_ge3: "card X \ 3" by simp + thus False + proof cases + assume m_less_n: "m < n" + then obtain k where k_pos: "k < m \ (m < k \ k < n) \ (n < k \ k < card X)" + using is_free_nat m_less_n n_less_card card_ge3 by blast + have k1: "k < m \ord (f k) (f m) (f n)" using m_less_n n_less_card betw_3n by simp + have k2: "m < k \ k < n \ ord (f m) (f k) (f n)" using m_less_n n_less_card betw_3n by simp + have k3: "n < k \ k < card X \ ord (f m) (f n) (f k)" using m_less_n betw_3n by simp + have "f m \ f n" using k1 k2 k3 k_pos ord_distinct by auto + thus False using f_eq by simp + next + assume "\ m < n" + then have n_less_m: "n < m" using m_not_n by simp + then obtain k where k_pos: "k < n \ (n < k \ k < m) \ (m < k \ k < card X)" + using is_free_nat n_less_m m_less_card card_ge3 by blast + have k1: "k < n \ord (f k) (f n) (f m)" using n_less_m m_less_card betw_3n by simp + have k2: "n < k \ k < m \ ord (f n) (f k) (f m)" using n_less_m m_less_card betw_3n by simp + have k3: "m < k \ k < card X \ ord (f n) (f m) (f k)" using n_less_m betw_3n by simp + have "f n \ f m" using k1 k2 k3 k_pos ord_distinct by auto + thus False using f_eq by simp + qed + qed +qed + +lemma ordering_inj: + assumes "ordering f ord X" + and "\a b c. (ord a b c \ a \ b \ a \ c \ b \ c)" + and "finite X \ m < card X" + and "finite X \ n < card X" + and "f m = f n" + shows "m = n" + using inf_ordering_inj' finite_ordering_inj assms by blast + +lemma ordering_sym: + assumes ord_sym: "\a b c. ord a b c \ ord c b a" + and "finite X" + and "ordering f ord X" + shows "ordering (\n. f (card X - 1 - n)) ord X" +unfolding ordering_def using assms(2) + apply auto + apply (metis ordering_def assms(3) card_0_eq card_gt_0_iff diff_Suc_less gr_implies_not0) +proof - + fix x + assume "finite X" + assume "x \ X" + obtain n where "finite X \ n < card X" and "f n = x" + by (metis ordering_def \x \ X\ assms(3)) + have "f (card X - ((card X - 1 - n) + 1)) = x" + by (simp add: Suc_leI \f n = x\ \finite X \ n < card X\ assms(2)) + thus "\nx \ X\ add.commute assms(2) card_Diff_singleton card_Suc_Diff1 diff_less_Suc plus_1_eq_Suc) +next + fix n n' n'' + assume "finite X" + assume "n'' < card X" "n < n'" "n' < n''" + have "ord (f (card X - Suc n'')) (f (card X - Suc n')) (f (card X - Suc n))" + using assms(3) unfolding ordering_def + using \n < n'\ \n' < n''\ \n'' < card X\ diff_less_mono2 by auto + thus " ord (f (card X - Suc n)) (f (card X - Suc n')) (f (card X - Suc n''))" + using ord_sym by blast +qed + +lemma zero_into_ordering: + assumes "ordering f betw X" + and "X \ {}" + shows "(f 0) \ X" + using ordering_def + by (metis assms card_eq_0_iff gr_implies_not0 linorder_neqE_nat) + + +section "Locally ordered chains" +text \Definitions for Schutz-like chains, with local order only.\ + +definition local_ordering :: "(nat \ 'a) \ ('a \ 'a \ 'a \ bool) \ 'a set \ bool" + where "local_ordering f ord X + \ (\n. (finite X \ n < card X) \ f n \ X) \ + (\x\X. \n. (finite X \ n < card X) \ f n = x) \ + (\n. (finite X \ Suc (Suc n) < card X) \ ord (f n) (f (Suc n)) (f (Suc (Suc n))))" + +lemma finite_local_ordering_intro: + assumes "finite X" + and "\n < card X. f n \ X" + and "\x \ X. \n < card X. f n = x" + and "\n n' n''. Suc n = n' \ Suc n' = n'' \ n'' < card X \ ord (f n) (f n') (f n'')" + shows "local_ordering f ord X" + unfolding local_ordering_def by (simp add: assms) + +lemma infinite_local_ordering_intro: + assumes "infinite X" + and "\n::nat. f n \ X" + and "\x \ X. \n::nat. f n = x" + and "\n n' n''. Suc n = n' \ Suc n' = n'' \ ord (f n) (f n') (f n'')" + shows "local_ordering f ord X" + using assms unfolding local_ordering_def by metis + +lemma total_implies_local: + "ordering f ord X \ local_ordering f ord X" + unfolding ordering_def local_ordering_def + using lessI by presburger + +lemma ordering_ord_ijk_loc: + assumes "local_ordering f ord X" + and "finite X \ Suc (Suc i) < card X" + shows "ord (f i) (f (Suc i)) (f (Suc (Suc i)))" + by (metis local_ordering_def assms) + +lemma empty_ordering_loc [simp]: + "\f. local_ordering f ord {}" + by (simp add: local_ordering_def) + +lemma singleton_ordered_loc [simp]: + "local_ordering f ord {f 0}" + unfolding local_ordering_def by simp + +lemma singleton_ordering_loc [simp]: + "\f. local_ordering f ord {a}" + using singleton_ordered_loc by fast + +lemma two_ordered_loc: + assumes "a = f 0" and "b = f 1" + shows "local_ordering f ord {a, b}" +proof cases + assume "a = b" + thus ?thesis using assms singleton_ordered_loc by (metis insert_absorb2) +next + assume a_neq_b: "a \ b" + hence "(\n. (finite {a,b} \ n < card {a,b}) \ f n \ {a,b})" + using assms by (metis One_nat_def card.infinite card_2_iff fact_0 fact_2 insert_iff less_2_cases_iff) + moreover have "(\x\{a,b}. \n. (finite {a,b} \ n < card {a,b}) \ f n = x)" + using assms a_neq_b all_not_in_conv card_Suc_eq card_0_eq card_gt_0_iff insert_iff lessI by auto + moreover have "(\n. (finite {a,b} \ Suc (Suc n) < card {a,b}) + \ ord (f n) (f (Suc n)) (f (Suc (Suc n))))" + using a_neq_b by auto + ultimately have "local_ordering f ord {a, b}" + using local_ordering_def by blast + thus ?thesis by auto +qed + +lemma two_ordering_loc [simp]: + "\f. local_ordering f ord {a, b}" + using total_implies_local two_ordering by fastforce + +lemma card_le2_ordering_loc: + assumes finiteX: "finite X" + and card_le2: "card X \ 2" + shows "\f. local_ordering f ord X" + using assms total_implies_local card_le2_ordering by metis + +lemma ord_ordered_loc: + assumes abc: "ord a b c" + and abc_neq: "a \ b \ a \ c \ b \ c" + shows "\f. local_ordering f ord {a,b,c}" + using assms total_implies_local ord_ordered by metis + +lemma overlap_ordering_loc: + assumes abc: "ord a b c" + and bcd: "ord b c d" + and abd: "ord a b d" + and acd: "ord a c d" + and abc_neq: "a \ b \ a \ c \ a \ d \ b \ c \ b \ d \ c \ d" + shows "\f. local_ordering f ord {a,b,c,d}" + using overlap_ordering[OF assms] total_implies_local by blast + +lemma ordering_sym_loc: + assumes ord_sym: "\a b c. ord a b c \ ord c b a" + and "finite X" + and "local_ordering f ord X" + shows "local_ordering (\n. f (card X - 1 - n)) ord X" + unfolding local_ordering_def using assms(2) apply auto + apply (metis local_ordering_def assms(3) card_0_eq card_gt_0_iff diff_Suc_less gr_implies_not0) +proof - + fix x + assume "finite X" + assume "x \ X" + obtain n where "finite X \ n < card X" and "f n = x" + by (metis local_ordering_def \x \ X\ assms(3)) + have "f (card X - ((card X - 1 - n) + 1)) = x" + by (simp add: Suc_leI \f n = x\ \finite X \ n < card X\ assms(2)) + thus "\nx \ X\ add.commute assms(2) card_Diff_singleton card_Suc_Diff1 diff_less_Suc plus_1_eq_Suc) +next + fix n + let ?n1 = "Suc n" + let ?n2 = "Suc ?n1" + assume "finite X" + assume "Suc (Suc n) < card X" + have "ord (f (card X - Suc ?n2)) (f (card X - Suc ?n1)) (f (card X - Suc n))" + using assms(3) unfolding local_ordering_def + using \Suc (Suc n) < card X\ by (metis + Suc_diff_Suc Suc_lessD card_eq_0_iff card_gt_0_iff diff_less gr_implies_not0 zero_less_Suc) + thus " ord (f (card X - Suc n)) (f (card X - Suc ?n1)) (f (card X - Suc ?n2))" + using ord_sym by blast +qed + +lemma zero_into_ordering_loc: + assumes "local_ordering f betw X" + and "X \ {}" + shows "(f 0) \ X" + using local_ordering_def by (metis assms card_eq_0_iff gr_implies_not0 linorder_neqE_nat) + end \ No newline at end of file diff --git a/thys/Schutz_Spacetime/Util.thy b/thys/Schutz_Spacetime/Util.thy --- a/thys/Schutz_Spacetime/Util.thy +++ b/thys/Schutz_Spacetime/Util.thy @@ -1,56 +1,129 @@ -(* Title: Schutz_Spacetime/Util.thy - Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot - University of Edinburgh, 2021 -*) -theory Util - imports Main - -(* Some "utility" proofs -- little proofs that come in handy every now and then. *) - -begin - -text \ - We need this in order to obtain a natural number which can be passed to the ordering function, - distinct from two others, in the case of a finite set of events with cardinality a least 3. -\ - -lemma is_free_nat: - assumes "(m::nat) < n" - and "n < c" - and "c \ 3" - shows "\k::nat. k < m \ (m < k \ k < n) \ (n < k \ k < c)" -using assms by presburger - -text \Helpful proofs on sets.\ - -lemma set_le_two [simp]: "card {a, b} \ 2" - by (simp add: card_insert_if) - -lemma set_le_three [simp]: "card {a, b, c} \ 3" - by (simp add: card_insert_if) - -lemma card_subset: "\card Y = n; Y \ X\ \ card X \ n \ infinite X" - using card_mono by blast - -lemma card_subset_finite: "\finite X; card Y = n; Y \ X\ \ card X \ n" - using card_subset by auto - -lemma three_subset: "\x \ y; x \ z; y \ z; {x,y,z} \ X\ \ card X \ 3 \ infinite X" - apply (case_tac "finite X") - apply (auto simp : card_mono) - apply (erule_tac Y = "{x,y,z}" in card_subset_finite) - by auto - -lemma card_Collect_nat: - assumes "(j::nat)>i" - shows "card {i..j} = j-i+1" - using card_atLeastAtMost - using Suc_diff_le assms le_eq_less_or_eq by presburger - - -text \These lemmas make life easier with some of the ordering proofs.\ - -lemma less_3_cases: "n < 3 \ n = 0 \ n = Suc 0 \ n = Suc (Suc 0)" -by auto - +(* Title: Schutz_Spacetime/Util.thy + Authors: Richard Schmoetten, Jake Palmer and Jacques D. Fleuriot + University of Edinburgh, 2021 +*) +theory Util +imports Main + +begin + +text \Some "utility" proofs -- little proofs that come in handy every now and then.\ + +text \ + We need this in order to obtain a natural number which can be passed to the ordering function, + distinct from two others, in the case of a finite set of events with cardinality a least 3. +\ + +lemma is_free_nat: + assumes "(m::nat) < n" + and "n < c" + and "c \ 3" + shows "\k::nat. k < m \ (m < k \ k < n) \ (n < k \ k < c)" +using assms by presburger + +text \Helpful proofs on sets.\ + +lemma set_le_two [simp]: "card {a, b} \ 2" + by (simp add: card_insert_if) + +lemma set_le_three [simp]: "card {a, b, c} \ 3" + by (simp add: card_insert_if) + +lemma card_subset: "\card Y = n; Y \ X\ \ card X \ n \ infinite X" + using card_mono by blast + +lemma card_subset_finite: "\finite X; card Y = n; Y \ X\ \ card X \ n" + using card_subset by auto + +lemma three_subset: "\x \ y; x \ z; y \ z; {x,y,z} \ X\ \ card X \ 3 \ infinite X" + apply (case_tac "finite X") + apply (auto simp : card_mono) + apply (erule_tac Y = "{x,y,z}" in card_subset_finite) + by auto + +lemma three_in_set3: + assumes "card X \ 3" + obtains x y z where "x\X" and "y\X" and "z\X" and "x\y" and "x\z" and "y\z" + using assms by (auto simp add: card_le_Suc_iff numeral_3_eq_3) + +lemma card_Collect_nat: + assumes "(j::nat)>i" + shows "card {i..j} = j-i+1" + using card_atLeastAtMost + using Suc_diff_le assms le_eq_less_or_eq by presburger + +lemma inf_3_elms: assumes "infinite X" shows "(\x\X. \y\X. \z\X. x \ y \ y \ z \ x \ z)" +proof - + obtain x y where 1: "x\X" "y\X" "y\x" + by (metis assms finite.emptyI finite.insertI rev_finite_subset singleton_iff subsetI) + have "infinite (X-{x,y})" + using infinite_remove by (simp add: assms) + then obtain z where 2: "z\X" "x\z" "z\y" + using infinite_imp_nonempty by (metis Diff_eq_empty_iff insertCI subset_eq) + show ?thesis using 1 2 by blast +qed + +lemma card_3_dist: "card {x,y,z} = 3 \ x\y \ x\z \ y\z" + by (simp add: eval_nat_numeral card_insert_if) + +lemma card_3_eq: + "card X = 3 \ (\x y z. X={x,y,z} \ x \ y \ y \ z \ x \ z)" + (is "card X = 3 \ ?card3 X") +proof + assume asm: "card X = 3" hence "card X \ 3" by simp + then obtain x y z where "x \ y \ y \ z \ x \ z" "{x,y,z} \ X" + apply (simp add: eval_nat_numeral) + by (auto simp add: card_le_Suc_iff) + thus "?card3 X" + using Finite_Set.card_subset_eq \card X = 3\ + apply (simp add: eval_nat_numeral) + by (smt (verit, ccfv_threshold) \{x, y, z} \ X\ card.empty card.infinite card_insert_if + card_subset_eq empty_iff finite.emptyI insertE nat.distinct(1)) +next + show "?card3 X \ card X = 3" + by (smt (z3) card.empty card.insert eval_nat_numeral(2) finite.intros(1) finite_insert insertE + insert_absorb insert_not_empty numeral_3_eq_3 semiring_norm(26,27)) +qed + + +lemma card_3_eq': + "\card X = 3; card {a,b,c} = 3; {a,b,c} \X\ \ X = {a,b,c}" + "\card X = 3; a \ X; b \ X; c \ X; a \ b; a \ c; b \ c\ \ X = {a,b,c}" +proof - + show "\card X = 3; card {a,b,c} = 3; {a,b,c} \X\ \ X = {a,b,c}" + by (metis card.infinite card_subset_eq zero_neq_numeral) + thus "\card X = 3; a \ X; b \ X; c \ X; a \ b; a \ c; b \ c\ \ X = {a,b,c}" + by (meson card_3_dist empty_subsetI insert_subset) +qed + +lemma card_4_eq: + "card X = 4 \ (\S\<^sub>1. \S\<^sub>2. \S\<^sub>3. \S\<^sub>4. X = {S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4} \ + S\<^sub>1 \ S\<^sub>2 \ S\<^sub>1 \ S\<^sub>3 \ S\<^sub>1 \ S\<^sub>4 \ S\<^sub>2 \ S\<^sub>3 \ S\<^sub>2 \ S\<^sub>4 \ S\<^sub>3 \ S\<^sub>4)" + (is "card X = 4 \ ?card4 X") +proof + assume "card X = 4" + hence "card X \ 4" by auto + then obtain S\<^sub>1 S\<^sub>2 S\<^sub>3 S\<^sub>4 where + 0: "S\<^sub>1\X \ S\<^sub>2\X \ S\<^sub>3\X \ S\<^sub>4\X" and + 1: "S\<^sub>1 \ S\<^sub>2 \ S\<^sub>1 \ S\<^sub>3 \ S\<^sub>1 \ S\<^sub>4 \ S\<^sub>2 \ S\<^sub>3 \ S\<^sub>2 \ S\<^sub>4 \ S\<^sub>3 \ S\<^sub>4" + apply (simp add: eval_nat_numeral) + by (auto simp add: card_le_Suc_iff) + then have 2: "{S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4} \ X" "card {S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4} = 4" by auto + have "X = {S\<^sub>1, S\<^sub>2, S\<^sub>3, S\<^sub>4}" + using Finite_Set.card_subset_eq \card X = 4\ + apply (simp add: eval_nat_numeral) + by (smt (z3) \card X = 4\ 2 card.infinite card_subset_eq nat.distinct(1)) + thus "?card4 X" using 1 by blast +next + show "?card4 X \ card X = 4" + by (smt (z3) card.empty card.insert eval_nat_numeral(2) finite.intros(1) finite_insert insertE + insert_absorb insert_not_empty numeral_3_eq_3 semiring_norm(26,27)) +qed + + +text \These lemmas make life easier with some of the ordering proofs.\ + +lemma less_3_cases: "n < 3 \ n = 0 \ n = Suc 0 \ n = Suc (Suc 0)" + by auto + end \ No newline at end of file diff --git a/thys/Schutz_Spacetime/document/root.bib b/thys/Schutz_Spacetime/document/root.bib --- a/thys/Schutz_Spacetime/document/root.bib +++ b/thys/Schutz_Spacetime/document/root.bib @@ -1,11 +1,11 @@ -@book{schutz1997, - title = {Independent {{Axioms}} for {{Minkowski Space}}-{{Time}}}, - author = {Schutz, John W.}, - year = {1997}, - month = oct, - publisher = {{CRC Press}}, - abstract = {The primary aim of this monograph is to clarify the undefined primitive concepts and the axioms which form the basis of Einstein's theory of special relativity. Minkowski space-time is developed from a set of independent axioms, stated in terms of a single relation of betweenness. It is shown that all models are isomorphic to the usual coordinate model, and the axioms are consistent relative to the reals.}, - isbn = {978-0-582-31760-4}, - keywords = {Science / Physics / Mathematical \& Computational}, - language = {en} -} +@book{schutz1997, + title = {Independent {{Axioms}} for {{Minkowski Space}}-{{Time}}}, + author = {Schutz, John W.}, + year = {1997}, + month = oct, + publisher = {{CRC Press}}, + abstract = {The primary aim of this monograph is to clarify the undefined primitive concepts and the axioms which form the basis of Einstein's theory of special relativity. Minkowski space-time is developed from a set of independent axioms, stated in terms of a single relation of betweenness. It is shown that all models are isomorphic to the usual coordinate model, and the axioms are consistent relative to the reals.}, + isbn = {978-0-582-31760-4}, + keywords = {Science / Physics / Mathematical \& Computational}, + language = {en} +} diff --git a/thys/Schutz_Spacetime/document/root.tex b/thys/Schutz_Spacetime/document/root.tex --- a/thys/Schutz_Spacetime/document/root.tex +++ b/thys/Schutz_Spacetime/document/root.tex @@ -1,76 +1,76 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym} - -% further packages required for unusual symbols (see also -% isabellesym.sty), use only when needed - -\usepackage{amssymb} - %for \, \, \, \, \, \, - %\, \, \, \, \, - %\, \, \ - -%\usepackage{eurosym} - %for \ - -%\usepackage[only,bigsqcap]{stmaryrd} - %for \ - -%\usepackage{eufrak} - %for \ ... \, \ ... \ (also included in amssymb) - -%\usepackage{textcomp} - %for \, \, \, \, \, - %\ - -% this should be the last package used -\usepackage{pdfsetup} - -% urls in roman style, theory text in math-similar italics -\urlstyle{rm} -\isabellestyle{it} - -% for uniform font size -%\renewcommand{\isastyle}{\isastyleminor} - -% shortcuts for event and path sets -\newcommand{\E}{\mathcal{E}} -\renewcommand{\P}{\mathcal{P}} -\newcommand{\spray}{\ensuremath{\textrm{SPRAY}}} -\newcommand{\ord}[3]{\ensuremath{[#1\;#2\;#3]}} - - -\begin{document} - -\title{Geometric Axioms for Minkowski Spacetime} -\author{Richard Schmoetten, Jake Palmer, Jacques Fleuriot} -\maketitle -\begin{abstract} -This is a formalisation of Schutz' system of axioms for Minkowski spacetime \cite{schutz1997}, as well as the results in his third chapter (``Temporal Order on a Path''), with the exception of the second part of Theorem 12. Many results are proven here that cannot be found in Schutz, either preceding the theorem they are needed for, or in their own thematic section. -\end{abstract} - -\tableofcontents\newpage - -% sane default for proof documents -\parindent 0pt\parskip 0.5ex - -% generated text of all theories -%\input{session} -\input{Util.tex} % Should be empty -%\chapter{Ternary Ordering} -\input{TernaryOrdering.tex}\newpage -%\chapter{Primitives and Axioms} -\input{Minkowski.tex}\newpage -%\chapter{Temporal Order on a Path} -\input{TemporalOrderOnPath.tex} - - -% optional bibliography -\bibliographystyle{abbrv} -\bibliography{root} - -\end{document} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + +% shortcuts for event and path sets +\newcommand{\E}{\mathcal{E}} +\renewcommand{\P}{\mathcal{P}} +\newcommand{\spray}{\ensuremath{\textrm{SPRAY}}} +\newcommand{\ord}[3]{\ensuremath{[#1\;#2\;#3]}} + + +\begin{document} + +\title{Geometric Axioms for Minkowski Spacetime} +\author{Richard Schmoetten, Jake Palmer, Jacques Fleuriot} +\maketitle +\begin{abstract} +This is a formalisation of Schutz' system of axioms for Minkowski spacetime \cite{schutz1997}, as well as the results in his third chapter (``Temporal Order on a Path''), with the exception of the second part of Theorem 12. Many results are proven here that cannot be found in Schutz, either preceding the theorem they are needed for, or in their own thematic section. +\end{abstract} + +\tableofcontents\newpage + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +%\input{session} +\input{Util.tex} % Should be empty +%\chapter{Ternary Ordering} +\input{TernaryOrdering.tex}\newpage +%\chapter{Primitives and Axioms} +\input{Minkowski.tex}\newpage +%\chapter{Temporal Order on a Path} +\input{TemporalOrderOnPath.tex} + + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: