diff --git a/thys/Chandy_Lamport/Distributed_System.thy b/thys/Chandy_Lamport/Distributed_System.thy --- a/thys/Chandy_Lamport/Distributed_System.thy +++ b/thys/Chandy_Lamport/Distributed_System.thy @@ -1,1829 +1,1829 @@ section \Modelling distributed systems\ text \We assume familiarity with Chandy and Lamport's paper \emph{Distributed Snapshots: Determining Global States of Distributed Systems}~\<^cite>\"chandy"\.\ theory Distributed_System imports Main begin type_synonym 'a fifo = "'a list" type_synonym channel_id = nat datatype 'm message = Marker | Msg 'm datatype recording_state = NotStarted | Recording | Done text \We characterize distributed systems by three underlying type variables: Type variable 'p captures the processes of the underlying system. Type variable 's describes the possible states of the processes. Finally, type variable 'm describes all possible messages in said system. Each process is in exactly one state at any point in time of the system. Processes are interconnected by directed channels, which hold messages in-flight between connected processes. There can be an arbitrary number of channels between different processes. The entire state of the system including the (potentially unfinished) snapshot state is called \emph{configuration}.\ record ('p, 's, 'm) configuration = states :: "'p \ 's" msgs :: "channel_id \ 'm message fifo" process_snapshot :: "'p \ 's option" channel_snapshot :: "channel_id \ 'm fifo * recording_state" text \An event in Chandy and Lamport's formalization describes a process' state transition, optionally producing or consuming (but not both) a message on a channel. Additionally, a process may either initiate a snapshot spontaneously, or is forced to do so by receiving a snapshot \emph{marker} on one of it's incoming channels.\ datatype ('p, 's, 'm) event = isTrans: Trans (occurs_on: 'p) 's 's | isSend: Send (getId: channel_id) (occurs_on: 'p) (partner: 'p) 's 's (getMsg: 'm) | isRecv: Recv (getId: channel_id) (occurs_on: 'p) (partner: 'p) 's 's (getMsg: 'm) | isSnapshot: Snapshot (occurs_on: 'p) | isRecvMarker: RecvMarker (getId: channel_id) (occurs_on: 'p) (partner: 'p) text \We introduce abbreviations and type synoyms for commonly used terms.\ type_synonym ('p, 's, 'm) trace = "('p, 's, 'm) event list" abbreviation ps where "ps \ process_snapshot" abbreviation cs where "cs \ channel_snapshot" abbreviation no_snapshot_change where "no_snapshot_change c c' \ ((\p'. ps c p' = ps c' p') \ (\i'. cs c i' = cs c' i'))" abbreviation has_snapshotted where "has_snapshotted c p \ process_snapshot c p \ None" text \A regular event is an event as described in Chandy and Lamport's original paper: A state transition accompanied by the emission or receiving of a message. Nonregular events are related to snapshotting and receiving markers along communication channels.\ definition regular_event[simp]: "regular_event ev \ (isTrans ev \ isSend ev \ isRecv ev)" lemma nonregular_event: "~ regular_event ev = (isSnapshot ev \ isRecvMarker ev)" by (meson event.distinct_disc event.exhaust_disc regular_event) lemma event_occurs_on_unique: assumes "p \ q" "occurs_on ev = p" shows "occurs_on ev \ q" using assms by (cases ev, auto) subsection \The distributed system locale\ text \In order to capture Chandy and Lamport's computation system we introduce two locales. The distributed system locale describes global truths, such as the mapping from channel IDs to sender and receiver processes, the transition relations for the underlying computation system and the core assumption that no process has a channel to itself. While not explicitly mentioned in Chandy's and Lamport's work, it makes sense to assume that a channel need not communicate to itself via messages, since it shares memory with itself.\ locale distributed_system = fixes channel :: "channel_id \ ('p * 'p) option" and trans :: "'p \ 's \ 's \ bool" and send :: "channel_id \ 'p \ 'p \ 's \ 's \ 'm \ bool" and recv :: "channel_id \ 'p \ 'p \ 's \ 's \ 'm \ bool" assumes no_self_channel: "\i. \p. channel i = Some (p, p)" begin subsubsection \State transitions\ definition can_occur :: "('p, 's, 'm) event \ ('p, 's, 'm) configuration \ bool" where "can_occur ev c \ (case ev of Trans p s s' \ states c p = s \ trans p s s' | Send i p q s s' msg \ states c p = s \ channel i = Some (p, q) \ send i p q s s' msg | Recv i p q s s' msg \ states c p = s \ channel i = Some (q, p) \ length (msgs c i) > 0 \ hd (msgs c i) = Msg msg \ recv i p q s s' msg | Snapshot p \ \ has_snapshotted c p | RecvMarker i p q \ channel i = Some (q, p) \ length (msgs c i) > 0 \ hd (msgs c i) = Marker)" definition src where "src i p \ (\q. channel i = Some (p, q))" definition dest where "dest i q \ (\p. channel i = Some (p, q))" lemma can_occur_Recv: assumes "can_occur (Recv i p q s s' m) c" shows "states c p = s \ channel i = Some (q, p) \ (\xs. msgs c i = Msg m # xs) \ recv i p q s s' m" proof - have "\xs. msgs c i = Msg m # xs" using assms can_occur_def by (metis (mono_tags, lifting) event.case(3) hd_Cons_tl length_greater_0_conv) then show ?thesis using assms can_occur_def by auto qed abbreviation check_snapshot_occur where "check_snapshot_occur c c' p \ (can_occur (Snapshot p) c \ (ps c' p = Some (states c p)) \ (\p'. states c p' = states c' p') \ (\p'. (p' \ p) \ ps c' p' = ps c p') \ (\i. (\q. channel i = Some (p, q)) \ msgs c' i = msgs c i @ [Marker]) \ (\i. (\q. channel i = Some (q, p)) \ channel_snapshot c' i = (fst (channel_snapshot c i), Recording)) \ (\i. (\q. channel i = Some (p, q)) \ msgs c' i = msgs c i) \ (\i. (\q. channel i = Some (q, p)) \ channel_snapshot c' i = channel_snapshot c i))" abbreviation check_recv_marker_occur where "check_recv_marker_occur c c' i p q \ (can_occur (RecvMarker i p q) c \ (\r. states c r = states c' r) \ (\r. (r \ p) \ process_snapshot c r = process_snapshot c' r) \ (Marker # msgs c' i = msgs c i) \ (channel_snapshot c' i = (fst (channel_snapshot c i), Done)) \ (if has_snapshotted c p then (process_snapshot c p = process_snapshot c' p) \ (\i'. (i' \ i) \ msgs c' i' = msgs c i') \ (\i'. (i' \ i) \ channel_snapshot c i' = channel_snapshot c' i') else (process_snapshot c' p = Some (states c p)) \ (\i'. i' \ i \ (\r. channel i' = Some (p, r)) \ msgs c' i' = msgs c i' @ [Marker]) \ (\i'. i' \ i \ (\r. channel i' = Some (r, p)) \ channel_snapshot c' i' = (fst (channel_snapshot c i'), Recording)) \ (\i'. i' \ i \ (\r. channel i' = Some (p, r)) \ msgs c' i' = msgs c i') \ (\i'. i' \ i \ (\r. channel i' = Some (r, p)) \ channel_snapshot c' i' = channel_snapshot c i')))" abbreviation check_trans_occur where "check_trans_occur c c' p s s'\ (can_occur (Trans p s s') c \ (states c' p = s') \ (\r. (r \ p) \ states c' r = states c r) \ (\i. msgs c' i = msgs c i) \ (no_snapshot_change c c'))" abbreviation check_send_occur where "check_send_occur c c' i p q s s' msg \ (can_occur (Send i p q s s' msg) c \ (states c' p = s') \ (\r. (r \ p) \ states c' r = states c r) \ (msgs c' i = msgs c i @ [Msg msg]) \ (\i'. i \ i' \ msgs c' i' = msgs c i') \ (no_snapshot_change c c'))" abbreviation check_recv_occur where "check_recv_occur c c' i p q s s' msg \ (can_occur (Recv i p q s s' msg) c \ (states c p = s \ states c' p = s') \ (\r. (r \ p) \ states c' r = states c r) \ (msgs c i = Msg msg # msgs c' i) \ (\i'. i \ i' \ msgs c' i' = msgs c i') \ (\r. process_snapshot c r = process_snapshot c' r) \ (\i'. i' \ i \ channel_snapshot c i' = channel_snapshot c' i') \ (if snd (channel_snapshot c i) = Recording then channel_snapshot c' i = (fst (channel_snapshot c i) @ [msg], Recording) else channel_snapshot c i = channel_snapshot c' i))" text \The \emph{next} predicate lets us express configuration transitions using events. The predicate $next(s_1, e, s_2)$ denotes the transition of the configuration $s_1$ to $s_2$ via the event $e$. It ensures that $e$ can occur in state $s_1$ and the state $s_2$ is correctly constructed from $s_1$.\ primrec "next" :: "('p, 's, 'm) configuration \ ('p, 's, 'm) event \ ('p, 's, 'm) configuration \ bool" ("_ \ _ \ _" [70, 70, 70]) where next_snapshot: "c \ Snapshot p \ c' = check_snapshot_occur c c' p" | next_recv_marker: "c \ RecvMarker i p q \ c' = check_recv_marker_occur c c' i p q" | next_trans: "c \ Trans p s s' \ c' = check_trans_occur c c' p s s'" | next_send: "c \ Send i p q s s' msg \ c' = check_send_occur c c' i p q s s' msg" | next_recv: "c \ Recv i p q s s' msg \ c' = check_recv_occur c c' i p q s s' msg" text \Useful lemmas about state transitions\ lemma state_and_event_determine_next: assumes "c \ ev \ c'" and "c \ ev \ c''" shows "c' = c''" proof (cases ev) case (Snapshot p) then have "states c' = states c''" using assms by auto moreover have "msgs c' = msgs c''" proof (rule ext) fix i show "msgs c' i = msgs c'' i" proof (cases "channel i = None") case True then show ?thesis using Snapshot assms by auto next case False then obtain r s where "channel i = Some (r, s)" by auto with assms Snapshot show ?thesis by (cases "r = p", simp_all) qed qed moreover have "process_snapshot c' = process_snapshot c''" by (metis Snapshot assms next_snapshot ext) moreover have "channel_snapshot c' = channel_snapshot c''" proof (rule ext) fix i show "channel_snapshot c' i = channel_snapshot c'' i" proof (cases "channel i = None") case True then show ?thesis using assms Snapshot by simp next case False then obtain r s where "channel i = Some (r, s)" by auto with assms Snapshot show ?thesis by (cases "s = p", simp_all) qed qed ultimately show "c' = c''" by simp next case (RecvMarker i p) then have "states c' = states c''" using assms by auto moreover have "msgs c' = msgs c''" proof (rule ext) fix i' show "msgs c' i' = msgs c'' i'" proof (cases "i' = i") case True then have "Marker # msgs c' i' = msgs c i'" using assms RecvMarker by simp also have "... = Marker # msgs c'' i'" using assms RecvMarker \i' = i\ by simp finally show ?thesis by simp next case False then show ?thesis proof (cases "has_snapshotted c p") case True then show ?thesis using assms RecvMarker \i' \ i\ by simp next case no_snap: False then show ?thesis proof (cases "channel i' = None") case True then show ?thesis using assms RecvMarker \i' \ i\ no_snap by simp next case False then obtain r s where "channel i' = Some (r, s)" by auto with assms RecvMarker no_snap \i' \ i\ show ?thesis by (cases "r = p"; simp_all) qed qed qed qed moreover have "process_snapshot c' = process_snapshot c''" proof (rule ext) fix r show "ps c' r = ps c'' r" proof (cases "r \ p") case True then show ?thesis using assms RecvMarker by simp next case False with assms RecvMarker \~ r \ p\ show ?thesis by (cases "has_snapshotted c r", auto) qed qed moreover have "channel_snapshot c' = channel_snapshot c''" proof (rule ext) fix i' show "cs c' i' = cs c'' i'" proof (cases "i' = i") case True then show ?thesis using assms RecvMarker by simp next case False then show ?thesis proof (cases "has_snapshotted c p") case True then show ?thesis using assms RecvMarker \i' \ i\ by simp next case no_snap: False then show ?thesis proof (cases "channel i' = None") case True then show ?thesis using assms RecvMarker \i' \ i\ no_snap by simp next case False then obtain r s where "channel i' = Some (r, s)" by auto with assms RecvMarker no_snap \i' \ i\ show ?thesis by (cases "s = p"; simp_all) qed qed qed qed ultimately show "c' = c''" by simp next case (Trans p s s') then have "states c' = states c''" by (metis (no_types, lifting) assms next_trans ext) moreover have "msgs c' = msgs c''" using assms Trans by auto moreover have "process_snapshot c' = process_snapshot c''" using assms Trans by auto moreover have "channel_snapshot c' = channel_snapshot c''" using assms Trans by auto ultimately show "c' = c''" by simp next case (Send i p s s' m) then have "states c' = states c''" by (metis (no_types, lifting) assms next_send ext) moreover have "msgs c' = msgs c''" proof (rule ext) fix i' from assms Send show "msgs c' i' = msgs c'' i'" by (cases "i' = i", simp_all) qed moreover have "process_snapshot c' = process_snapshot c''" using assms Send by auto moreover have "channel_snapshot c' = channel_snapshot c''" using assms Send by auto ultimately show "c' = c''" by simp next case (Recv i p s s' m) then have "states c' = states c''" by (metis (no_types, lifting) assms next_recv ext) moreover have "msgs c' = msgs c''" proof (rule ext) fix i' from assms Recv show "msgs c' i' = msgs c'' i'" by (cases "i' = i", simp_all) qed moreover have "process_snapshot c' = process_snapshot c''" using assms Recv by auto moreover have "channel_snapshot c' = channel_snapshot c''" proof (rule ext) fix i' show "cs c' i' = cs c'' i'" proof (cases "i' \ i") case True then show ?thesis using assms Recv by simp next case False with assms Recv show ?thesis by (cases "snd (cs c i') = Recording", auto) qed qed ultimately show "c' = c''" by simp qed lemma exists_next_if_can_occur: assumes "can_occur ev c" shows "\c'. c \ ev \ c'" proof (cases ev) case (Snapshot p) let ?c = "\ states = states c, msgs = %i. if (\q. channel i = Some (p, q)) then msgs c i @ [Marker] else msgs c i, process_snapshot = %r. if r = p then Some (states c p) else ps c r, channel_snapshot = %i. if (\q. channel i = Some (q, p)) then (fst (cs c i), Recording) else cs c i \" have "c \ ev \ ?c" using Snapshot assms by auto then show ?thesis by blast next case (RecvMarker i p q) show ?thesis proof (cases "has_snapshotted c p") case True let ?c = "\ states = states c, msgs = %i'. if i = i' then tl (msgs c i') else msgs c i', process_snapshot = ps c, channel_snapshot = %i'. if i = i' then (fst (cs c i'), Done) else cs c i' \" have "msgs c i = Marker # msgs ?c i" using assms can_occur_def RecvMarker hd_Cons_tl by fastforce then have "c \ ev \ ?c" using True RecvMarker assms by auto then show ?thesis by blast next case False let ?c = "\ states = states c, msgs = %i'. if i' = i then tl (msgs c i') else if (\r. channel i' = Some (p, r)) then msgs c i' @ [Marker] else msgs c i', process_snapshot = %r. if r = p then Some (states c r) else ps c r, channel_snapshot = %i'. if i = i' then (fst (cs c i'), Done) else if (\r. channel i' = Some (r, p)) then (fst (cs c i'), Recording) else cs c i' \" have "msgs c i = Marker # msgs ?c i" using assms can_occur_def RecvMarker hd_Cons_tl by fastforce moreover have "ps ?c p = Some (states c p)" by simp ultimately have "c \ ev \ ?c" using RecvMarker assms False by auto then show ?thesis by blast qed next case (Trans p s s') let ?c = "\ states = %r. if r = p then s' else states c r, msgs = msgs c, process_snapshot = ps c, channel_snapshot = cs c \" have "c \ ev \ ?c" using Trans assms by auto then show ?thesis by blast next case (Send i p q s s' msg) let ?c = "\ states = %r. if r = p then s' else states c r, msgs = %i'. if i = i' then msgs c i' @ [Msg msg] else msgs c i', process_snapshot = ps c, channel_snapshot = cs c \" have "c \ ev \ ?c" using Send assms by auto then show ?thesis by blast next case (Recv i p q s s' msg) then show ?thesis proof (cases "snd (cs c i)") case Recording let ?c = "\ states = %r. if r = p then s' else states c r, msgs = %i'. if i = i' then tl (msgs c i') else msgs c i', process_snapshot = ps c, channel_snapshot = %i'. if i = i' then (fst (cs c i') @ [msg], Recording) else cs c i'\" have "c \ ev \ ?c" using Recv Recording assms can_occur_Recv by fastforce then show ?thesis by blast next case Done let ?c = "\ states = %r. if r = p then s' else states c r, msgs = %i'. if i = i' then tl (msgs c i') else msgs c i', process_snapshot = ps c, channel_snapshot = cs c \" have "c \ ev \ ?c" using Done Recv assms can_occur_Recv by fastforce then show ?thesis by blast next case NotStarted let ?c = "\ states = %r. if r = p then s' else states c r, msgs = %i'. if i = i' then tl (msgs c i') else msgs c i', process_snapshot = ps c, channel_snapshot = cs c \" have "c \ ev \ ?c" using NotStarted Recv assms can_occur_Recv by fastforce then show ?thesis by blast qed qed lemma exists_exactly_one_following_state: "can_occur ev c \ \!c'. c \ ev \ c'" using exists_next_if_can_occur state_and_event_determine_next by blast lemma no_state_change_if_no_event: assumes "c \ ev \ c'" and "occurs_on ev \ p" shows "states c p = states c' p \ process_snapshot c p = process_snapshot c' p" using assms by (cases ev, auto) lemma no_msgs_change_if_no_channel: assumes "c \ ev \ c'" and "channel i = None" shows "msgs c i = msgs c' i" using assms proof (cases ev) case (RecvMarker cid p) then have "cid \ i" using assms RecvMarker can_occur_def by fastforce with assms RecvMarker show ?thesis by (cases "has_snapshotted c p", auto) next case (Send cid p s s' m) then have "cid \ i" using assms Send can_occur_def by fastforce then show ?thesis using assms Send by auto next case (Recv cid p s s' m) then have "cid \ i" using assms Recv can_occur_def by fastforce then show ?thesis using assms Recv by simp qed simp_all lemma no_cs_change_if_no_channel: assumes "c \ ev \ c'" and "channel i = None" shows "cs c i = cs c' i" using assms proof (cases ev) case (RecvMarker cid p) then have "cid \ i" using assms RecvMarker can_occur_def by fastforce with assms RecvMarker show ?thesis by (cases "has_snapshotted c p", auto) next case (Send cid p s s' m) then have "cid \ i" using assms Send can_occur_def by fastforce then show ?thesis using assms Send by auto next case (Recv cid p s s' m) then have "cid \ i" using assms Recv can_occur_def by fastforce then show ?thesis using assms Recv by simp qed simp_all lemma no_msg_change_if_no_event: assumes "c \ ev \ c'" and "isSend ev \ getId ev \ i" and "isRecv ev \ getId ev \ i" and "regular_event ev" shows "msgs c i = msgs c' i" proof (cases "channel i = None") case True then show ?thesis using assms no_msgs_change_if_no_channel by simp next have "isTrans ev \ isSend ev \ isRecv ev" using assms by simp then show ?thesis proof (elim disjE) assume "isTrans ev" then show ?thesis by (metis assms(1) event.collapse(1) next_trans) next assume "isSend ev" then obtain i' r s u u' m where Send: "ev = Send i' r s u u' m" by (meson isSend_def) then show ?thesis using Send assms by auto next assume "isRecv ev" then obtain i' r s u u' m where "ev = Recv i' r s u u' m" by (meson isRecv_def) then show ?thesis using assms by auto qed qed lemma no_cs_change_if_no_event: assumes "c \ ev \ c'" and "isRecv ev \ getId ev \ i" and "regular_event ev" shows "cs c i = cs c' i" proof - have "isTrans ev \ isSend ev \ isRecv ev" using assms by simp then show ?thesis proof (elim disjE) assume "isTrans ev" then show ?thesis by (metis assms(1) event.collapse(1) next_trans) next assume "isSend ev" then obtain i' r s u u' m where "ev = Send i' r s u u' m" by (meson isSend_def) then show ?thesis using assms by auto next assume "isRecv ev" then obtain i r s u u' m where "ev = Recv i r s u u' m" by (meson isRecv_def) then show ?thesis using assms by auto qed qed lemma happen_implies_can_occur: assumes "c \ ev \ c'" shows "can_occur ev c" proof - show ?thesis using assms by (cases ev, auto) qed lemma snapshot_increases_message_length: assumes "ev = Snapshot p" and "c \ ev \ c'" and "channel i = Some (q, r)" shows "length (msgs c i) \ length (msgs c' i)" using assms by (cases "p = q", auto) lemma recv_marker_changes_head_only_at_i: assumes "ev = RecvMarker i p q" and "c \ ev \ c'" and "i' \ i" shows "msgs c i' = [] \ hd (msgs c i') = hd (msgs c' i')" proof (cases "channel i' = None") case True then show ?thesis using assms no_msgs_change_if_no_channel by presburger next case False then show ?thesis proof (cases "msgs c i'") case Nil then show ?thesis by simp next case (Cons m xs) then obtain r s where "channel i' = Some (r, s)" using False by auto then show ?thesis proof (cases "has_snapshotted c p") case True then show ?thesis using assms by auto next case False with assms show ?thesis by (cases "r = p", auto) qed qed qed lemma recv_marker_other_channels_not_shrinking: assumes "ev = RecvMarker i p q" and "c \ ev \ c'" shows "length (msgs c i') \ length (msgs c' i') \ i \ i'" proof (rule iffI) show "length (msgs c i') \ length (msgs c' i') \ i \ i'" proof (rule ccontr) assume asm: "~ i \ i'" "length (msgs c i') \ length (msgs c' i')" then have "msgs c i = Marker # msgs c' i" using assms by auto then have "length (msgs c i) > length (msgs c' i)" by simp then have "length (msgs c i') > length (msgs c' i')" using asm by simp then show False using asm by simp qed next show "i \ i' \ length (msgs c i') \ length (msgs c' i')" proof - assume "i \ i'" then show ?thesis proof (cases "channel i' = None") case True then show ?thesis using assms no_msgs_change_if_no_channel by presburger next case False then obtain r s where chan: "channel i' = Some (r, s)" by auto then show ?thesis proof (cases "has_snapshotted c p") case True with assms \i \ i'\ show ?thesis by auto next case no_snap: False then show ?thesis proof (cases "p = r") case True then have "msgs c' i' = msgs c i' @ [Marker]" using \i \ i'\ assms no_snap chan by auto then show ?thesis by auto next case False then show ?thesis using assms \i \ i'\ chan no_snap by auto qed qed qed qed qed lemma regular_event_cannot_induce_snapshot: assumes "~ has_snapshotted c p" and "c \ ev \ c'" shows "regular_event ev \ ~ has_snapshotted c' p" proof (cases ev) case (Trans q s s') then show ?thesis using assms(1) assms(2) by auto next case (Send q r s s' m) then show ?thesis using assms by auto next case (Recv q r s s' m) then show ?thesis using assms by auto qed simp_all lemma regular_event_preserves_process_snapshots: assumes "c \ ev \ c'" shows "regular_event ev \ ps c r = ps c' r" proof (cases ev) case (Trans p s s') then show ?thesis using assms by auto next case (Send p q s s' m) then show ?thesis using assms by auto next case (Recv p q s s' m) then show ?thesis using assms by auto qed simp_all lemma no_state_change_if_nonregular_event: assumes "~ regular_event ev" and "c \ ev \ c'" shows "states c p = states c' p" proof - have "isSnapshot ev \ isRecvMarker ev" using nonregular_event assms by auto then show ?thesis proof (elim disjE, goal_cases) case 1 then obtain q where "ev = Snapshot q" by (meson isSnapshot_def) then show ?thesis using assms(2) by auto next case 2 then obtain i q r where "ev = RecvMarker i q r" by (meson isRecvMarker_def) then show ?thesis using assms(2) by auto qed qed lemma nonregular_event_induces_snapshot: assumes "~ has_snapshotted c p" and "c \ ev \ c'" and "occurs_on ev = p" and "~ regular_event ev" shows "~ regular_event ev \ has_snapshotted c' p" proof (cases ev) case (Snapshot q) then have "q = p" using assms by auto then show ?thesis using Snapshot assms(2) by auto next case (RecvMarker i q r) then have "q = p" using assms by auto then show ?thesis using RecvMarker assms by auto qed (simp_all add: assms) lemma snapshot_state_unchanged: assumes step: "c \ ev \ c'" and "has_snapshotted c p" shows "ps c p = ps c' p" proof (cases "occurs_on ev = p") case False then show ?thesis using local.step no_state_change_if_no_event by auto next case True then show ?thesis proof (cases "regular_event ev") case True then show ?thesis using local.step regular_event_preserves_process_snapshots by auto next case False have "isRecvMarker ev" proof (rule ccontr) have "isSnapshot ev \ isRecvMarker ev" using False nonregular_event by blast moreover assume "~ isRecvMarker ev" ultimately have "isSnapshot ev" by simp then have "ev = Snapshot p" by (metis True event.collapse(4)) then have "can_occur ev c" using happen_implies_can_occur local.step by blast then have "~ has_snapshotted c p" unfolding can_occur_def by (simp add: \ev = Snapshot p\) then show False using assms by auto qed then show ?thesis (* z3 sledgehammer fails for Isabelle2019 *) proof - have "\n pa. c \ RecvMarker n p pa \ c'" by (metis True \isRecvMarker ev\ event.collapse(5) local.step) then show ?thesis using assms(2) by force qed qed qed lemma message_must_be_delivered: assumes valid: "c \ ev \ c'" and delivered: "(msgs c i \ [] \ hd (msgs c i) = m) \ (msgs c' i = [] \ hd (msgs c' i) \ m)" shows "(\p q. ev = RecvMarker i p q \ m = Marker) \ (\p q s s' m'. ev = Recv i p q s s' m' \ m = Msg m')" proof (cases ev) case (Snapshot p) then show ?thesis proof (cases "msgs c i") case Nil then show ?thesis using delivered by simp next case (Cons m xs) with assms Snapshot show ?thesis proof (cases "channel i = None") case True then show ?thesis using assms Snapshot by auto next case False then obtain r s where chan: "channel i = Some (r, s)" by auto then show ?thesis proof (cases "r = p") case True then have "msgs c' i = msgs c i @ [Marker]" using assms(1) Snapshot chan by auto then show ?thesis using delivered by auto next case False then have "msgs c' i = msgs c i" using assms Snapshot chan by simp then show ?thesis using delivered Cons by simp qed qed qed next case (RecvMarker i' p q) then have "i' = i" by (metis assms(1) delivered le_0_eq length_greater_0_conv list.size(3) recv_marker_changes_head_only_at_i recv_marker_other_channels_not_shrinking) moreover have "Marker = m" using \i' = i\ RecvMarker assms(1) can_occur_def delivered by auto moreover have "channel i = Some (q, p)" using RecvMarker assms(1) calculation(1) can_occur_def by auto ultimately show ?thesis using RecvMarker by simp next case (Trans p' s s') then show ?thesis using valid delivered by auto next case (Send p' q' s s' m') then show ?thesis by (metis (no_types, lifting) delivered distributed_system.next.simps(4) distributed_system_axioms hd_append2 snoc_eq_iff_butlast valid) next case (Recv i' p q s s' m') then have "i = i'" using assms(1) delivered by auto also have "m = Msg m'" by (metis (no_types, lifting) Recv delivered list.sel(1) next_recv valid) ultimately show ?thesis using Recv by auto qed lemma message_must_be_delivered_2: assumes "c \ ev \ c'" "m : set (msgs c i)" "m \ set (msgs c' i)" shows "(\p q. ev = RecvMarker i p q \ m = Marker) \ (\p q s s' m'. ev = Recv i p q s s' m' \ m = Msg m')" proof - have uneq_sets: "set (msgs c i) \ set (msgs c' i)" using assms(2) assms(3) by blast then obtain p q where chan: "channel i = Some (p, q)" using assms no_msgs_change_if_no_channel by fastforce then show ?thesis proof (cases ev) case (Snapshot p') with Snapshot assms chan have "set (msgs c' i) = set (msgs c i)" by (cases "p' = p", auto) then show ?thesis using uneq_sets by simp next case (Trans p' s s') then show ?thesis using uneq_sets assms by simp next case (Send i' p' q' s s' m) then show ?thesis by (metis (no_types, lifting) UnCI assms(1) assms(2) assms(3) local.next.simps(4) set_append) next case (RecvMarker i' p' q') have "i' = i" proof (rule ccontr) assume "~ i' = i" show False using assms chan RecvMarker proof (cases "has_snapshotted c p'") case True then show False using assms chan RecvMarker \~ i' = i\ by simp next case False then show False using assms chan RecvMarker \~ i' = i\ by (cases "p' = p", simp_all) qed qed moreover have "m = Marker" proof - have "msgs c i' = Marker # msgs c' i'" using assms chan RecvMarker by auto then show ?thesis using assms \i' = i\ by simp qed ultimately show ?thesis using RecvMarker by simp next case (Recv i' p' q' s s' m') have "i' = i" proof (rule ccontr) assume "~ i' = i" then show False using Recv assms(1) uneq_sets by auto qed then have "i' = i \ m = Msg m'" using Recv assms by auto then show ?thesis using Recv by simp qed qed lemma recv_marker_means_snapshotted_1: assumes "ev = RecvMarker i p q" and "c \ ev \ c'" shows "has_snapshotted c' p" using assms snapshot_state_unchanged by (cases "has_snapshotted c p", auto) lemma recv_marker_means_snapshotted_2: fixes c c' :: "('p, 's, 'm) configuration" and ev :: "('p, 's, 'm) event" and i :: channel_id assumes "c \ ev \ c'" and "Marker : set (msgs c i)" and "Marker \ set (msgs c' i)" and "channel i = Some (q, p)" shows "has_snapshotted c' p" proof - have "\p q. ev = RecvMarker i p q" using assms message_must_be_delivered_2 by blast then obtain r s where RecvMarker: "ev = RecvMarker i r s" by blast then have "r = p" using assms(1) assms(4) can_occur_def by auto then show ?thesis using recv_marker_means_snapshotted_1 assms RecvMarker by blast qed lemma event_stays_valid_if_no_occurrence: assumes "c \ ev \ c'" and "occurs_on ev \ occurs_on ev'" and "can_occur ev' c" shows "can_occur ev' c'" proof (cases ev') case (Trans p s s') have "states c p = states c' p" using Trans assms(1) assms(2) no_state_change_if_no_event by auto moreover have "states c p = s" using can_occur_def assms Trans by simp ultimately have "states c' p = s" by simp moreover have "trans p s s'" using Trans assms(3) can_occur_def by auto ultimately show ?thesis by (simp add: Trans can_occur_def) next case (Recv i p q s s' m) then have "hd (msgs c i) = Msg m" proof - from Recv have "length (msgs c i) > 0" using assms(3) can_occur_def by auto then obtain m' xs where mcqp: "msgs c i = m' # xs" by (metis list.size(3) nat_less_le neq_Nil_conv) then have "Msg m = m'" proof (cases m', auto) case Marker then have "msgs c i = Marker # xs" by (simp add:mcqp) then have "~ can_occur ev' c" using Recv can_occur_def by simp then show False using assms(3) by simp next case (Msg msg) then have "msgs c i = Msg msg # xs" by (simp add: mcqp) then show "m = msg" using Recv can_occur_def assms(3) by simp qed then show ?thesis by (simp add: mcqp) qed show ?thesis proof (rule ccontr) assume asm: "~ can_occur ev' c'" then have "msgs c' i = [] \ hd (msgs c' i) \ Msg m" using Recv assms can_occur_def no_state_change_if_no_event distributed_system_axioms list.case_eq_if by fastforce then obtain i' p' q' s'' s''' m' where RMoR: "ev = RecvMarker i' p' q' \ ev = Recv i p' q' s'' s''' m'" by (metis Recv \hd (msgs c i) = Msg m\ assms(1) assms(3) can_occur_Recv list.discI message_must_be_delivered) then have "occurs_on ev = p" proof - have f1: "states c p = s \ channel i = Some (q, p) \ recv i p q s s' m \ 0 < length (msgs c i) \ hd (msgs c i) = Msg m" using Recv assms(3) can_occur_def by force have f2: "RecvMarker i' p' q' = ev \ states c p' = s'' \ channel i = Some (q', p') \ recv i p' q' s'' s''' m' \ 0 < length (msgs c i) \ hd (msgs c i) = Msg m'" using RMoR assms(1) can_occur_def by force have "\e n c. \p pa s sa m. \ca cb. (\ c \ e \ ca \ msgs ca n \ [] \ hd (msgs c n) = Marker \ msgs c n = [] \ Recv n p pa s sa m = e) \ (\ c \ e \ cb \ hd (msgs c n) = Marker \ hd (msgs cb n) = hd (msgs c n) \ msgs c n = [] \ Recv n p pa s sa m = e)" by (metis (no_types) message_must_be_delivered) then show ?thesis using f2 f1 by (metis RMoR \msgs c' i = [] \ hd (msgs c' i) \ Msg m\ assms(1) event.disc(13,15) event.sel(3,5) length_greater_0_conv message.distinct(1) option.inject prod.inject) qed then show False using assms Recv by simp qed next case (Send i p q s s' m) then have "states c p = states c' p" using assms no_state_change_if_no_event by auto then show "can_occur ev' c'" using assms assms(3) can_occur_def Send by auto next case (RecvMarker i p q) then have msgs_ci: "hd (msgs c i) = Marker \ length (msgs c i) > 0" proof - from RecvMarker have "length (msgs c i) > 0" using assms(3) can_occur_def by auto then obtain m' xs where mci: "msgs c i = m' # xs" by (metis list.size(3) nat_less_le neq_Nil_conv) then have m_mark: "Marker = m'" proof (cases m', auto) case (Msg msg) then have "msgs c i = Msg msg # xs" by (simp add:mci) then have "~ can_occur ev' c" using RecvMarker can_occur_def by simp then show False using assms(3) by simp qed then show ?thesis by (simp add: mci) qed show ?thesis proof (rule ccontr) assume asm: "~ can_occur ev' c'" then have "msgs c' i = [] \ hd (msgs c' i) \ Marker" using RecvMarker assms(3) can_occur_def list.case_eq_if by fastforce then have "\p q. ev = RecvMarker i p q \ Marker = Marker" using message_must_be_delivered msgs_ci assms by blast then obtain r s where RecvMarker_ev: "ev = RecvMarker i r s" by blast then have "p = r \ q = s" using RecvMarker assms(1) assms(3) can_occur_def by auto then have "occurs_on ev = p" using assms RecvMarker_ev by auto then show False using assms using RecvMarker by auto qed next case (Snapshot p) then have "~ has_snapshotted c p" using assms assms(3) can_occur_def by simp show ?thesis proof (rule ccontr) assume asm: "~ can_occur ev' c'" then have "has_snapshotted c' p" using can_occur_def Snapshot by simp then have "occurs_on ev = p" using \\ has_snapshotted c p\ assms(1) no_state_change_if_no_event by fastforce then show False using assms(2) Snapshot by auto qed qed lemma msgs_unchanged_for_other_is: assumes "c \ ev \ c'" and "regular_event ev" and "getId ev = i" and "i' \ i" shows "msgs c i' = msgs c' i'" proof - have "isTrans ev \ isSend ev \ isRecv ev" using assms by simp then show ?thesis proof (elim disjE, goal_cases) case 1 then obtain p s s' where "ev = Trans p s s'" by (meson isTrans_def) then show ?thesis using assms by simp next case 2 then obtain i' p q s s' m where "ev = Send i' p q s s' m" by (meson isSend_def) then show ?thesis using assms by simp next case 3 then obtain i' p q s s' m where "ev = Recv i' p q s s' m" by (meson isRecv_def) with assms show ?thesis by auto qed qed lemma msgs_unchanged_if_snapshotted_RecvMarker_for_other_is: assumes "c \ ev \ c'" and "ev = RecvMarker i p q" and "has_snapshotted c p" and "i' \ i" shows "msgs c i' = msgs c' i'" using assms by auto lemma event_can_go_back_if_no_sender: assumes "c \ ev \ c'" and "occurs_on ev \ occurs_on ev'" and "can_occur ev' c'" and "~ isRecvMarker ev'" and "~ isSend ev" shows "can_occur ev' c" proof (cases ev') case (Snapshot p) then have "~ has_snapshotted c' p" using assms(3) can_occur_def by simp then have "~ has_snapshotted c p" using assms(1) snapshot_state_unchanged by force then show ?thesis using can_occur_def Snapshot by simp next case (RecvMarker i p q) then show ?thesis using assms(4) by auto next case (Trans p s s') then show ?thesis using assms(1) assms(2) can_occur_def no_state_change_if_no_event assms(3) by auto next case (Send p q s s' m) then show ?thesis using assms(1) assms(2) can_occur_def no_state_change_if_no_event assms(3) by auto next case (Recv i p q s s' m) have "msgs c' i \ Nil" using Recv can_occur_def assms by auto moreover have "hd (msgs c' i) = Msg m \ length (msgs c' i) > 0" proof - from Recv have "length (msgs c' i) > 0" using assms(3) can_occur_def by auto then obtain m' xs where mcqp: "msgs c' i = m' # xs" by (metis list.size(3) nat_less_le neq_Nil_conv) then have "Msg m = m'" proof (cases m', auto) case Marker then have "msgs c' i = Marker # xs" by (simp add:mcqp) then have "~ can_occur ev' c'" using Recv can_occur_def by simp then show False using assms(3) by simp next case (Msg msg) then have "msgs c' i = Msg msg # xs" by (simp add: mcqp) then show "m = msg" using Recv can_occur_def assms(3) by simp qed then show ?thesis by (simp add: mcqp) qed moreover have "msgs c i \ Nil \ hd (msgs c' i) = hd (msgs c i)" proof (cases ev) case (Snapshot p') then have "p' \ p" using assms Recv by simp have chan: "channel i = Some (q, p)" by (metis Recv assms(3) distributed_system.can_occur_Recv distributed_system_axioms) with Snapshot assms have "length (msgs c i) > 0 \ hd (msgs c i) = hd (msgs c' i)" proof (cases "q = p'") case True then have "msgs c' i = msgs c i @ [Marker]" using Snapshot chan assms by simp then show ?thesis by (metis append_self_conv2 calculation(2) hd_append2 length_greater_0_conv list.sel(1) message.simps(3)) next case False then have "msgs c' i = msgs c i" using Snapshot chan assms by simp then show ?thesis using calculation by simp qed then show ?thesis by simp next case (RecvMarker i' p' q') then have "i' \ i" using Recv assms(1) assms(2) assms(3) can_occur_def by force then show ?thesis proof (cases "has_snapshotted c p'") case True then have "msgs c i = msgs c' i" using \i' \ i\ RecvMarker assms by simp then show ?thesis using calculation by simp next case no_snap: False then have chan: "channel i = Some (q, p)" by (metis Recv assms(3) distributed_system.can_occur_Recv distributed_system_axioms) then show ?thesis proof (cases "q = p'") case True then have "msgs c' i = msgs c i @ [Marker]" using no_snap RecvMarker \i' \ i\ assms(1) chan by auto then show ?thesis by (metis append_self_conv2 calculation(2) hd_append2 list.sel(1) message.simps(3)) next case False then have "msgs c' i = msgs c i" using RecvMarker no_snap False chan assms \i' \ i\ by simp then show ?thesis using calculation by simp qed qed next case (Trans p' s'' s''') then show ?thesis using assms(1) \msgs c' i \ Nil\ by auto next case (Send i' p' q' s'' s''' m'') have "p' \ p" using Recv Send assms(2) by auto then show ?thesis using Recv Send assms(1) assms(5) calculation(1) by auto next case (Recv i' p' q' s'' s''' m'') then have "i' \ i" using assms \ev' = Recv i p q s s' m\ by (metis distributed_system.can_occur_Recv distributed_system_axioms event.sel(3) next_recv option.inject prod.inject) have "msgs c i = msgs c' i" using msgs_unchanged_for_other_is Recv \i' \ i\ assms(1) by auto then show ?thesis using \msgs c' i \ Nil\ by simp qed moreover have "states c p = states c' p" using no_state_change_if_no_event assms Recv by simp ultimately show ?thesis using Recv assms(3) can_occur_def list.case_eq_if by fastforce qed lemma nonregular_event_can_go_back_if_in_distinct_processes: assumes "c \ ev \ c'" and "regular_event ev" and "~ regular_event ev'" and "can_occur ev' c'" and "occurs_on ev \ occurs_on ev'" shows "can_occur ev' c" proof - let ?p = "occurs_on ev" let ?q = "occurs_on ev'" have "isTrans ev \ isSend ev \ isRecv ev" using assms by simp moreover have "isSnapshot ev' \ isRecvMarker ev'" using assms nonregular_event by auto ultimately show ?thesis proof (elim disjE, goal_cases) case 1 then show ?case using assms(1) assms(4) assms(5) event_can_go_back_if_no_sender by blast next case 2 then obtain s s' where Trans: "ev = Trans ?p s s'" by (metis event.collapse(1)) obtain i r where RecvMarker: "ev' = RecvMarker i ?q r" using 2 by (metis event.collapse(5)) have "msgs c i = msgs c' i" using "2"(1) assms(1) assms(2) no_msg_change_if_no_event by blast moreover have "can_occur ev' c'" using assms by simp ultimately show ?thesis using can_occur_def RecvMarker by (metis (mono_tags, lifting) "2"(2) event.case_eq_if event.distinct_disc(13) event.distinct_disc(17) event.distinct_disc(19) event.distinct_disc(7) event.sel(10)) next case 3 then have "ev' = Snapshot ?q" by (metis event.collapse(4)) have "~ has_snapshotted c' ?q" by (metis (mono_tags, lifting) "3"(1) assms(4) can_occur_def event.case_eq_if event.distinct_disc(11) event.distinct_disc(16) event.distinct_disc(6)) then have "~ has_snapshotted c ?q" using assms(1) assms(2) regular_event_preserves_process_snapshots by auto then show ?case unfolding can_occur_def using \ev' = Snapshot ?q\ by (metis (mono_tags, lifting) event.simps(29)) next case 4 then have "ev' = Snapshot ?q" by (metis event.collapse(4)) have "~ has_snapshotted c' ?q" by (metis (mono_tags, lifting) \ev' = Snapshot (occurs_on ev')\ assms(4) can_occur_def event.simps(29)) then have "~ has_snapshotted c ?q" using assms(1) assms(2) regular_event_preserves_process_snapshots by auto then show ?case unfolding can_occur_def by (metis (mono_tags, lifting) \ev' = Snapshot (occurs_on ev')\ event.simps(29)) next case 5 then obtain i s u u' m where "ev = Send i ?p s u u' m" by (metis event.collapse(2)) from 5 obtain i' r where "ev' = RecvMarker i' ?q r" by (metis event.collapse(5)) then have pre: "hd (msgs c' i') = Marker \ length (msgs c' i') > 0" by (metis (mono_tags, lifting) assms(4) can_occur_def event.simps(30)) have "hd (msgs c i') = Marker \ length (msgs c i') > 0" proof (cases "i' = i") case False then have "msgs c i' = msgs c' i'" by (metis \ev = Send i (occurs_on ev) s u u' m\ assms(1) assms(2) event.sel(8) msgs_unchanged_for_other_is) then show ?thesis using pre by auto next case True then have "msgs c' i' = msgs c i' @ [Msg m]" by (metis \ev = Send i (occurs_on ev) s u u' m\ assms(1) next_send) then have "length (msgs c' i') > 1" using pre by fastforce then have "length (msgs c i') > 0" by (simp add: \msgs c' i' = msgs c i' @ [Msg m]\) then show ?thesis using \msgs c' i' = msgs c i' @ [Msg m]\ pre by auto qed then show ?case unfolding can_occur_def using \ev' = RecvMarker i' ?q r\ by (metis (mono_tags, lifting) assms(4) can_occur_def event.simps(30)) next case 6 then obtain i s u u' m where "ev = Recv i ?p s u u' m" by (metis event.collapse(3)) from 6 obtain i' r where "ev' = RecvMarker i' ?q r" by (metis event.collapse(5)) then have "i' \ i" proof - have "?p \ ?q" using assms by simp moreover have "channel i = Some (s, ?p)" by (metis \ev = Recv i (occurs_on ev) s u u' m\ assms(1) distributed_system.can_occur_Recv distributed_system_axioms happen_implies_can_occur) moreover have "channel i' = Some (r, ?q)" by (metis (mono_tags, lifting) \ev' = RecvMarker i' (occurs_on ev') r\ assms(4) can_occur_def event.case_eq_if event.disc(5,10,15,20) event.sel(5,10,13)) ultimately show ?thesis by auto qed then show ?case by (metis (mono_tags, lifting) "6"(1) \ev = Recv i (occurs_on ev) s u u' m\ \ev' = RecvMarker i' (occurs_on ev') r\ assms(1) assms(4) can_occur_def event.case_eq_if event.distinct_disc(13) event.distinct_disc(17) event.distinct_disc(7) event.sel(10) next_recv) qed qed lemma same_state_implies_same_result_state: assumes "states c p = states d p" "c \ ev \ c'" and "d \ ev \ d'" shows "states d' p = states c' p" proof (cases "occurs_on ev = p") case False then show ?thesis by (metis assms(1-3) distributed_system.no_state_change_if_no_event distributed_system_axioms) next case True then show ?thesis using assms by (cases ev, auto) qed lemma same_snapshot_state_implies_same_result_snapshot_state: assumes "ps c p = ps d p" and "states c p = states d p" and "c \ ev \ c'" and "d \ ev \ d'" shows "ps d' p = ps c' p" proof (cases "occurs_on ev = p") case False then show ?thesis using assms no_state_change_if_no_event by auto next case True then show ?thesis proof (cases ev) case (Snapshot q) then have "p = q" using True by auto then show ?thesis using Snapshot assms(2) assms(3) assms(4) by auto next case (RecvMarker i q r) then have "p = q" using True by auto then show ?thesis proof - have f1: "\c ca. \ c \ ev \ ca \ ps c p = None \ ps c p = ps ca p" using RecvMarker \p = q\ by force have "\c ca. ps c p \ None \ \ c \ ev \ ca \ ps ca p = Some (states c p)" using RecvMarker \p = q\ by force then show ?thesis using f1 by (metis (no_types) assms(1) assms(2) assms(3) assms(4)) qed next case (Trans q s s') then have "p = q" using True by auto then show ?thesis using Trans assms(1) assms(3) assms(4) by auto next case (Send i q r u u' m) then have "p = q" using True by auto then show ?thesis using Send assms(1) assms(3) assms(4) by auto next case (Recv i q r u u' m) then have "p = q" using True by auto then show ?thesis using Recv assms(1) assms(3) assms(4) by auto qed qed lemma same_messages_imply_same_resulting_messages: assumes "msgs c i = msgs d i" "c \ ev \ c'" and "d \ ev \ d'" and "regular_event ev" shows "msgs c' i = msgs d' i" proof - have "isTrans ev \ isSend ev \ isRecv ev" using assms by simp then show ?thesis proof (elim disjE) assume "isTrans ev" then show ?thesis by (metis assms(1) assms(2) assms(3) isTrans_def next_trans) next assume "isSend ev" then obtain i' r s u u' m where "ev = Send i' r s u u' m" by (metis event.collapse(2)) with assms show ?thesis by (cases "i = i'", auto) next assume "isRecv ev" then obtain i' r s u u' m where Recv: "ev = Recv i' r s u u' m" by (metis event.collapse(3)) with assms show ?thesis by (cases "i = i'", auto) qed qed lemma Trans_msg: assumes "c \ ev \ c'" and "isTrans ev" shows "msgs c i = msgs c' i" using assms(1) assms(2) no_msg_change_if_no_event regular_event by blast lemma new_msg_in_set_implies_occurrence: assumes "c \ ev \ c'" and "m \ set (msgs c i)" and "m \ set (msgs c' i)" and "channel i = Some (p, q)" shows "occurs_on ev = p" (is ?P) proof (rule ccontr) assume "~ ?P" have "set (msgs c' i) \ set (msgs c i)" proof (cases ev) case (Snapshot r) then have "msgs c' i = msgs c i" using \~ ?P\ assms by simp then show ?thesis by auto next case (RecvMarker i' r s) then show ?thesis proof (cases "has_snapshotted c r") case True then show ?thesis proof (cases "i' = i") case True then have "Marker # msgs c' i = msgs c i" using RecvMarker True assms by simp then show ?thesis by (metis set_subset_Cons) next case False then show ?thesis using RecvMarker True assms by simp qed next case no_snap: False have chan: "channel i' = Some (s, r)" using RecvMarker assms(1) can_occur_def by auto then show ?thesis proof (cases "i' = i") case True then have "Marker # msgs c' i = msgs c i" using RecvMarker assms by simp then show ?thesis by (metis set_subset_Cons) next case False then have "msgs c' i = msgs c i" using \~ ?P\ RecvMarker assms no_snap by simp then show ?thesis by simp qed qed next case (Trans r u u') then show ?thesis using assms \~ ?P\ by simp next case (Send i' r s u u' m') then have "i' \ i" using \~ ?P\ can_occur_def assms by auto then have "msgs c i = msgs c' i" using \~ ?P\ assms Send by simp then show ?thesis by simp next case (Recv i' r s u u' m') then show ?thesis by (metis (no_types, lifting) assms(1) eq_iff local.next.simps(5) set_subset_Cons) qed moreover have "~ set (msgs c' i) \ set (msgs c i)" using assms by blast ultimately show False by simp qed lemma new_Marker_in_set_implies_nonregular_occurence: assumes "c \ ev \ c'" and "Marker \ set (msgs c i)" and "Marker \ set (msgs c' i)" and "channel i = Some (p, q)" shows "~ regular_event ev" (is ?P) proof (rule ccontr) have "occurs_on ev = p" using assms new_msg_in_set_implies_occurrence by blast assume "~ ?P" then have "isTrans ev \ isSend ev \ isRecv ev" by simp then have "Marker \ set (msgs c' i)" proof (elim disjE, goal_cases) case 1 then obtain r u u' where "ev = Trans r u u'" by (metis event.collapse(1)) then show ?thesis using assms(1) assms(2) by auto next case 2 then obtain i' r q u u' m where "ev = Send i' r q u u' m" by (metis event.collapse(2)) then show ?thesis by (metis (no_types, lifting) Un_iff assms(1) assms(2) empty_iff empty_set insert_iff list.set(2) message.distinct(1) next_send set_append) next case 3 then obtain i' r q u u' m where "ev = Recv i' r q u u' m" by (metis event.collapse(3)) then show ?thesis by (metis assms(1) assms(2) list.set_intros(2) next_recv) qed then show False using assms by simp qed lemma RecvMarker_implies_Marker_in_set: assumes "c \ ev \ c'" and "ev = RecvMarker cid p q" shows "Marker \ set (msgs c cid)" by (metis (mono_tags, lifting) assms(1) assms(2) can_occur_def distributed_system.happen_implies_can_occur distributed_system_axioms event.simps(30) list.set_sel(1) list.size(3) nat_less_le) lemma RecvMarker_given_channel: assumes "isRecvMarker ev" and "getId ev = cid" and "channel cid = Some (p, q)" and "can_occur ev c" shows "ev = RecvMarker cid q p" by (metis (mono_tags, lifting) assms(1) assms(2) assms(3) assms(4) can_occur_def event.case_eq_if event.collapse(5) event.distinct_disc(8,14,18,20) option.inject prod.inject) lemma Recv_given_channel: assumes "isRecv ev" and "getId ev = cid" and "channel cid = Some (p, q)" and "can_occur ev c" shows "\s s' m. ev = Recv cid q p s s' m" by (metis assms(1) assms(2) assms(3) assms(4) distributed_system.can_occur_Recv distributed_system_axioms event.collapse(3) option.inject prod.inject) lemma same_cs_if_not_recv: assumes "c \ ev \ c'" and "~ isRecv ev" shows "fst (cs c cid) = fst (cs c' cid)" proof (cases "channel cid = None") case True then show ?thesis using assms(1) no_cs_change_if_no_channel by auto next case False then obtain p q where chan: "channel cid = Some (p, q)" by auto then show ?thesis proof (cases ev) case (Snapshot r) with Snapshot assms chan show ?thesis by (cases "r = q", auto) next case (RecvMarker cid' r s) then show ?thesis proof (cases "has_snapshotted c r") case True with assms RecvMarker chan show ?thesis by (cases "cid' = cid", auto) next case no_snap: False then show ?thesis proof (cases "cid' = cid") case True then show ?thesis using RecvMarker assms chan by auto next case False with assms RecvMarker chan no_snap show ?thesis by (cases "r = q", auto) qed qed next case (Trans r u u') then show ?thesis using assms by auto next case (Send r s u u') then show ?thesis using assms by auto qed (metis assms(2) isRecv_def) qed lemma done_only_from_recv_marker: assumes "c \ ev \ c'" and "channel cid = Some (p, q)" and "snd (cs c cid) \ Done" and "snd (cs c' cid) = Done" shows "ev = RecvMarker cid q p" proof (rule ccontr) assume "~ ev = RecvMarker cid q p" then show False proof (cases "isRecvMarker ev") case True then obtain cid' s r where RecvMarker: "ev = RecvMarker cid' s r" by (meson isRecvMarker_def) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then show False using \ev = RecvMarker cid' s r\ \ev \ RecvMarker cid q p\ assms(1) assms(2) can_occur_def by auto qed then have "snd (cs c' cid) \ Done" proof (cases "has_snapshotted c s") case True then show ?thesis using RecvMarker assms \cid \ cid'\ by simp next case False with RecvMarker assms \cid \ cid'\ show ?thesis by (cases "s = q", auto) qed then show False using assms by auto next case False then have "isSnapshot ev \ isTrans ev \ isSend ev \ isRecv ev" using event.exhaust_disc by blast then have "snd (cs c' cid) \ Done" proof (elim disjE, goal_cases) case 1 then obtain r where Snapshot: "ev = Snapshot r" by (meson isSnapshot_def) with assms show ?thesis by (cases "q = r", auto) next case 2 then obtain r u u' where "ev = Trans r u u'" by (meson isTrans_def) then show ?case using assms by auto next case 3 then obtain cid' r s u u' m where "ev = Send cid' r s u u' m" by (meson isSend_def) then show ?thesis using assms by auto next case 4 then obtain cid' r s u u' m where Recv: "ev = Recv cid' r s u u' m" by (meson isRecv_def) show ?thesis using Recv assms proof (cases "cid = cid'") case True then have "snd (cs c cid) = NotStarted \ snd (cs c cid) = Recording" using assms(3) recording_state.exhaust by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then have "snd (cs c' cid') = NotStarted" using True Recv assms(1) by auto then show ?case using True by auto next case 2 then have "snd (cs c' cid') = Recording" using True Recv assms(1) by auto then show ?case using True by auto qed qed auto qed then show False using assms by auto qed qed lemma cs_not_not_started_stable: assumes "c \ ev \ c'" and "snd (cs c cid) \ NotStarted" and "channel cid = Some (p, q)" shows "snd (cs c' cid) \ NotStarted" using assms proof (cases ev) case (Snapshot r) then show ?thesis by (metis assms(1) assms(2) next_snapshot recording_state.simps(2) sndI) next case (RecvMarker cid' r s) then show ?thesis proof (cases "has_snapshotted c r") case True with RecvMarker assms show ?thesis by (cases "cid = cid'", auto) next case no_snap: False then show ?thesis proof (cases "cid = cid'") case True then show ?thesis using RecvMarker assms by auto next case False with RecvMarker assms no_snap show ?thesis by (cases "s = p", auto) qed qed next case (Recv cid' r s u u' m) then have "snd (cs c cid) = Recording \ snd (cs c cid) = Done" using assms(2) recording_state.exhaust by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then show ?thesis by (metis (no_types, lifting) Recv assms(1) eq_snd_iff next_recv recording_state.distinct(1)) next case 2 with Recv assms show ?thesis by (cases "cid = cid'", auto) qed qed auto lemma fst_cs_changed_by_recv_recording: assumes step: "c \ ev \ c'" and "fst (cs c cid) \ fst (cs c' cid)" and "channel cid = Some (p, q)" shows "snd (cs c cid) = Recording \ (\p q u u' m. ev = Recv cid q p u u' m)" proof - have oc_on: "occurs_on ev = q" proof - obtain nn :: "('p, 's, 'm) event \ nat" and aa :: "('p, 's, 'm) event \ 'p" and aaa :: "('p, 's, 'm) event \ 'p" and bb :: "('p, 's, 'm) event \ 's" and bba :: "('p, 's, 'm) event \ 's" and cc :: "('p, 's, 'm) event \ 'm" where f1: "\e. (\ isRecv e \ e = Recv (nn e) (aa e) (aaa e) (bb e) (bba e) (cc e)) \ (isRecv e \ (\n a aa b ba c. e \ Recv n a aa b ba c))" - using isRecv_def by moura + by (metis isRecv_def) then have f2: "c \ Recv (nn ev) (aa ev) (aaa ev) (bb ev) (bba ev) (cc ev) \ c'" by (metis (no_types) assms(2) local.step same_cs_if_not_recv) have f3: "\x0 x1 x7 x8. (x0 \ x7 \ cs (x8::('p, 's, 'm) configuration) x0 = cs (x1::('p, 's, _) configuration) x0) = (x0 = x7 \ cs x8 x0 = cs x1 x0)" by auto have f4: "\x0 x1 x7 x8. (x7 \ x0 \ msgs (x1::('p, 's, 'm) configuration) x0 = msgs (x8::('p, 's, _) configuration) x0) = (x7 = x0 \ msgs x1 x0 = msgs x8 x0)" by auto have "\x0 x1 x6 x8. (x0 \ x6 \ states (x1::('p, 's, 'm) configuration) x0 = states (x8::(_, _, 'm) configuration) x0) = (x0 = x6 \ states x1 x0 = states x8 x0)" by fastforce then have "can_occur (Recv (nn ev) (aa ev) (aaa ev) (bb ev) (bba ev) (cc ev)) c \ states c (aa ev) = bb ev \ states c' (aa ev) = bba ev \ (\a. a = aa ev \ states c' a = states c a) \ msgs c (nn ev) = Msg (cc ev) # msgs c' (nn ev) \ (\n. nn ev = n \ msgs c' n = msgs c n) \ (\a. ps c a = ps c' a) \ (\n. n = nn ev \ cs c n = cs c' n) \ (if snd (cs c (nn ev)) = Recording then cs c' (nn ev) = (fst (cs c (nn ev)) @ [cc ev], Recording) else cs c (nn ev) = cs c' (nn ev))" using f4 f3 f2 by force then show ?thesis using f1 by (metis (no_types) Pair_inject assms(2) assms(3) can_occur_Recv event.sel(3) local.step option.sel same_cs_if_not_recv) qed have "isRecv ev" (is ?P) proof (rule ccontr) assume "~ ?P" then have "fst (cs c cid) = fst (cs c' cid)" by (metis local.step same_cs_if_not_recv) then show False using assms by simp qed then obtain cid' r s u u' m where Recv: "ev = Recv cid' r s u u' m" by (meson isRecv_def) have "cid = cid'" proof (rule ccontr) assume "~ cid = cid'" then have "fst (cs c cid) = fst (cs c' cid)" using Recv step by auto then show False using assms by simp qed moreover have "snd (cs c cid) = Recording" proof (rule ccontr) assume "~ snd (cs c cid) = Recording" then have "fst (cs c cid) = fst (cs c' cid)" using Recv step \cid = cid'\ by auto then show False using assms by simp qed ultimately show ?thesis using Recv by simp qed lemma no_marker_and_snapshotted_implies_no_more_markers: assumes "c \ ev \ c'" and "has_snapshotted c p" and "Marker \ set (msgs c cid)" and "channel cid = Some (p, q)" shows "Marker \ set (msgs c' cid)" proof (cases ev) case (Snapshot r) then have "r \ p" using assms(1) assms(2) can_occur_def by auto then have "msgs c cid = msgs c' cid" using assms Snapshot by simp then show ?thesis using assms by simp next case (RecvMarker cid' r s) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" moreover have "can_occur ev c" using happen_implies_can_occur assms by blast ultimately have "Marker : set (msgs c cid)" using can_occur_def RecvMarker by (metis (mono_tags, lifting) assms(1) event.simps(30) hd_in_set list.size(3) recv_marker_other_channels_not_shrinking zero_order(1)) then show False using assms by simp qed then have "msgs c cid = msgs c' cid" proof (cases "r = p") case True then show ?thesis using RecvMarker \cid \ cid'\ assms(1) assms(2) msgs_unchanged_if_snapshotted_RecvMarker_for_other_is by blast next case False with RecvMarker \cid \ cid'\ step assms show ?thesis by (cases "has_snapshotted c r", auto) qed then show ?thesis using assms by simp next case (Trans r u u') then show ?thesis using assms by auto next case (Send cid' r s u u' m) with assms Send show ?thesis by (cases "cid = cid'", auto) next case (Recv cid' r s u u' m) with assms Recv show ?thesis by (cases "cid = cid'", auto) qed lemma same_messages_if_no_occurrence: assumes "c \ ev \ c'" and "~ occurs_on ev = p" and "~ occurs_on ev = q" and "channel cid = Some (p, q)" shows "msgs c cid = msgs c' cid \ cs c cid = cs c' cid" proof (cases ev) case (Snapshot r) then show ?thesis using assms by auto next case (RecvMarker cid' r s) have "cid \ cid'" by (metis RecvMarker_given_channel assms(1) assms(3) assms(4) RecvMarker event.sel(5,10) happen_implies_can_occur isRecvMarker_def) have "\a. channel cid = Some (r, q)" using assms(2) assms(4) RecvMarker by auto with RecvMarker assms \cid \ cid'\ show ?thesis by (cases "has_snapshotted c r", auto) next case (Trans r u u') then show ?thesis using assms by auto next case (Send cid' r s u u' m) then have "cid \ cid'" by (metis (mono_tags, lifting) Pair_inject assms(1) assms(2) assms(4) can_occur_def event.sel(2) event.simps(27) happen_implies_can_occur option.inject) then show ?thesis using assms Send by simp next case (Recv cid' r s u u' m) then have "cid \ cid'" by (metis assms(1) assms(3) assms(4) distributed_system.can_occur_Recv distributed_system.happen_implies_can_occur distributed_system_axioms event.sel(3) option.inject prod.inject) then show ?thesis using assms Recv by simp qed end (* locale distributed_system *) end (* theory Distributed_System *) diff --git a/thys/Chandy_Lamport/Util.thy b/thys/Chandy_Lamport/Util.thy --- a/thys/Chandy_Lamport/Util.thy +++ b/thys/Chandy_Lamport/Util.thy @@ -1,366 +1,366 @@ section \Utilties\ theory Util imports Main "HOL-Library.Sublist" "HOL-Library.Multiset" begin abbreviation swap_events where "swap_events i j t \ take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t) @ drop (j+1) t" lemma swap_neighbors_2: shows "i+1 < length t \ swap_events i (i+1) t = (t[i := t ! (i+1)])[i+1 := t ! i]" proof (induct i arbitrary: t) case 0 then show ?case by (metis One_nat_def Suc_eq_plus1 add_lessD1 append.left_neutral append_Cons cancel_comm_monoid_add_class.diff_cancel drop_update_cancel length_list_update numeral_One take_0 take_Cons_numeral upd_conv_take_nth_drop zero_less_Suc) next case (Suc n) let ?t = "tl t" have "t = hd t # ?t" by (metis Suc.prems hd_Cons_tl list.size(3) not_less_zero) moreover have "swap_events n (n+1) ?t = (?t[n := ?t ! (n+1)])[n+1 := ?t ! n]" by (metis Suc.hyps Suc.prems Suc_eq_plus1 length_tl less_diff_conv) ultimately show ?case by (metis Suc_eq_plus1 append_Cons diff_self_eq_0 drop_Suc_Cons list_update_code(3) nth_Cons_Suc take_Suc_Cons) qed lemma swap_identical_length: assumes "i < j" and "j < length t" shows "length t = length (swap_events i j t)" proof - have "length (take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i+1)) (drop (i+1) t))" by simp then have "... = j+1" using assms(1) assms(2) by auto then show ?thesis using assms(2) by auto qed lemma swap_identical_heads: assumes "i < j" and "j < length t" shows "take i t = take i (swap_events i j t)" using assms by auto lemma swap_identical_tails: assumes "i < j" and "j < length t" shows "drop (j+1) t = drop (j+1) (swap_events i j t)" proof - have "length (take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i+1)) (drop (i+1) t))" by simp then have "... = j+1" using assms(1) assms(2) by auto then show ?thesis by (metis \length (take i t @ [t ! j, t ! i] @ take (j - (i + 1)) (drop (i + 1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i + 1)) (drop (i + 1) t))\ append.assoc append_eq_conv_conj) qed lemma swap_neighbors: shows "i+1 < length l \ (l[i := l ! (i+1)])[i+1 := l ! i] = take i l @ [l ! (i+1), l ! i] @ drop (i+2) l" proof (induct i arbitrary: l) case 0 then show ?case by (metis One_nat_def add.left_neutral add_lessD1 append_Cons append_Nil drop_update_cancel length_list_update one_add_one plus_1_eq_Suc take0 take_Suc_Cons upd_conv_take_nth_drop zero_less_two) next case (Suc n) let ?l = "tl l" have "(l[Suc n := l ! (Suc n + 1)])[Suc n + 1 := l ! Suc n] = hd l # (?l[n := ?l ! (n+1)])[n+1 := ?l ! n]" by (metis Suc.prems add.commute add_less_same_cancel2 list.collapse list.size(3) list_update_code(3) not_add_less2 nth_Cons_Suc plus_1_eq_Suc) have "n + 1 < length ?l" using Suc.prems by auto then have "(?l[n := ?l ! (n+1)])[n+1 := ?l ! n] = take n ?l @ [?l ! (n+1), ?l ! n] @ drop (n+2) ?l" using Suc.hyps by simp then show ?case by (cases l) auto qed lemma swap_events_perm: assumes "i < j" and "j < length t" shows "mset (swap_events i j t) = mset t" proof - have swap: "swap_events i j t = take i t @ [t ! j, t ! i] @ (take (j - (i+1)) (drop (i+1) t)) @ (drop (j+1) t)" by auto have reg: "t = take i t @ (take ((j+1) - i) (drop i t)) @ drop (j+1) t" by (metis add_diff_inverse_nat add_lessD1 append.assoc append_take_drop_id assms(1) less_imp_add_positive less_not_refl take_add) have "mset (take i t) = mset (take i t)" by simp moreover have "mset (drop (j+1) t) = mset (drop (j+1) t)" by simp moreover have "mset ([t ! j, t ! i] @ (take (j - (i+1)) (drop (i+1) t))) = mset (take ((j+1) - i) (drop i t))" proof - let ?l = "take (j - (i+1)) (drop (i+1) t)" have "take ((j+1) - i) (drop i t) = t ! i # ?l @ [t ! j]" proof - have f1: "Suc (j - Suc i) = j - i" by (meson Suc_diff_Suc assms(1)) have f2: "i < length t" using assms(1) assms(2) by linarith have f3: "j - (i + 1) + (i + 1) = j" using \i < j\ by simp then have "drop (j - (i + 1)) (drop (i + 1) t) = drop j t" by (metis drop_drop) then show ?thesis using f3 f2 f1 by (metis (no_types) Cons_nth_drop_Suc Suc_diff_le Suc_eq_plus1 assms(1) assms(2) hd_drop_conv_nth length_drop less_diff_conv nat_less_le take_Suc_Cons take_hd_drop) qed then show ?thesis by fastforce qed ultimately show ?thesis using swap reg by simp (metis mset_append union_mset_add_mset_left union_mset_add_mset_right) qed lemma sum_eq_if_same_subterms: fixes i :: nat shows "\k. i \ k \ k < j \ f k = f' k \ sum f {i..) a) l \ takeWhile ((\) a) l" shows "\i j. i < j \ j < length l \ l ! i = a \ l ! j \ a" (is ?P) proof (rule ccontr) assume "~ ?P" then have asm: "\i j. i < j \ j < length l \ l ! i \ a \ l ! j = a" (is ?Q) by simp then have "filter ((\) a) l = takeWhile ((\) a) l" proof (cases "a : set l") case False then have "\i. i < length l \ l ! i \ a" by auto then show ?thesis by (metis (mono_tags, lifting) False filter_True takeWhile_eq_all_conv) next case True then have ex_j: "\j. j < length l \ l ! j = a" by (simp add: in_set_conv_nth) let ?j = "Min {j. j < length l \ l ! j = a}" have fin_j: "finite {j. j < length l \ l ! j = a}" using finite_nat_set_iff_bounded by blast moreover have "{j. j < length l \ l ! j = a} \ empty" using ex_j by blast ultimately have "?j < length l" using Min_less_iff by blast have tail_all_a: "\j. j < length l \ j \ ?j \ l ! j = a" proof (rule allI, rule impI) fix j assume "j < length l \ j \ ?j" moreover have "\ ?Q; j < length l \ j \ ?j \ \ \k. k \ ?j \ k \ j \ l ! j = a" proof (induct "j - ?j") case 0 then have "j = ?j" using 0 by simp then show ?case using Min_in \{j. j < length l \ l ! j = a} \ {}\ fin_j by blast next case (Suc n) then have "\k. k \ ?j \ k < j \ l ! j = a" by (metis (mono_tags, lifting) Min_in \{j. j < length l \ l ! j = a} \ {}\ fin_j le_eq_less_or_eq mem_Collect_eq) then show ?case using Suc.hyps(2) by auto qed ultimately show "l ! j = a" using asm by blast qed moreover have head_all_not_a: "\i. i < ?j \ l ! i \ a" using asm calculation by (metis (mono_tags, lifting) Min_le \Min {j. j < length l \ l ! j = a} < length l\ fin_j leD less_trans mem_Collect_eq) ultimately have "takeWhile ((\) a) l = take ?j l" proof - have "length (takeWhile ((\) a) l) = ?j" proof - have "length (takeWhile ((\) a) l) \ ?j" (is ?S) proof (rule ccontr) assume "\ ?S" then have "l ! ?j \ a" by (metis (mono_tags, lifting) not_le_imp_less nth_mem set_takeWhileD takeWhile_nth) then show False using \Min {j. j < length l \ l ! j = a} < length l\ tail_all_a by blast qed moreover have "length (takeWhile ((\) a) l) \ ?j" (is ?T) proof (rule ccontr) assume "\ ?T" then have "\j. j < ?j \ l ! j = a" by (metis (mono_tags, lifting) \Min {j. j < length l \ l ! j = a} < length l\ calculation le_less_trans not_le_imp_less nth_length_takeWhile) then show False using head_all_not_a by blast qed ultimately show ?thesis by simp qed moreover have "length (take ?j l) = ?j" by (metis calculation takeWhile_eq_take) ultimately show ?thesis by (metis takeWhile_eq_take) qed moreover have "filter ((\) a) l = take ?j l" proof - have "filter ((\) a) l = filter ((\) a) (take ?j l) @ filter ((\) a) (drop ?j l)" by (metis append_take_drop_id filter_append) moreover have "filter ((\) a) (take ?j l) = take ?j l" using head_all_not_a by (metis \takeWhile ((\) a) l = take (Min {j. j < length l \ l ! j = a}) l\ filter_id_conv set_takeWhileD) moreover have "filter ((\) a) (drop ?j l) = []" proof - have "\j. j \ ?j \ j < length l \ l ! j = drop ?j l ! (j - ?j)" by simp then have "\j. j < length l - ?j \ drop ?j l ! j = a" using tail_all_a by (metis (no_types, lifting) Groups.add_ac(2) \Min {j. j < length l \ l ! j = a} < length l\ less_diff_conv less_imp_le_nat not_add_less2 not_le nth_drop) then show ?thesis proof - obtain aa :: "('a \ bool) \ 'a list \ 'a" where "\x0 x1. (\v2. v2 \ set x1 \ x0 v2) = (aa x0 x1 \ set x1 \ x0 (aa x0 x1))" by moura then have f1: "\as p. aa p as \ set as \ p (aa p as) \ filter p as = []" by (metis (full_types) filter_False) obtain nn :: "'a list \ 'a \ nat" where f2: "\x0 x1. (\v2 x0 ! nn x0 x1 = x1)" - by moura + by moura { assume "drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) = a" then have "filter ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l) = [] \ \ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" using f1 by (metis (full_types)) } moreover { assume "\ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length l - Min {n. n < length l \ l ! n = a}" then have "\ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" by simp } ultimately have "filter ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l) = [] \ \ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" using \\j l ! j = a}. drop (Min {j. j < length l \ l ! j = a}) l ! j = a\ by blast then show ?thesis using f2 f1 by (meson in_set_conv_nth) qed qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed then show False using assms by simp qed lemma util_exactly_one_element: assumes "m \ set l" and "l' = l @ [m]" shows "\!j. j < length l' \ l' ! j = m" (is ?P) proof - have "\j. j < length l' - 1 \ l' ! j \ m" by (metis assms(1) assms(2) butlast_snoc length_butlast nth_append nth_mem) then have one_j: "\j. j < length l' \ l' ! j = m \ j = (length l' - 1)" by (metis (no_types, opaque_lifting) diff_Suc_1 lessE) show ?thesis proof (rule ccontr) assume "~ ?P" then obtain i j where "i \ j" "i < length l'" "j < length l'" "l' ! i = m" "l' ! j = m" using assms by auto then show False using one_j by blast qed qed lemma exists_one_iff_filter_one: shows "(\!j. j < length l \ l ! j = a) \ length (filter ((=) a) l) = 1" proof (rule iffI) assume "\!j. j < length l \ l ! j = a" then obtain j where "j < length l" "l ! j = a" by blast moreover have "\k. k \ j \ k < length l \ l ! k \ a" using \\!j. j < length l \ l ! j = a\ \j < length l\ \l ! j = a\ by blast moreover have "l = take j l @ [l ! j] @ drop (j+1) l" by (metis Cons_eq_appendI Cons_nth_drop_Suc Suc_eq_plus1 append_self_conv2 append_take_drop_id calculation(1) calculation(2)) moreover have "filter ((=) a) (take j l) = []" proof - have "\k. k < length (take j l) \ (take j l) ! k \ a" using calculation(3) by auto then show ?thesis by (metis (full_types) filter_False in_set_conv_nth) qed moreover have "filter ((=) a) (drop (j+1) l) = []" proof - have "\k. k < length (drop (j+1) l) \ (drop (j+1) l) ! k \ a" using calculation(3) by auto then show ?thesis by (metis (full_types) filter_False in_set_conv_nth) qed ultimately show "length (filter ((=) a) l) = 1" by (metis (mono_tags, lifting) One_nat_def Suc_eq_plus1 append_Cons append_self_conv2 filter.simps(2) filter_append list.size(3) list.size(4)) next assume asm: "length (filter ((=) a) l) = 1" then have "filter ((=) a) l = [a]" proof - let ?xs = "filter ((=) a) l" have "length ?xs = 1" using asm by blast then show ?thesis by (metis (full_types) Cons_eq_filterD One_nat_def length_0_conv length_Suc_conv) qed then have "\j. j < length l \ l ! j = a" by (metis (full_types) filter_False in_set_conv_nth list.discI) then obtain j where j: "j < length l" "l ! j = a" by blast moreover have "\k. k < length l \ k \ j \ l ! k \ a" proof (rule allI, rule impI) fix k assume assm: "k < length l \ k \ j" then have \k < length l\ .. show "l ! k \ a" proof (rule ccontr) assume lka: "\ l ! k \ a" then have \l ! k = a\ by simp show False proof (cases "k < j") let ?xs = "take (k+1) l" let ?ys = "drop (k+1) l" case True then have "length (filter ((=) a) ?xs) > 0" using \k < length l\ \l ! k = a\ by (auto simp add: filter_empty_conv in_set_conv_nth) moreover have "length (filter ((=) a) ?ys) > 0" proof - have "?ys ! (j - (k+1)) = l ! j" using True assm by auto moreover have "j - (k+1) < length ?ys" using True \j < length l\ by auto ultimately show ?thesis by (metis (full_types) \l ! j = a\ filter_empty_conv length_greater_0_conv nth_mem) qed moreover have "?xs @ ?ys = l" using append_take_drop_id by blast ultimately have "length (filter ((=) a) l) > 1" by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 asm filter_append length_append less_add_eq_less less_one nat_neq_iff) then show False using asm by simp next let ?xs = "take (j+1) l" let ?ys = "drop (j+1) l" case False then have "length (filter ((=) a) ?xs) > 0" using \k < length l\ \l ! j = a\ by (auto simp add: filter_empty_conv in_set_conv_nth) moreover have "length (filter ((=) a) ?ys) > 0" proof - have "?ys ! (k - (j+1)) = l ! k" using False assm by auto moreover have "k - (j+1) < length ?ys" using False assm by auto ultimately show ?thesis by (metis (full_types) filter_empty_conv length_greater_0_conv lka nth_mem) qed moreover have "?xs @ ?ys = l" using append_take_drop_id by blast ultimately have "length (filter ((=) a) l) > 1" by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 asm filter_append length_append less_add_eq_less less_one nat_neq_iff) then show False using asm by simp qed qed qed ultimately show "\!j. j < length l \ l ! j = a" by blast qed end diff --git a/thys/CryptHOL/Misc_CryptHOL.thy b/thys/CryptHOL/Misc_CryptHOL.thy --- a/thys/CryptHOL/Misc_CryptHOL.thy +++ b/thys/CryptHOL/Misc_CryptHOL.thy @@ -1,2521 +1,2517 @@ (* Title: Misc_CryptHOL.thy Author: Andreas Lochbihler, ETH Zurich *) section \Miscellaneous library additions\ theory Misc_CryptHOL imports Probabilistic_While.While_SPMF "HOL-Library.Rewrite" "HOL-Library.Simps_Case_Conv" "HOL-Library.Type_Length" "HOL-Eisbach.Eisbach" Coinductive.TLList Monad_Normalisation.Monad_Normalisation Monomorphic_Monad.Monomorphic_Monad Applicative_Lifting.Applicative begin hide_const (open) Henstock_Kurzweil_Integration.negligible declare eq_on_def [simp del] subsection \HOL\ lemma asm_rl_conv: "(PROP P \ PROP P) \ Trueprop True" by(rule equal_intr_rule) iprover+ named_theorems if_distribs "Distributivity theorems for If" lemma if_mono_cong: "\b \ x \ x'; \ b \ y \ y' \ \ If b x y \ If b x' y'" by simp lemma if_cong_then: "\ b = b'; b' \ t = t'; e = e' \ \ If b t e = If b' t' e'" by simp lemma if_False_eq: "\ b \ False; e = e' \ \ If b t e = e'" by auto lemma imp_OO_imp [simp]: "(\) OO (\) = (\)" by auto lemma inj_on_fun_updD: "\ inj_on (f(x := y)) A; x \ A \ \ inj_on f A" by(auto simp add: inj_on_def split: if_split_asm) lemma disjoint_notin1: "\ A \ B = {}; x \ B \ \ x \ A" by auto lemma Least_le_Least: fixes x :: "'a :: wellorder" assumes "Q x" and Q: "\x. Q x \ \y\x. P y" shows "Least P \ Least Q" -proof - - obtain f :: "'a \ 'a" where "\a. \ Q a \ f a \ a \ P (f a)" using Q by moura - moreover have "Q (Least Q)" using \Q x\ by(rule LeastI) - ultimately show ?thesis by (metis (full_types) le_cases le_less less_le_trans not_less_Least) -qed + by (metis assms order_trans wellorder_Least_lemma) lemma is_empty_image [simp]: "Set.is_empty (f ` A) = Set.is_empty A" by(auto simp add: Set.is_empty_def) subsection \Relations\ inductive Imagep :: "('a \ 'b \ bool) \ ('a \ bool) \ 'b \ bool" for R P where ImagepI: "\ P x; R x y \ \ Imagep R P y" lemma r_r_into_tranclp: "\ r x y; r y z \ \ r^++ x z" by(rule tranclp.trancl_into_trancl)(rule tranclp.r_into_trancl) lemma transp_tranclp_id: assumes "transp R" shows "tranclp R = R" proof(intro ext iffI) fix x y assume "R^++ x y" thus "R x y" by induction(blast dest: transpD[OF assms])+ qed simp lemma transp_inv_image: "transp r \ transp (\x y. r (f x) (f y))" using trans_inv_image[where r="{(x, y). r x y}" and f = f] by(simp add: transp_trans inv_image_def) lemma Domainp_conversep: "Domainp R\\ = Rangep R" by(auto) lemma bi_unique_rel_set_bij_betw: assumes unique: "bi_unique R" and rel: "rel_set R A B" shows "\f. bij_betw f A B \ (\x\A. R x (f x))" proof - from assms obtain f where f: "\x. x \ A \ R x (f x)" and B: "\x. x \ A \ f x \ B" apply(atomize_elim) apply(fold all_conj_distrib) apply(subst choice_iff[symmetric]) apply(auto dest: rel_setD1) done have "inj_on f A" by(rule inj_onI)(auto dest!: f dest: bi_uniqueDl[OF unique]) moreover have "f ` A = B" using rel by(auto 4 3 intro: B dest: rel_setD2 f bi_uniqueDr[OF unique]) ultimately have "bij_betw f A B" unfolding bij_betw_def .. thus ?thesis using f by blast qed definition restrict_relp :: "('a \ 'b \ bool) \ ('a \ bool) \ ('b \ bool) \ 'a \ 'b \ bool" ("_ \ (_ \ _)" [53, 54, 54] 53) where "restrict_relp R P Q = (\x y. R x y \ P x \ Q y)" lemma restrict_relp_apply [simp]: "(R \ P \ Q) x y \ R x y \ P x \ Q y" by(simp add: restrict_relp_def) lemma restrict_relpI [intro?]: "\ R x y; P x; Q y \ \ (R \ P \ Q) x y" by(simp add: restrict_relp_def) lemma restrict_relpE [elim?, cases pred]: assumes "(R \ P \ Q) x y" obtains (restrict_relp) "R x y" "P x" "Q y" using assms by(simp add: restrict_relp_def) lemma conversep_restrict_relp [simp]: "(R \ P \ Q)\\ = R\\ \ Q \ P" by(auto simp add: fun_eq_iff) lemma restrict_relp_restrict_relp [simp]: "R \ P \ Q \ P' \ Q' = R \ inf P P' \ inf Q Q'" by(auto simp add: fun_eq_iff) lemma restrict_relp_cong: "\ P = P'; Q = Q'; \x y. \ P x; Q y \ \ R x y = R' x y \ \ R \ P \ Q = R' \ P' \ Q'" by(auto simp add: fun_eq_iff) lemma restrict_relp_cong_simp: "\ P = P'; Q = Q'; \x y. P x =simp=> Q y =simp=> R x y = R' x y \ \ R \ P \ Q = R' \ P' \ Q'" by(rule restrict_relp_cong; simp add: simp_implies_def) lemma restrict_relp_parametric [transfer_rule]: includes lifting_syntax shows "((A ===> B ===> (=)) ===> (A ===> (=)) ===> (B ===> (=)) ===> A ===> B ===> (=)) restrict_relp restrict_relp" unfolding restrict_relp_def[abs_def] by transfer_prover lemma restrict_relp_mono: "\ R \ R'; P \ P'; Q \ Q' \ \ R \ P \ Q \ R' \ P' \ Q'" by(simp add: le_fun_def) lemma restrict_relp_mono': "\ (R \ P \ Q) x y; \ R x y; P x; Q y \ \ R' x y &&& P' x &&& Q' y \ \ (R' \ P' \ Q') x y" by(auto dest: conjunctionD1 conjunctionD2) lemma restrict_relp_DomainpD: "Domainp (R \ P \ Q) x \ Domainp R x \ P x" by(auto simp add: Domainp.simps) lemma restrict_relp_True: "R \ (\_. True) \ (\_. True) = R" by(simp add: fun_eq_iff) lemma restrict_relp_False1: "R \ (\_. False) \ Q = bot" by(simp add: fun_eq_iff) lemma restrict_relp_False2: "R \ P \ (\_. False) = bot" by(simp add: fun_eq_iff) definition rel_prod2 :: "('a \ 'b \ bool) \ 'a \ ('c \ 'b) \ bool" where "rel_prod2 R a = (\(c, b). R a b)" lemma rel_prod2_simps [simp]: "rel_prod2 R a (c, b) \ R a b" by(simp add: rel_prod2_def) lemma restrict_rel_prod: "rel_prod (R \ I1 \ I2) (S \ I1' \ I2') = rel_prod R S \ pred_prod I1 I1' \ pred_prod I2 I2'" by(auto simp add: fun_eq_iff) lemma restrict_rel_prod1: "rel_prod (R \ I1 \ I2) S = rel_prod R S \ pred_prod I1 (\_. True) \ pred_prod I2 (\_. True)" by(simp add: restrict_rel_prod[symmetric] restrict_relp_True) lemma restrict_rel_prod2: "rel_prod R (S \ I1 \ I2) = rel_prod R S \ pred_prod (\_. True) I1 \ pred_prod (\_. True) I2" by(simp add: restrict_rel_prod[symmetric] restrict_relp_True) consts relcompp_witness :: "('a \ 'b \ bool) \ ('b \ 'c \ bool) \ 'a \ 'c \ 'b" specification (relcompp_witness) relcompp_witness1: "(A OO B) (fst xy) (snd xy) \ A (fst xy) (relcompp_witness A B xy)" relcompp_witness2: "(A OO B) (fst xy) (snd xy) \ B (relcompp_witness A B xy) (snd xy)" apply(fold all_conj_distrib) apply(rule choice allI)+ by(auto intro: choice allI) lemmas relcompp_witness[of _ _ "(x, y)" for x y, simplified] = relcompp_witness1 relcompp_witness2 hide_fact (open) relcompp_witness1 relcompp_witness2 lemma relcompp_witness_eq [simp]: "relcompp_witness (=) (=) (x, x) = x" using relcompp_witness(1)[of "(=)" "(=)" x x] by(simp add: eq_OO) subsection \Pairs\ lemma split_apfst [simp]: "case_prod h (apfst f xy) = case_prod (h \ f) xy" by(cases xy) simp definition corec_prod :: "('s \ 'a) \ ('s \ 'b) \ 's \ 'a \ 'b" where "corec_prod f g = (\s. (f s, g s))" lemma corec_prod_apply: "corec_prod f g s = (f s, g s)" by(simp add: corec_prod_def) lemma corec_prod_sel [simp]: shows fst_corec_prod: "fst (corec_prod f g s) = f s" and snd_corec_prod: "snd (corec_prod f g s) = g s" by(simp_all add: corec_prod_apply) lemma apfst_corec_prod [simp]: "apfst h (corec_prod f g s) = corec_prod (h \ f) g s" by(simp add: corec_prod_apply) lemma apsnd_corec_prod [simp]: "apsnd h (corec_prod f g s) = corec_prod f (h \ g) s" by(simp add: corec_prod_apply) lemma map_corec_prod [simp]: "map_prod f g (corec_prod h k s) = corec_prod (f \ h) (g \ k) s" by(simp add: corec_prod_apply) lemma split_corec_prod [simp]: "case_prod h (corec_prod f g s) = h (f s) (g s)" by(simp add: corec_prod_apply) lemma Pair_fst_Unity: "(fst x, ()) = x" by(cases x) simp definition rprodl :: "('a \ 'b) \ 'c \ 'a \ ('b \ 'c)" where "rprodl = (\((a, b), c). (a, (b, c)))" lemma rprodl_simps [simp]: "rprodl ((a, b), c) = (a, (b, c))" by(simp add: rprodl_def) lemma rprodl_parametric [transfer_rule]: includes lifting_syntax shows "(rel_prod (rel_prod A B) C ===> rel_prod A (rel_prod B C)) rprodl rprodl" unfolding rprodl_def by transfer_prover definition lprodr :: "'a \ ('b \ 'c) \ ('a \ 'b) \ 'c" where "lprodr = (\(a, b, c). ((a, b), c))" lemma lprodr_simps [simp]: "lprodr (a, b, c) = ((a, b), c)" by(simp add: lprodr_def) lemma lprodr_parametric [transfer_rule]: includes lifting_syntax shows "(rel_prod A (rel_prod B C) ===> rel_prod (rel_prod A B) C) lprodr lprodr" unfolding lprodr_def by transfer_prover lemma lprodr_inverse [simp]: "rprodl (lprodr x) = x" by(cases x) auto lemma rprodl_inverse [simp]: "lprodr (rprodl x) = x" by(cases x) auto lemma pred_prod_mono' [mono]: "pred_prod A B xy \ pred_prod A' B' xy" if "\x. A x \ A' x" "\y. B y \ B' y" using that by(cases xy) auto fun rel_witness_prod :: "('a \ 'b) \ ('c \ 'd) \ (('a \ 'c) \ ('b \ 'd))" where "rel_witness_prod ((a, b), (c, d)) = ((a, c), (b, d))" subsection \Sums\ lemma islE: assumes "isl x" obtains l where "x = Inl l" using assms by(cases x) auto lemma Inl_in_Plus [simp]: "Inl x \ A <+> B \ x \ A" by auto lemma Inr_in_Plus [simp]: "Inr x \ A <+> B \ x \ B" by auto lemma Inl_eq_map_sum_iff: "Inl x = map_sum f g y \ (\z. y = Inl z \ x = f z)" by(cases y) auto lemma Inr_eq_map_sum_iff: "Inr x = map_sum f g y \ (\z. y = Inr z \ x = g z)" by(cases y) auto lemma inj_on_map_sum [simp]: "\ inj_on f A; inj_on g B \ \ inj_on (map_sum f g) (A <+> B)" proof(rule inj_onI, goal_cases) case (1 x y) then show ?case by(cases x; cases y; auto simp add: inj_on_def) qed lemma inv_into_map_sum: "inv_into (A <+> B) (map_sum f g) x = map_sum (inv_into A f) (inv_into B g) x" if "x \ f ` A <+> g ` B" "inj_on f A" "inj_on g B" using that by(cases rule: PlusE[consumes 1])(auto simp add: inv_into_f_eq f_inv_into_f) fun rsuml :: "('a + 'b) + 'c \ 'a + ('b + 'c)" where "rsuml (Inl (Inl a)) = Inl a" | "rsuml (Inl (Inr b)) = Inr (Inl b)" | "rsuml (Inr c) = Inr (Inr c)" fun lsumr :: "'a + ('b + 'c) \ ('a + 'b) + 'c" where "lsumr (Inl a) = Inl (Inl a)" | "lsumr (Inr (Inl b)) = Inl (Inr b)" | "lsumr (Inr (Inr c)) = Inr c" lemma rsuml_lsumr [simp]: "rsuml (lsumr x) = x" by(cases x rule: lsumr.cases) simp_all lemma lsumr_rsuml [simp]: "lsumr (rsuml x) = x" by(cases x rule: rsuml.cases) simp_all subsection \Option\ declare is_none_bind [simp] lemma case_option_collapse: "case_option x (\_. x) y = x" by(simp split: option.split) lemma indicator_single_Some: "indicator {Some x} (Some y) = indicator {x} y" by(simp split: split_indicator) subsubsection \Predicator and relator\ lemma option_pred_mono_strong: "\ pred_option P x; \a. \ a \ set_option x; P a \ \ P' a \ \ pred_option P' x" by(fact option.pred_mono_strong) lemma option_pred_map [simp]: "pred_option P (map_option f x) = pred_option (P \ f) x" by(fact option.pred_map) lemma option_pred_o_map [simp]: "pred_option P \ map_option f = pred_option (P \ f)" by(simp add: fun_eq_iff) lemma option_pred_bind [simp]: "pred_option P (Option.bind x f) = pred_option (pred_option P \ f) x" by(simp add: pred_option_def) lemma pred_option_conj [simp]: "pred_option (\x. P x \ Q x) = (\x. pred_option P x \ pred_option Q x)" by(auto simp add: pred_option_def) lemma pred_option_top [simp]: "pred_option (\_. True) = (\_. True)" by(fact option.pred_True) lemma rel_option_restrict_relpI [intro?]: "\ rel_option R x y; pred_option P x; pred_option Q y \ \ rel_option (R \ P \ Q) x y" by(erule option.rel_mono_strong) simp lemma rel_option_restrict_relpE [elim?]: assumes "rel_option (R \ P \ Q) x y" obtains "rel_option R x y" "pred_option P x" "pred_option Q y" proof show "rel_option R x y" using assms by(auto elim!: option.rel_mono_strong) have "pred_option (Domainp (R \ P \ Q)) x" using assms by(fold option.Domainp_rel) blast then show "pred_option P x" by(rule option_pred_mono_strong)(blast dest!: restrict_relp_DomainpD) have "pred_option (Domainp (R \ P \ Q)\\) y" using assms by(fold option.Domainp_rel)(auto simp only: option.rel_conversep Domainp_conversep) then show "pred_option Q y" by(rule option_pred_mono_strong)(auto dest!: restrict_relp_DomainpD) qed lemma rel_option_restrict_relp_iff: "rel_option (R \ P \ Q) x y \ rel_option R x y \ pred_option P x \ pred_option Q y" by(blast intro: rel_option_restrict_relpI elim: rel_option_restrict_relpE) lemma option_rel_map_restrict_relp: shows option_rel_map_restrict_relp1: "rel_option (R \ P \ Q) (map_option f x) = rel_option (R \ f \ P \ f \ Q) x" and option_rel_map_restrict_relp2: "rel_option (R \ P \ Q) x (map_option g y) = rel_option ((\x. R x \ g) \ P \ Q \ g) x y" by(simp_all add: option.rel_map restrict_relp_def fun_eq_iff) fun rel_witness_option :: "'a option \ 'b option \ ('a \ 'b) option" where "rel_witness_option (Some x, Some y) = Some (x, y)" | "rel_witness_option (None, None) = None" | "rel_witness_option _ = None" \ \Just to make the definition complete\ lemma rel_witness_option: shows set_rel_witness_option: "\ rel_option A x y; (a, b) \ set_option (rel_witness_option (x, y)) \ \ A a b" and map1_rel_witness_option: "rel_option A x y \ map_option fst (rel_witness_option (x, y)) = x" and map2_rel_witness_option: "rel_option A x y \ map_option snd (rel_witness_option (x, y)) = y" by(cases "(x, y)" rule: rel_witness_option.cases; simp; fail)+ lemma rel_witness_option1: assumes "rel_option A x y" shows "rel_option (\a (a', b). a = a' \ A a' b) x (rel_witness_option (x, y))" using map1_rel_witness_option[OF assms, symmetric] unfolding option.rel_eq[symmetric] option.rel_map by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms]) lemma rel_witness_option2: assumes "rel_option A x y" shows "rel_option (\(a, b') b. b = b' \ A a b') (rel_witness_option (x, y)) y" using map2_rel_witness_option[OF assms] unfolding option.rel_eq[symmetric] option.rel_map by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms]) subsubsection \Orders on option\ abbreviation le_option :: "'a option \ 'a option \ bool" where "le_option \ ord_option (=)" lemma le_option_bind_mono: "\ le_option x y; \a. a \ set_option x \ le_option (f a) (g a) \ \ le_option (Option.bind x f) (Option.bind y g)" by(cases x) simp_all lemma le_option_refl [simp]: "le_option x x" by(cases x) simp_all lemma le_option_conv_option_ord: "le_option = option_ord" by(auto simp add: fun_eq_iff flat_ord_def elim: ord_option.cases) definition pcr_Some :: "('a \ 'b \ bool) \ 'a \ 'b option \ bool" where "pcr_Some R x y \ (\z. y = Some z \ R x z)" lemma pcr_Some_simps [simp]: "pcr_Some R x (Some y) \ R x y" by(simp add: pcr_Some_def) lemma pcr_SomeE [cases pred]: assumes "pcr_Some R x y" obtains (pcr_Some) z where "y = Some z" "R x z" using assms by(auto simp add: pcr_Some_def) subsubsection \Filter for option\ fun filter_option :: "('a \ bool) \ 'a option \ 'a option" where "filter_option P None = None" | "filter_option P (Some x) = (if P x then Some x else None)" lemma set_filter_option [simp]: "set_option (filter_option P x) = {y \ set_option x. P y}" by(cases x) auto lemma filter_map_option: "filter_option P (map_option f x) = map_option f (filter_option (P \ f) x)" by(cases x) simp_all lemma is_none_filter_option [simp]: "Option.is_none (filter_option P x) \ Option.is_none x \ \ P (the x)" by(cases x) simp_all lemma filter_option_eq_Some_iff [simp]: "filter_option P x = Some y \ x = Some y \ P y" by(cases x) auto lemma Some_eq_filter_option_iff [simp]: "Some y = filter_option P x \ x = Some y \ P y" by(cases x) auto lemma filter_conv_bind_option: "filter_option P x = Option.bind x (\y. if P y then Some y else None)" by(cases x) simp_all subsubsection \Assert for option\ primrec assert_option :: "bool \ unit option" where "assert_option True = Some ()" | "assert_option False = None" lemma set_assert_option_conv: "set_option (assert_option b) = (if b then {()} else {})" by(simp) lemma in_set_assert_option [simp]: "x \ set_option (assert_option b) \ b" by(cases b) simp_all subsubsection \Join on options\ definition join_option :: "'a option option \ 'a option" where "join_option x = (case x of Some y \ y | None \ None)" simps_of_case join_simps [simp, code]: join_option_def lemma set_join_option [simp]: "set_option (join_option x) = \(set_option ` set_option x)" by(cases x)(simp_all) lemma in_set_join_option: "x \ set_option (join_option (Some (Some x)))" by simp lemma map_join_option: "map_option f (join_option x) = join_option (map_option (map_option f) x)" by(cases x) simp_all lemma bind_conv_join_option: "Option.bind x f = join_option (map_option f x)" by(cases x) simp_all lemma join_conv_bind_option: "join_option x = Option.bind x id" by(cases x) simp_all lemma join_option_parametric [transfer_rule]: includes lifting_syntax shows "(rel_option (rel_option R) ===> rel_option R) join_option join_option" unfolding join_conv_bind_option[abs_def] by transfer_prover lemma join_option_eq_Some [simp]: "join_option x = Some y \ x = Some (Some y)" by(cases x) simp_all lemma Some_eq_join_option [simp]: "Some y = join_option x \ x = Some (Some y)" by(cases x) auto lemma join_option_eq_None: "join_option x = None \ x = None \ x = Some None" by(cases x) simp_all lemma None_eq_join_option: "None = join_option x \ x = None \ x = Some None" by(cases x) auto subsubsection \Zip on options\ function zip_option :: "'a option \ 'b option \ ('a \ 'b) option" where "zip_option (Some x) (Some y) = Some (x, y)" | "zip_option _ None = None" | "zip_option None _ = None" by pat_completeness auto termination by lexicographic_order lemma zip_option_eq_Some_iff [iff]: "zip_option x y = Some (a, b) \ x = Some a \ y = Some b" by(cases "(x, y)" rule: zip_option.cases) simp_all lemma set_zip_option [simp]: "set_option (zip_option x y) = set_option x \ set_option y" by auto lemma zip_map_option1: "zip_option (map_option f x) y = map_option (apfst f) (zip_option x y)" by(cases "(x, y)" rule: zip_option.cases) simp_all lemma zip_map_option2: "zip_option x (map_option g y) = map_option (apsnd g) (zip_option x y)" by(cases "(x, y)" rule: zip_option.cases) simp_all lemma map_zip_option: "map_option (map_prod f g) (zip_option x y) = zip_option (map_option f x) (map_option g y)" by(simp add: zip_map_option1 zip_map_option2 option.map_comp apfst_def apsnd_def o_def prod.map_comp) lemma zip_conv_bind_option: "zip_option x y = Option.bind x (\x. Option.bind y (\y. Some (x, y)))" by(cases "(x, y)" rule: zip_option.cases) simp_all lemma zip_option_parametric [transfer_rule]: includes lifting_syntax shows "(rel_option R ===> rel_option Q ===> rel_option (rel_prod R Q)) zip_option zip_option" unfolding zip_conv_bind_option[abs_def] by transfer_prover lemma rel_option_eqI [simp]: "rel_option (=) x x" by(simp add: option.rel_eq) subsubsection \Binary supremum on @{typ "'a option"}\ primrec sup_option :: "'a option \ 'a option \ 'a option" where "sup_option x None = x" | "sup_option x (Some y) = (Some y)" lemma sup_option_idem [simp]: "sup_option x x = x" by(cases x) simp_all lemma sup_option_assoc: "sup_option (sup_option x y) z = sup_option x (sup_option y z)" by(cases z) simp_all lemma sup_option_left_idem: "sup_option x (sup_option x y) = sup_option x y" by(rewrite sup_option_assoc[symmetric])(simp) lemmas sup_option_ai = sup_option_assoc sup_option_left_idem lemma sup_option_None [simp]: "sup_option None y = y" by(cases y) simp_all subsubsection \Restriction on @{typ "'a option"}\ primrec (transfer) enforce_option :: "('a \ bool) \ 'a option \ 'a option" where "enforce_option P (Some x) = (if P x then Some x else None)" | "enforce_option P None = None" lemma set_enforce_option [simp]: "set_option (enforce_option P x) = {a \ set_option x. P a}" by(cases x) auto lemma enforce_map_option: "enforce_option P (map_option f x) = map_option f (enforce_option (P \ f) x)" by(cases x) auto lemma enforce_bind_option [simp]: "enforce_option P (Option.bind x f) = Option.bind x (enforce_option P \ f)" by(cases x) auto lemma enforce_option_alt_def: "enforce_option P x = Option.bind x (\a. Option.bind (assert_option (P a)) (\_ :: unit. Some a))" by(cases x) simp_all lemma enforce_option_eq_None_iff [simp]: "enforce_option P x = None \ (\a. x = Some a \ \ P a)" by(cases x) auto lemma enforce_option_eq_Some_iff [simp]: "enforce_option P x = Some y \ x = Some y \ P y" by(cases x) auto lemma Some_eq_enforce_option_iff [simp]: "Some y = enforce_option P x \ x = Some y \ P y" by(cases x) auto lemma enforce_option_top [simp]: "enforce_option \ = id" by(rule ext; rename_tac x; case_tac x; simp) lemma enforce_option_K_True [simp]: "enforce_option (\_. True) x = x" by(cases x) simp_all lemma enforce_option_bot [simp]: "enforce_option \ = (\_. None)" by(simp add: fun_eq_iff) lemma enforce_option_K_False [simp]: "enforce_option (\_. False) x = None" by simp lemma enforce_pred_id_option: "pred_option P x \ enforce_option P x = x" by(cases x) auto subsubsection \Maps\ lemma map_add_apply: "(m1 ++ m2) x = sup_option (m1 x) (m2 x)" by(simp add: map_add_def split: option.split) lemma map_le_map_upd2: "\ f \\<^sub>m g; \y'. f x = Some y' \ y' = y \ \ f \\<^sub>m g(x \ y)" by(cases "x \ dom f")(auto simp add: map_le_def Ball_def) lemma eq_None_iff_not_dom: "f x = None \ x \ dom f" by auto lemma card_ran_le_dom: "finite (dom m) \ card (ran m) \ card (dom m)" by(simp add: ran_alt_def card_image_le) lemma dom_subset_ran_iff: assumes "finite (ran m)" shows "dom m \ ran m \ dom m = ran m" proof assume le: "dom m \ ran m" then have "card (dom m) \ card (ran m)" by(simp add: card_mono assms) moreover have "card (ran m) \ card (dom m)" by(simp add: finite_subset[OF le assms] card_ran_le_dom) ultimately show "dom m = ran m" using card_subset_eq[OF assms le] by simp qed simp text \ We need a polymorphic constant for the empty map such that \transfer_prover\ can use a custom transfer rule for @{const Map.empty} \ definition Map_empty where [simp]: "Map_empty \ Map.empty" lemma map_le_Some1D: "\ m \\<^sub>m m'; m x = Some y \ \ m' x = Some y" by(auto simp add: map_le_def Ball_def) lemma map_le_fun_upd2: "\ f \\<^sub>m g; x \ dom f \ \ f \\<^sub>m g(x := y)" by(auto simp add: map_le_def) lemma map_eqI: "\x\dom m \ dom m'. m x = m' x \ m = m'" by(auto simp add: fun_eq_iff domIff intro: option.expand) subsection \Countable\ lemma countable_lfp: assumes step: "\Y. countable Y \ countable (F Y)" and cont: "Order_Continuity.sup_continuous F" shows "countable (lfp F)" by(subst sup_continuous_lfp[OF cont])(simp add: countable_funpow[OF step]) lemma countable_lfp_apply: assumes step: "\Y x. (\x. countable (Y x)) \ countable (F Y x)" and cont: "Order_Continuity.sup_continuous F" shows "countable (lfp F x)" proof - { fix n have "\x. countable ((F ^^ n) bot x)" by(induct n)(auto intro: step) } thus ?thesis using cont by(simp add: sup_continuous_lfp) qed subsection \ Extended naturals \ lemma idiff_enat_eq_enat_iff: "x - enat n = enat m \ (\k. x = enat k \ k - n = m)" by (cases x) simp_all lemma eSuc_SUP: "A \ {} \ eSuc (\ (f ` A)) = (\x\A. eSuc (f x))" by (subst eSuc_Sup) (simp_all add: image_comp) lemma ereal_of_enat_1: "ereal_of_enat 1 = ereal 1" by (simp add: one_enat_def) lemma ennreal_real_conv_ennreal_of_enat: "ennreal (real n) = ennreal_of_enat n" by (simp add: ennreal_of_nat_eq_real_of_nat) lemma enat_add_sub_same2: "b \ \ \ a + b - b = (a :: enat)" by (cases a; cases b) simp_all lemma enat_sub_add: "y \ x \ x - y + z = x + z - (y :: enat)" by (cases x; cases y; cases z) simp_all lemma SUP_enat_eq_0_iff [simp]: "\ (f ` A) = (0 :: enat) \ (\x\A. f x = 0)" by (simp add: bot_enat_def [symmetric]) lemma SUP_enat_add_left: assumes "I \ {}" shows "(SUP i\I. f i + c :: enat) = (SUP i\I. f i) + c" (is "?lhs = ?rhs") proof(cases "c", rule antisym) case (enat n) show "?lhs \ ?rhs" by(auto 4 3 intro: SUP_upper intro: SUP_least) have "(SUP i\I. f i) \ ?lhs - c" using enat by(auto simp add: enat_add_sub_same2 intro!: SUP_least order_trans[OF _ SUP_upper[THEN enat_minus_mono1]]) note add_right_mono[OF this, of c] also have "\ + c \ ?lhs" using assms by(subst enat_sub_add)(auto intro: SUP_upper2 simp add: enat_add_sub_same2 enat) finally show "?rhs \ ?lhs" . qed(simp add: assms SUP_constant) lemma SUP_enat_add_right: assumes "I \ {}" shows "(SUP i\I. c + f i :: enat) = c + (SUP i\I. f i)" using SUP_enat_add_left[OF assms, of f c] by(simp add: add.commute) lemma iadd_SUP_le_iff: "n + (SUP x\A. f x :: enat) \ y \ (if A = {} then n \ y else \x\A. n + f x \ y)" by(simp add: bot_enat_def SUP_enat_add_right[symmetric] SUP_le_iff) lemma SUP_iadd_le_iff: "(SUP x\A. f x :: enat) + n \ y \ (if A = {} then n \ y else \x\A. f x + n \ y)" using iadd_SUP_le_iff[of n f A y] by(simp add: add.commute) subsection \Extended non-negative reals\ lemma (in finite_measure) nn_integral_indicator_neq_infty: "f -` A \ sets M \ (\\<^sup>+ x. indicator A (f x) \M) \ \" unfolding ennreal_indicator[symmetric] apply(rule integrableD) apply(rule integrable_const_bound[where B=1]) apply(simp_all add: indicator_vimage[symmetric]) done lemma (in finite_measure) nn_integral_indicator_neq_top: "f -` A \ sets M \ (\\<^sup>+ x. indicator A (f x) \M) \ \" by(drule nn_integral_indicator_neq_infty) simp lemma nn_integral_indicator_map: assumes [measurable]: "f \ measurable M N" "{x\space N. P x} \ sets N" shows "(\\<^sup>+x. indicator {x\space N. P x} (f x) \M) = emeasure M {x\space M. P (f x)}" using assms(1)[THEN measurable_space] by (subst nn_integral_indicator[symmetric]) (auto intro!: nn_integral_cong split: split_indicator simp del: nn_integral_indicator) subsection \BNF material\ lemma transp_rel_fun: "\ is_equality Q; transp R \ \ transp (rel_fun Q R)" by(rule transpI)(auto dest: transpD rel_funD simp add: is_equality_def) lemma rel_fun_inf: "inf (rel_fun Q R) (rel_fun Q R') = rel_fun Q (inf R R')" by(rule antisym)(auto elim: rel_fun_mono dest: rel_funD) lemma reflp_fun1: includes lifting_syntax shows "\ is_equality A; reflp B \ \ reflp (A ===> B)" by(simp add: reflp_def rel_fun_def is_equality_def) lemma type_copy_id': "type_definition (\x. x) (\x. x) UNIV" by unfold_locales simp_all lemma type_copy_id: "type_definition id id UNIV" by(simp add: id_def type_copy_id') lemma GrpE [cases pred]: assumes "BNF_Def.Grp A f x y" obtains (Grp) "y = f x" "x \ A" using assms by(simp add: Grp_def) lemma rel_fun_Grp_copy_Abs: includes lifting_syntax assumes "type_definition Rep Abs A" shows "rel_fun (BNF_Def.Grp A Abs) (BNF_Def.Grp B g) = BNF_Def.Grp {f. f ` A \ B} (Rep ---> g)" proof - interpret type_definition Rep Abs A by fact show ?thesis by(auto simp add: rel_fun_def Grp_def fun_eq_iff Abs_inverse Rep_inverse intro!: Rep) qed lemma rel_set_Grp: "rel_set (BNF_Def.Grp A f) = BNF_Def.Grp {B. B \ A} (image f)" by(auto simp add: rel_set_def BNF_Def.Grp_def fun_eq_iff) lemma rel_set_comp_Grp: "rel_set R = (BNF_Def.Grp {x. x \ {(x, y). R x y}} ((`) fst))\\ OO BNF_Def.Grp {x. x \ {(x, y). R x y}} ((`) snd)" apply(auto 4 4 del: ext intro!: ext simp add: BNF_Def.Grp_def intro!: rel_setI intro: rev_bexI) apply(simp add: relcompp_apply) subgoal for A B apply(rule exI[where x="A \ B \ {(x, y). R x y}"]) apply(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI) done done lemma Domainp_Grp: "Domainp (BNF_Def.Grp A f) = (\x. x \ A)" by(auto simp add: fun_eq_iff Grp_def) lemma pred_prod_conj [simp]: shows pred_prod_conj1: "\P Q R. pred_prod (\x. P x \ Q x) R = (\x. pred_prod P R x \ pred_prod Q R x)" and pred_prod_conj2: "\P Q R. pred_prod P (\x. Q x \ R x) = (\x. pred_prod P Q x \ pred_prod P R x)" by(auto simp add: pred_prod.simps) lemma pred_sum_conj [simp]: shows pred_sum_conj1: "\P Q R. pred_sum (\x. P x \ Q x) R = (\x. pred_sum P R x \ pred_sum Q R x)" and pred_sum_conj2: "\P Q R. pred_sum P (\x. Q x \ R x) = (\x. pred_sum P Q x \ pred_sum P R x)" by(auto simp add: pred_sum.simps fun_eq_iff) lemma pred_list_conj [simp]: "list_all (\x. P x \ Q x) = (\x. list_all P x \ list_all Q x)" by(auto simp add: list_all_def) lemma pred_prod_top [simp]: "pred_prod (\_. True) (\_. True) = (\_. True)" by(simp add: pred_prod.simps fun_eq_iff) lemma rel_fun_conversep: includes lifting_syntax shows "(A^--1 ===> B^--1) = (A ===> B)^--1" by(auto simp add: rel_fun_def fun_eq_iff) lemma left_unique_Grp [iff]: "left_unique (BNF_Def.Grp A f) \ inj_on f A" unfolding Grp_def left_unique_def by(auto simp add: inj_on_def) lemma right_unique_Grp [simp, intro!]: "right_unique (BNF_Def.Grp A f)" by(simp add: Grp_def right_unique_def) lemma bi_unique_Grp [iff]: "bi_unique (BNF_Def.Grp A f) \ inj_on f A" by(simp add: bi_unique_alt_def) lemma left_total_Grp [iff]: "left_total (BNF_Def.Grp A f) \ A = UNIV" by(auto simp add: left_total_def Grp_def) lemma right_total_Grp [iff]: "right_total (BNF_Def.Grp A f) \ f ` A = UNIV" by(auto simp add: right_total_def BNF_Def.Grp_def image_def) lemma bi_total_Grp [iff]: "bi_total (BNF_Def.Grp A f) \ A = UNIV \ surj f" by(auto simp add: bi_total_alt_def) lemma left_unique_vimage2p [simp]: "\ left_unique P; inj f \ \ left_unique (BNF_Def.vimage2p f g P)" unfolding vimage2p_Grp by(intro left_unique_OO) simp_all lemma right_unique_vimage2p [simp]: "\ right_unique P; inj g \ \ right_unique (BNF_Def.vimage2p f g P)" unfolding vimage2p_Grp by(intro right_unique_OO) simp_all lemma bi_unique_vimage2p [simp]: "\ bi_unique P; inj f; inj g \ \ bi_unique (BNF_Def.vimage2p f g P)" unfolding bi_unique_alt_def by simp lemma left_total_vimage2p [simp]: "\ left_total P; surj g \ \ left_total (BNF_Def.vimage2p f g P)" unfolding vimage2p_Grp by(intro left_total_OO) simp_all lemma right_total_vimage2p [simp]: "\ right_total P; surj f \ \ right_total (BNF_Def.vimage2p f g P)" unfolding vimage2p_Grp by(intro right_total_OO) simp_all lemma bi_total_vimage2p [simp]: "\ bi_total P; surj f; surj g \ \ bi_total (BNF_Def.vimage2p f g P)" unfolding bi_total_alt_def by simp lemma vimage2p_eq [simp]: "inj f \ BNF_Def.vimage2p f f (=) = (=)" by(auto simp add: vimage2p_def fun_eq_iff inj_on_def) lemma vimage2p_conversep: "BNF_Def.vimage2p f g R^--1 = (BNF_Def.vimage2p g f R)^--1" by(simp add: vimage2p_def fun_eq_iff) lemma rel_fun_refl: "\ A \ (=); (=) \ B \ \ (=) \ rel_fun A B" by(subst fun.rel_eq[symmetric])(rule fun_mono) lemma rel_fun_mono_strong: "\ rel_fun A B f g; A' \ A; \x y. \ x \ f ` {x. Domainp A' x}; y \ g ` {x. Rangep A' x}; B x y \ \ B' x y \ \ rel_fun A' B' f g" by(auto simp add: rel_fun_def) fastforce lemma rel_fun_refl_strong: assumes "A \ (=)" "\x. x \ f ` {x. Domainp A x} \ B x x" shows "rel_fun A B f f" proof - have "rel_fun (=) (=) f f" by(simp add: rel_fun_eq) then show ?thesis using assms(1) by(rule rel_fun_mono_strong) (auto intro: assms(2)) qed lemma Grp_iff: "BNF_Def.Grp B g x y \ y = g x \ x \ B" by(simp add: Grp_def) lemma Rangep_Grp: "Rangep (BNF_Def.Grp A f) = (\x. x \ f ` A)" by(auto simp add: fun_eq_iff Grp_iff) lemma rel_fun_Grp: "rel_fun (BNF_Def.Grp UNIV h)\\ (BNF_Def.Grp A g) = BNF_Def.Grp {f. f ` range h \ A} (map_fun h g)" by(auto simp add: rel_fun_def fun_eq_iff Grp_iff) subsection \Transfer and lifting material\ context includes lifting_syntax begin lemma monotone_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone" unfolding monotone_def[abs_def] by transfer_prover lemma fun_ord_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total C" shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord" unfolding fun_ord_def[abs_def] by transfer_prover lemma Plus_parametric [transfer_rule]: "(rel_set A ===> rel_set B ===> rel_set (rel_sum A B)) (<+>) (<+>)" unfolding Plus_def[abs_def] by transfer_prover lemma pred_fun_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" shows "((A ===> (=)) ===> (B ===> (=)) ===> (A ===> B) ===> (=)) pred_fun pred_fun" unfolding pred_fun_def by(transfer_prover) lemma rel_fun_eq_OO: "((=) ===> A) OO ((=) ===> B) = ((=) ===> A OO B)" by(clarsimp simp add: rel_fun_def fun_eq_iff relcompp.simps) metis end lemma Quotient_set_rel_eq: includes lifting_syntax assumes "Quotient R Abs Rep T" shows "(rel_set T ===> rel_set T ===> (=)) (rel_set R) (=)" proof(rule rel_funI iffI)+ fix A B C D assume AB: "rel_set T A B" and CD: "rel_set T C D" have *: "\x y. R x y = (T x (Abs x) \ T y (Abs y) \ Abs x = Abs y)" "\a b. T a b \ Abs a = b" using assms unfolding Quotient_alt_def by simp_all { assume [simp]: "B = D" thus "rel_set R A C" by(auto 4 4 intro!: rel_setI dest: rel_setD1[OF AB, simplified] rel_setD2[OF AB, simplified] rel_setD2[OF CD] rel_setD1[OF CD] simp add: * elim!: rev_bexI) next assume AC: "rel_set R A C" show "B = D" apply safe apply(drule rel_setD2[OF AB], erule bexE) apply(drule rel_setD1[OF AC], erule bexE) apply(drule rel_setD1[OF CD], erule bexE) apply(simp add: *) apply(drule rel_setD2[OF CD], erule bexE) apply(drule rel_setD2[OF AC], erule bexE) apply(drule rel_setD1[OF AB], erule bexE) apply(simp add: *) done } qed lemma Domainp_eq: "Domainp (=) = (\_. True)" by(simp add: Domainp.simps fun_eq_iff) lemma rel_fun_eq_onpI: "eq_onp (pred_fun P Q) f g \ rel_fun (eq_onp P) (eq_onp Q) f g" by(auto simp add: eq_onp_def rel_fun_def) lemma bi_unique_eq_onp: "bi_unique (eq_onp P)" by(simp add: bi_unique_def eq_onp_def) lemma rel_fun_eq_conversep: includes lifting_syntax shows "(A\\ ===> (=)) = (A ===> (=))\\" by(auto simp add: fun_eq_iff rel_fun_def) lemma rel_fun_comp: "\f g h. rel_fun A B (f \ g) h = rel_fun A (\x. B (f x)) g h" "\f g h. rel_fun A B f (g \ h) = rel_fun A (\x y. B x (g y)) f h" by(auto simp add: rel_fun_def) lemma rel_fun_map_fun1: "rel_fun (BNF_Def.Grp UNIV h)\\ A f g \ rel_fun (=) A (map_fun h id f) g" by(auto simp add: rel_fun_def Grp_def) lemma map_fun2_id: "map_fun f g x = g \ map_fun f id x" by(simp add: map_fun_def o_assoc) lemma map_fun_id2_in: "map_fun g h f = map_fun g id (h \ f)" by(simp add: map_fun_def) lemma Domainp_rel_fun_le: "Domainp (rel_fun A B) \ pred_fun (Domainp A) (Domainp B)" by(auto dest: rel_funD) definition rel_witness_fun :: "('a \ 'b \ bool) \ ('b \ 'c \ bool) \ ('a \ 'd) \ ('c \ 'e) \ ('b \ 'd \ 'e)" where "rel_witness_fun A A' = (\(f, g) b. (f (THE a. A a b), g (THE c. A' b c)))" lemma assumes fg: "rel_fun (A OO A') B f g" and A: "left_unique A" "right_total A" and A': "right_unique A'" "left_total A'" shows rel_witness_fun1: "rel_fun A (\x (x', y). x = x' \ B x' y) f (rel_witness_fun A A' (f, g))" and rel_witness_fun2: "rel_fun A' (\(x, y') y. y = y' \ B x y') (rel_witness_fun A A' (f, g)) g" proof (goal_cases) case 1 have "A x y \ f x = f (THE a. A a y) \ B (f (THE a. A a y)) (g (The (A' y)))" for x y by(rule left_totalE[OF A'(2)]; erule meta_allE[of _ y]; erule exE; frule (1) fg[THEN rel_funD, OF relcomppI]) (auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1]) with 1 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def) next case 2 have "A' x y \ g y = g (The (A' x)) \ B (f (THE a. A a x)) (g (The (A' x)))" for x y by(rule right_totalE[OF A(2), of x]; frule (1) fg[THEN rel_funD, OF relcomppI]) (auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1]) with 2 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def) qed lemma rel_witness_fun_eq [simp]: "rel_witness_fun (=) (=) (f, g) = (\x. (f x, g x))" by(simp add: rel_witness_fun_def) subsection \Arithmetic\ lemma abs_diff_triangle_ineq2: "\a - b :: _ :: ordered_ab_group_add_abs\ \ \a - c\ + \c - b\" by(rule order_trans[OF _ abs_diff_triangle_ineq]) simp lemma (in ordered_ab_semigroup_add) add_left_mono_trans: "\ x \ a + b; b \ c \ \ x \ a + c" by(erule order_trans)(rule add_left_mono) lemma of_nat_le_one_cancel_iff [simp]: fixes n :: nat shows "real n \ 1 \ n \ 1" by linarith lemma (in linordered_semidom) mult_right_le: "c \ 1 \ 0 \ a \ c * a \ a" by(subst mult.commute)(rule mult_left_le) subsection \Chain-complete partial orders and \partial_function\\ lemma fun_ordD: "fun_ord ord f g \ ord (f x) (g x)" by(simp add: fun_ord_def) lemma parallel_fixp_induct_strong: assumes ccpo1: "class.ccpo luba orda (mk_less orda)" and ccpo2: "class.ccpo lubb ordb (mk_less ordb)" and adm: "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\x. P (fst x) (snd x))" and f: "monotone orda orda f" and g: "monotone ordb ordb g" and bot: "P (luba {}) (lubb {})" and step: "\x y. \ orda x (ccpo.fixp luba orda f); ordb y (ccpo.fixp lubb ordb g); P x y \ \ P (f x) (g y)" shows "P (ccpo.fixp luba orda f) (ccpo.fixp lubb ordb g)" proof - let ?P="\x y. orda x (ccpo.fixp luba orda f) \ ordb y (ccpo.fixp lubb ordb g) \ P x y" show ?thesis using ccpo1 ccpo2 _ f g proof(rule parallel_fixp_induct[where P="?P", THEN conjunct2, THEN conjunct2]) note [cont_intro] = admissible_leI[OF ccpo1] ccpo.mcont_const[OF ccpo1] admissible_leI[OF ccpo2] ccpo.mcont_const[OF ccpo2] show "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\xy. ?P (fst xy) (snd xy))" using adm by simp show "?P (luba {}) (lubb {})" using bot by(auto intro: ccpo.ccpo_Sup_least ccpo1 ccpo2 chain_empty) show "?P (f x) (g y)" if "?P x y" for x y using that apply(subst ccpo.fixp_unfold[OF ccpo1 f]) apply(subst ccpo.fixp_unfold[OF ccpo2 g]) apply(auto intro: step monotoneD[OF f] monotoneD[OF g]) done qed qed lemma parallel_fixp_induct_strong_uc: assumes a: "partial_function_definitions orda luba" and b: "partial_function_definitions ordb lubb" and F: "\x. monotone (fun_ord orda) orda (\f. U1 (F (C1 f)) x)" and G: "\y. monotone (fun_ord ordb) ordb (\g. U2 (G (C2 g)) y)" and eq1: "f \ C1 (ccpo.fixp (fun_lub luba) (fun_ord orda) (\f. U1 (F (C1 f))))" and eq2: "g \ C2 (ccpo.fixp (fun_lub lubb) (fun_ord ordb) (\g. U2 (G (C2 g))))" and inverse: "\f. U1 (C1 f) = f" and inverse2: "\g. U2 (C2 g) = g" and adm: "ccpo.admissible (prod_lub (fun_lub luba) (fun_lub lubb)) (rel_prod (fun_ord orda) (fun_ord ordb)) (\x. P (fst x) (snd x))" and bot: "P (\_. luba {}) (\_. lubb {})" and step: "\f' g'. \ \x. orda (U1 f' x) (U1 f x); \y. ordb (U2 g' y) (U2 g y); P (U1 f') (U2 g') \ \ P (U1 (F f')) (U2 (G g'))" shows "P (U1 f) (U2 g)" apply(unfold eq1 eq2 inverse inverse2) apply(rule parallel_fixp_induct_strong[OF partial_function_definitions.ccpo[OF a] partial_function_definitions.ccpo[OF b] adm]) using F apply(simp add: monotone_def fun_ord_def) using G apply(simp add: monotone_def fun_ord_def) apply(simp add: fun_lub_def bot) apply(rule step; simp add: inverse inverse2 eq1 eq2 fun_ordD) done lemmas parallel_fixp_induct_strong_1_1 = parallel_fixp_induct_strong_uc[ of _ _ _ _ "\x. x" _ "\x. x" "\x. x" _ "\x. x", OF _ _ _ _ _ _ refl refl] lemmas parallel_fixp_induct_strong_2_2 = parallel_fixp_induct_strong_uc[ of _ _ _ _ "case_prod" _ "curry" "case_prod" _ "curry", where P="\f g. P (curry f) (curry g)", unfolded case_prod_curry curry_case_prod curry_K, OF _ _ _ _ _ _ refl refl, split_format (complete), unfolded prod.case] for P lemma fixp_induct_option': \ \Stronger induction rule\ fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a option" and C :: "('b \ 'a option) \ 'c" and P :: "'b \ 'a \ bool" assumes mono: "\x. mono_option (\f. U (F (C f)) x)" assumes eq: "f \ C (ccpo.fixp (fun_lub (flat_lub None)) (fun_ord option_ord) (\f. U (F (C f))))" assumes inverse2: "\f. U (C f) = f" assumes step: "\g x y. \ \x y. U g x = Some y \ P x y; U (F g) x = Some y; \x. option_ord (U g x) (U f x) \ \ P x y" assumes defined: "U f x = Some y" shows "P x y" using step defined option.fixp_strong_induct_uc[of U F C, OF mono eq inverse2 option_admissible, of P] unfolding fun_lub_def flat_lub_def fun_ord_def by(simp (no_asm_use)) blast declaration \Partial_Function.init "option'" @{term option.fixp_fun} @{term option.mono_body} @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc} (SOME @{thm fixp_induct_option'})\ lemma bot_fun_least [simp]: "(\_. bot :: 'a :: order_bot) \ x" by(fold bot_fun_def) simp lemma fun_ord_conv_rel_fun: "fun_ord = rel_fun (=)" by(simp add: fun_ord_def fun_eq_iff rel_fun_def) inductive finite_chains :: "('a \ 'a \ bool) \ bool" for ord where finite_chainsI: "(\Y. Complete_Partial_Order.chain ord Y \ finite Y) \ finite_chains ord" lemma finite_chainsD: "\ finite_chains ord; Complete_Partial_Order.chain ord Y \ \ finite Y" by(rule finite_chains.cases) lemma finite_chains_flat_ord [simp, intro!]: "finite_chains (flat_ord x)" proof fix Y assume chain: "Complete_Partial_Order.chain (flat_ord x) Y" show "finite Y" proof(cases "\y \ Y. y \ x") case True then obtain y where y: "y \ Y" and yx: "y \ x" by blast hence "Y \ {x, y}" by(auto dest: chainD[OF chain] simp add: flat_ord_def) thus ?thesis by(rule finite_subset) simp next case False hence "Y \ {x}" by auto thus ?thesis by(rule finite_subset) simp qed qed lemma mcont_finite_chains: assumes finite: "finite_chains ord" and mono: "monotone ord ord' f" and ccpo: "class.ccpo lub ord (mk_less ord)" and ccpo': "class.ccpo lub' ord' (mk_less ord')" shows "mcont lub ord lub' ord' f" proof(intro mcontI contI) fix Y assume chain: "Complete_Partial_Order.chain ord Y" and Y: "Y \ {}" from finite chain have fin: "finite Y" by(rule finite_chainsD) from ccpo chain fin Y have lub: "lub Y \ Y" by(rule ccpo.in_chain_finite) interpret ccpo': ccpo lub' ord' "mk_less ord'" by(rule ccpo') have chain': "Complete_Partial_Order.chain ord' (f ` Y)" using chain by(rule chain_imageI)(rule monotoneD[OF mono]) have "ord' (f (lub Y)) (lub' (f ` Y))" using chain' by(rule ccpo'.ccpo_Sup_upper)(simp add: lub) moreover have "ord' (lub' (f ` Y)) (f (lub Y))" using chain' by(rule ccpo'.ccpo_Sup_least)(blast intro: monotoneD[OF mono] ccpo.ccpo_Sup_upper[OF ccpo chain]) ultimately show "f (lub Y) = lub' (f ` Y)" by(rule ccpo'.order.antisym) qed(fact mono) lemma rel_fun_curry: includes lifting_syntax shows "(A ===> B ===> C) f g \ (rel_prod A B ===> C) (case_prod f) (case_prod g)" by(auto simp add: rel_fun_def) lemma (in ccpo) Sup_image_mono: assumes ccpo: "class.ccpo luba orda lessa" and mono: "monotone orda (\) f" and chain: "Complete_Partial_Order.chain orda A" and "A \ {}" shows "Sup (f ` A) \ (f (luba A))" proof(rule ccpo_Sup_least) from chain show "Complete_Partial_Order.chain (\) (f ` A)" by(rule chain_imageI)(rule monotoneD[OF mono]) fix x assume "x \ f ` A" then obtain y where "x = f y" "y \ A" by blast from \y \ A\ have "orda y (luba A)" by(rule ccpo.ccpo_Sup_upper[OF ccpo chain]) hence "f y \ f (luba A)" by(rule monotoneD[OF mono]) thus "x \ f (luba A)" using \x = f y\ by simp qed lemma (in ccpo) admissible_le_mono: assumes "monotone (\) (\) f" shows "ccpo.admissible Sup (\) (\x. x \ f x)" proof(rule ccpo.admissibleI) fix Y assume chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" and le [rule_format]: "\x\Y. x \ f x" have "\Y \ \(f ` Y)" using chain by(rule ccpo_Sup_least)(rule order_trans[OF le]; blast intro!: ccpo_Sup_upper chain_imageI[OF chain] intro: monotoneD[OF assms]) also have "\ \ f (\Y)" by(rule Sup_image_mono[OF _ assms chain Y, where lessa="(<)"]) unfold_locales finally show "\Y \ \" . qed lemma (in ccpo) fixp_induct_strong2: assumes adm: "ccpo.admissible Sup (\) P" and mono: "monotone (\) (\) f" and bot: "P (\{})" and step: "\x. \ x \ ccpo_class.fixp f; x \ f x; P x \ \ P (f x)" shows "P (ccpo_class.fixp f)" proof(rule fixp_strong_induct[where P="\x. x \ f x \ P x", THEN conjunct2]) show "ccpo.admissible Sup (\) (\x. x \ f x \ P x)" using admissible_le_mono adm by(rule admissible_conj)(rule mono) next show "\{} \ f (\{}) \ P (\{})" by(auto simp add: bot chain_empty intro: ccpo_Sup_least) next fix x assume "x \ ccpo_class.fixp f" "x \ f x \ P x" thus "f x \ f (f x) \ P (f x)" by(auto dest: monotoneD[OF mono] intro: step) qed(rule mono) context partial_function_definitions begin lemma fixp_induct_strong2_uc: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a" and C :: "('b \ 'a) \ 'c" and P :: "('b \ 'a) \ bool" assumes mono: "\x. mono_body (\f. U (F (C f)) x)" and eq: "f \ C (fixp_fun (\f. U (F (C f))))" and inverse: "\f. U (C f) = f" and adm: "ccpo.admissible lub_fun le_fun P" and bot: "P (\_. lub {})" and step: "\f'. \ le_fun (U f') (U f); le_fun (U f') (U (F f')); P (U f') \ \ P (U (F f'))" shows "P (U f)" unfolding eq inverse apply (rule ccpo.fixp_induct_strong2[OF ccpo adm]) apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2] apply (rule_tac f'5="C x" in step) apply (simp_all add: inverse eq) done end lemmas parallel_fixp_induct_2_4 = parallel_fixp_induct_uc[ of _ _ _ _ "case_prod" _ "curry" "\f. case_prod (case_prod (case_prod f))" _ "\f. curry (curry (curry f))", where P="\f g. P (curry f) (curry (curry (curry g)))", unfolded case_prod_curry curry_case_prod curry_K, OF _ _ _ _ _ _ refl refl] for P lemma (in ccpo) fixp_greatest: assumes f: "monotone (\) (\) f" and ge: "\y. f y \ y \ x \ y" shows "x \ ccpo.fixp Sup (\) f" by(rule ge)(simp add: fixp_unfold[OF f, symmetric]) lemma fixp_rolling: assumes "class.ccpo lub1 leq1 (mk_less leq1)" and "class.ccpo lub2 leq2 (mk_less leq2)" and f: "monotone leq1 leq2 f" and g: "monotone leq2 leq1 g" shows "ccpo.fixp lub1 leq1 (\x. g (f x)) = g (ccpo.fixp lub2 leq2 (\x. f (g x)))" proof - interpret c1: ccpo lub1 leq1 "mk_less leq1" by fact interpret c2: ccpo lub2 leq2 "mk_less leq2" by fact show ?thesis proof(rule c1.order.antisym) have fg: "monotone leq2 leq2 (\x. f (g x))" using f g by(rule monotone2monotone) simp_all have gf: "monotone leq1 leq1 (\x. g (f x))" using g f by(rule monotone2monotone) simp_all show "leq1 (c1.fixp (\x. g (f x))) (g (c2.fixp (\x. f (g x))))" using gf by(rule c1.fixp_lowerbound)(subst (2) c2.fixp_unfold[OF fg], simp) show "leq1 (g (c2.fixp (\x. f (g x)))) (c1.fixp (\x. g (f x)))" using gf proof(rule c1.fixp_greatest) fix u assume u: "leq1 (g (f u)) u" have "leq1 (g (c2.fixp (\x. f (g x)))) (g (f u))" by(intro monotoneD[OF g] c2.fixp_lowerbound[OF fg] monotoneD[OF f u]) then show "leq1 (g (c2.fixp (\x. f (g x)))) u" using u by(rule c1.order_trans) qed qed qed lemma fixp_lfp_parametric_eq: includes lifting_syntax assumes f: "\x. lfp.mono_body (\f. F f x)" and g: "\x. lfp.mono_body (\f. G f x)" and param: "((A ===> (=)) ===> A ===> (=)) F G" shows "(A ===> (=)) (lfp.fixp_fun F) (lfp.fixp_fun G)" using f g proof(rule parallel_fixp_induct_1_1[OF complete_lattice_partial_function_definitions complete_lattice_partial_function_definitions _ _ reflexive reflexive, where P="(A ===> (=))"]) show "ccpo.admissible (prod_lub lfp.lub_fun lfp.lub_fun) (rel_prod lfp.le_fun lfp.le_fun) (\x. (A ===> (=)) (fst x) (snd x))" unfolding rel_fun_def by simp show "(A ===> (=)) (\_. \{}) (\_. \{})" by auto show "(A ===> (=)) (F f) (G g)" if "(A ===> (=)) f g" for f g using that by(rule rel_funD[OF param]) qed lemma mono2mono_map_option[THEN option.mono2mono, simp, cont_intro]: shows monotone_map_option: "monotone option_ord option_ord (map_option f)" by(rule monotoneI)(auto simp add: flat_ord_def) lemma mcont2mcont_map_option[THEN option.mcont2mcont, simp, cont_intro]: shows mcont_map_option: "mcont (flat_lub None) option_ord (flat_lub None) option_ord (map_option f)" by(rule mcont_finite_chains[OF _ _ flat_interpretation[THEN ccpo] flat_interpretation[THEN ccpo]]) simp_all lemma mono2mono_set_option [THEN lfp.mono2mono]: shows monotone_set_option: "monotone option_ord (\) set_option" by(auto intro!: monotoneI simp add: option_ord_Some1_iff) lemma mcont2mcont_set_option [THEN lfp.mcont2mcont, cont_intro, simp]: shows mcont_set_option: "mcont (flat_lub None) option_ord Union (\) set_option" by(rule mcont_finite_chains)(simp_all add: monotone_set_option ccpo option.partial_function_definitions_axioms) lemma eadd_gfp_partial_function_mono [partial_function_mono]: "\ monotone (fun_ord (\)) (\) f; monotone (fun_ord (\)) (\) g \ \ monotone (fun_ord (\)) (\) (\x. f x + g x :: enat)" by(rule mono2mono_gfp_eadd) lemma map_option_mono [partial_function_mono]: "mono_option B \ mono_option (\f. map_option g (B f))" unfolding map_conv_bind_option by(rule bind_mono) simp_all subsection \Folding over finite sets\ lemma (in comp_fun_commute) fold_invariant_remove [consumes 1, case_names start step]: assumes fin: "finite A" and start: "I A s" and step: "\x s A'. \ x \ A'; I A' s; A' \ A \ \ I (A' - {x}) (f x s)" shows "I {} (Finite_Set.fold f s A)" proof - define A' where "A' == A" with fin start have "finite A'" "A' \ A" "I A' s" by simp_all thus "I {} (Finite_Set.fold f s A')" proof(induction arbitrary: s) case empty thus ?case by simp next case (insert x A') let ?A' = "insert x A'" have "x \ ?A'" "I ?A' s" "?A' \ A" using insert by auto hence "I (?A' - {x}) (f x s)" by(rule step) with insert have "A' \ A" "I A' (f x s)" by auto hence "I {} (Finite_Set.fold f (f x s) A')" by(rule insert.IH) thus ?case using insert by(simp add: fold_insert2 del: fold_insert) qed qed lemma (in comp_fun_commute) fold_invariant_insert [consumes 1, case_names start step]: assumes fin: "finite A" and start: "I {} s" and step: "\x s A'. \ I A' s; x \ A'; x \ A; A' \ A \ \ I (insert x A') (f x s)" shows "I A (Finite_Set.fold f s A)" using fin start proof(rule fold_invariant_remove[where I="\A'. I (A - A')" and A=A and s=s, simplified]) fix x s A' assume *: "x \ A'" "I (A - A') s" "A' \ A" hence "x \ A - A'" "x \ A" "A - A' \ A" by auto with \I (A - A') s\ have "I (insert x (A - A')) (f x s)" by(rule step) also have "insert x (A - A') = A - (A' - {x})" using * by auto finally show "I \ (f x s)" . qed lemma (in comp_fun_idem) fold_set_union: assumes "finite A" "finite B" shows "Finite_Set.fold f z (A \ B) = Finite_Set.fold f (Finite_Set.fold f z A) B" using assms(2,1) by induction simp_all subsection \Parametrisation of transfer rules\ attribute_setup transfer_parametric = \ Attrib.thm >> (fn parametricity => Thm.rule_attribute [] (fn context => fn transfer_rule => let val ctxt = Context.proof_of context; val thm' = Lifting_Term.parametrize_transfer_rule ctxt transfer_rule in Lifting_Def.generate_parametric_transfer_rule ctxt thm' parametricity end handle Lifting_Term.MERGE_TRANSFER_REL msg => error (Pretty.string_of msg) )) \ "combine transfer rule with parametricity theorem" subsection \Lists\ lemma nth_eq_tlI: "xs ! n = z \ (x # xs) ! Suc n = z" by simp lemma list_all2_append': "length us = length vs \ list_all2 P (xs @ us) (ys @ vs) \ list_all2 P xs ys \ list_all2 P us vs" by(auto simp add: list_all2_append1 list_all2_append2 dest: list_all2_lengthD) definition disjointp :: "('a \ bool) list \ bool" where "disjointp xs = disjoint_family_on (\n. {x. (xs ! n) x}) {0.. disjointp xs; (xs ! n) x; (xs ! m) x; n < length xs; m < length xs \ \ n = m" by(auto 4 3 simp add: disjointp_def disjoint_family_on_def) lemma disjointpD': "\ disjointp xs; P x; Q x; xs ! n = P; xs ! m = Q; n < length xs; m < length xs \ \ n = m" by(auto 4 3 simp add: disjointp_def disjoint_family_on_def) lemma wf_strict_prefix: "wfP strict_prefix" proof - from wf have "wf (inv_image {(x, y). x < y} length)" by(rule wf_inv_image) moreover have "{(x, y). strict_prefix x y} \ inv_image {(x, y). x < y} length" by(auto intro: prefix_length_less) ultimately show ?thesis unfolding wfP_def by(rule wf_subset) qed lemma strict_prefix_setD: "strict_prefix xs ys \ set xs \ set ys" by(auto simp add: strict_prefix_def prefix_def) subsubsection \List of a given length\ inductive_set nlists :: "'a set \ nat \ 'a list set" for A n where nlists: "\ set xs \ A; length xs = n \ \ xs \ nlists A n" hide_fact (open) nlists lemma nlists_alt_def: "nlists A n = {xs. set xs \ A \ length xs = n}" by(auto simp add: nlists.simps) lemma nlists_empty: "nlists {} n = (if n = 0 then {[]} else {})" by(auto simp add: nlists_alt_def) lemma nlists_empty_gt0 [simp]: "n > 0 \ nlists {} n = {}" by(simp add: nlists_empty) lemma nlists_0 [simp]: "nlists A 0 = {[]}" by(auto simp add: nlists_alt_def) lemma Cons_in_nlists_Suc [simp]: "x # xs \ nlists A (Suc n) \ x \ A \ xs \ nlists A n" by(simp add: nlists_alt_def) lemma Nil_in_nlists [simp]: "[] \ nlists A n \ n = 0" by(auto simp add: nlists_alt_def) lemma Cons_in_nlists_iff: "x # xs \ nlists A n \ (\n'. n = Suc n' \ x \ A \ xs \ nlists A n')" by(cases n) simp_all lemma in_nlists_Suc_iff: "xs \ nlists A (Suc n) \ (\x xs'. xs = x # xs' \ x \ A \ xs' \ nlists A n)" by(cases xs) simp_all lemma nlists_Suc: "nlists A (Suc n) = (\x\A. (#) x ` nlists A n)" by(auto 4 3 simp add: in_nlists_Suc_iff intro: rev_image_eqI) lemma replicate_in_nlists [simp, intro]: "x \ A \ replicate n x \ nlists A n" by(simp add: nlists_alt_def set_replicate_conv_if) lemma nlists_eq_empty_iff [simp]: "nlists A n = {} \ n > 0 \ A = {}" using replicate_in_nlists by(cases n)(auto) lemma finite_nlists [simp]: "finite A \ finite (nlists A n)" by(induction n)(simp_all add: nlists_Suc) lemma finite_nlistsD: assumes "finite (nlists A n)" shows "finite A \ n = 0" proof(rule disjCI) assume "n \ 0" then obtain n' where n: "n = Suc n'" by(cases n)auto then have "A = hd ` nlists A n" by(auto 4 4 simp add: nlists_Suc intro: rev_image_eqI rev_bexI) also have "finite \" using assms .. finally show "finite A" . qed lemma finite_nlists_iff: "finite (nlists A n) \ finite A \ n = 0" by(auto dest: finite_nlistsD) lemma card_nlists: "card (nlists A n) = card A ^ n" proof(induction n) case (Suc n) have "card (\x\A. (#) x ` nlists A n) = card A * card (nlists A n)" proof(cases "finite A") case True then show ?thesis by(subst card_UN_disjoint)(auto simp add: card_image inj_on_def) next case False hence "\ finite (\x\A. (#) x ` nlists A n)" unfolding nlists_Suc[symmetric] by(auto dest: finite_nlistsD) then show ?thesis using False by simp qed then show ?case using Suc.IH by(simp add: nlists_Suc) qed simp lemma in_nlists_UNIV: "xs \ nlists UNIV n \ length xs = n" by(simp add: nlists_alt_def) subsubsection \ The type of lists of a given length \ typedef (overloaded) ('a, 'b :: len0) nlist = "nlists (UNIV :: 'a set) (LENGTH('b))" proof show "replicate LENGTH('b) undefined \ ?nlist" by simp qed setup_lifting type_definition_nlist subsection \Streams and infinite lists\ primrec sprefix :: "'a list \ 'a stream \ bool" where sprefix_Nil: "sprefix [] ys = True" | sprefix_Cons: "sprefix (x # xs) ys \ x = shd ys \ sprefix xs (stl ys)" lemma sprefix_append: "sprefix (xs @ ys) zs \ sprefix xs zs \ sprefix ys (sdrop (length xs) zs)" by(induct xs arbitrary: zs) simp_all lemma sprefix_stake_same [simp]: "sprefix (stake n xs) xs" by(induct n arbitrary: xs) simp_all lemma sprefix_same_imp_eq: assumes "sprefix xs ys" "sprefix xs' ys" and "length xs = length xs'" shows "xs = xs'" using assms(3,1,2) by(induct arbitrary: ys rule: list_induct2) auto lemma sprefix_shift_same [simp]: "sprefix xs (xs @- ys)" by(induct xs) simp_all lemma sprefix_shift [simp]: "length xs \ length ys \ sprefix xs (ys @- zs) \ prefix xs ys" by(induct xs arbitrary: ys)(simp, case_tac ys, auto) lemma prefixeq_stake2 [simp]: "prefix xs (stake n ys) \ length xs \ n \ sprefix xs ys" proof(induct xs arbitrary: n ys) case (Cons x xs) thus ?case by(cases ys n rule: stream.exhaust[case_product nat.exhaust]) auto qed simp lemma tlength_eq_infinity_iff: "tlength xs = \ \ \ tfinite xs" including tllist.lifting by transfer(simp add: llength_eq_infty_conv_lfinite) subsection \Monomorphic monads\ context includes lifting_syntax begin local_setup \Local_Theory.map_background_naming (Name_Space.mandatory_path "monad")\ definition bind_option :: "'m fail \ 'a option \ ('a \ 'm) \ 'm" where "bind_option fail x f = (case x of None \ fail | Some x' \ f x')" for fail simps_of_case bind_option_simps [simp]: bind_option_def lemma bind_option_parametric [transfer_rule]: "(M ===> rel_option B ===> (B ===> M) ===> M) bind_option bind_option" unfolding bind_option_def by transfer_prover lemma bind_option_K: "\monad. (x = None \ m = fail) \ bind_option fail x (\_. m) = m" by(cases x) simp_all end lemma bind_option_option [simp]: "monad.bind_option None = Option.bind" by(simp add: monad.bind_option_def fun_eq_iff split: option.split) context monad_fail_hom begin lemma hom_bind_option: "h (monad.bind_option fail1 x f) = monad.bind_option fail2 x (h \ f)" by(cases x)(simp_all) end lemma bind_option_set [simp]: "monad.bind_option fail_set = (\x f. \ (f ` set_option x))" by(simp add: monad.bind_option_def fun_eq_iff split: option.split) lemma run_bind_option_stateT [simp]: "\more. run_state (monad.bind_option (fail_state fail) x f) s = monad.bind_option fail x (\y. run_state (f y) s)" by(cases x) simp_all lemma run_bind_option_envT [simp]: "\more. run_env (monad.bind_option (fail_env fail) x f) s = monad.bind_option fail x (\y. run_env (f y) s)" by(cases x) simp_all subsection \Measures\ declare sets_restrict_space_count_space [measurable_cong] lemma (in sigma_algebra) sets_Collect_countable_Ex1: "(\i :: 'i :: countable. {x \ \. P i x} \ M) \ {x \ \. \!i. P i x} \ M" using sets_Collect_countable_Ex1'[of "UNIV :: 'i set"] by simp lemma pred_countable_Ex1 [measurable]: "(\i :: _ :: countable. Measurable.pred M (\x. P i x)) \ Measurable.pred M (\x. \!i. P i x)" unfolding pred_def by(rule sets.sets_Collect_countable_Ex1) lemma measurable_snd_count_space [measurable]: "A \ B \ snd \ measurable (M1 \\<^sub>M count_space A) (count_space B)" by(auto simp add: measurable_def space_pair_measure snd_vimage_eq_Times Times_Int_Times) lemma integrable_scale_measure [simp]: "\ integrable M f; r < \ \ \ integrable (scale_measure r M) f" for f :: "'a \ 'b::{banach, second_countable_topology}" by(auto simp add: integrable_iff_bounded nn_integral_scale_measure ennreal_mult_less_top) lemma integral_scale_measure: assumes "integrable M f" "r < \" shows "integral\<^sup>L (scale_measure r M) f = enn2real r * integral\<^sup>L M f" using assms apply(subst (1 2) real_lebesgue_integral_def) apply(simp_all add: nn_integral_scale_measure ennreal_enn2real_if) by(auto simp add: ennreal_mult_less_top ennreal_less_top_iff ennreal_mult_eq_top_iff enn2real_mult right_diff_distrib elim!: integrableE) subsection \Sequence space\ lemma (in sequence_space) nn_integral_split: assumes f[measurable]: "f \ borel_measurable S" shows "(\\<^sup>+\. f \ \S) = (\\<^sup>+\. (\\<^sup>+\'. f (comb_seq i \ \') \S) \S)" by (subst PiM_comb_seq[symmetric, where i=i]) (simp add: nn_integral_distr P.nn_integral_fst[symmetric]) lemma (in sequence_space) prob_Collect_split: assumes f[measurable]: "{x\space S. P x} \ sets S" shows "\

(x in S. P x) = (\\<^sup>+x. \

(x' in S. P (comb_seq i x x')) \S)" proof - have "\

(x in S. P x) = (\\<^sup>+x. (\\<^sup>+x'. indicator {x\space S. P x} (comb_seq i x x') \S) \S)" using nn_integral_split[of "indicator {x\space S. P x}"] by (auto simp: emeasure_eq_measure) also have "\ = (\\<^sup>+x. \

(x' in S. P (comb_seq i x x')) \S)" by (intro nn_integral_cong) (auto simp: emeasure_eq_measure nn_integral_indicator_map) finally show ?thesis . qed subsection \Probability mass functions\ lemma measure_map_pmf_conv_distr: "measure_pmf (map_pmf f p) = distr (measure_pmf p) (count_space UNIV) f" by(fact map_pmf_rep_eq) abbreviation coin_pmf :: "bool pmf" where "coin_pmf \ pmf_of_set UNIV" text \The rule @{thm [source] rel_pmf_bindI} is not complete as a program logic.\ notepad begin define x where "x = pmf_of_set {True, False}" define y where "y = pmf_of_set {True, False}" define f where "f x = pmf_of_set {True, False}" for x :: bool define g :: "bool \ bool pmf" where "g = return_pmf" define P :: "bool \ bool \ bool" where "P = (=)" have "rel_pmf P (bind_pmf x f) (bind_pmf y g)" by(simp add: P_def f_def[abs_def] g_def y_def bind_return_pmf' pmf.rel_eq) have "\ R x y" if "\x y. R x y \ rel_pmf P (f x) (g y)" for R x y \ \Only the empty relation satisfies @{thm [source] rel_pmf_bindI}'s second premise.\ proof assume "R x y" hence "rel_pmf P (f x) (g y)" by(rule that) thus False by(auto simp add: P_def f_def g_def rel_pmf_return_pmf2) qed define R where "R x y = False" for x y :: bool have "\ rel_pmf R x y" by(simp add: R_def[abs_def]) end lemma pred_rel_pmf: "\ pred_pmf P p; rel_pmf R p q \ \ pred_pmf (Imagep R P) q" unfolding pred_pmf_def apply(rule ballI) apply(unfold rel_pmf.simps) apply(erule exE conjE)+ apply hypsubst apply(unfold pmf.set_map) apply(erule imageE, hypsubst) apply(drule bspec) apply(erule rev_image_eqI) apply(rule refl) apply(erule Imagep.intros) apply(erule allE)+ apply(erule mp) apply(unfold prod.collapse) apply assumption done lemma pmf_rel_mono': "\ rel_pmf P x y; P \ Q \ \ rel_pmf Q x y" by(drule pmf.rel_mono) (auto) lemma rel_pmf_eqI [simp]: "rel_pmf (=) x x" by(simp add: pmf.rel_eq) lemma rel_pmf_bind_reflI: "(\x. x \ set_pmf p \ rel_pmf R (f x) (g x)) \ rel_pmf R (bind_pmf p f) (bind_pmf p g)" by(rule rel_pmf_bindI[where R="\x y. x = y \ x \ set_pmf p"])(auto intro: rel_pmf_reflI) lemma pmf_pred_mono_strong: "\ pred_pmf P p; \a. \ a \ set_pmf p; P a \ \ P' a \ \ pred_pmf P' p" by(simp add: pred_pmf_def) lemma rel_pmf_restrict_relpI [intro?]: "\ rel_pmf R x y; pred_pmf P x; pred_pmf Q y \ \ rel_pmf (R \ P \ Q) x y" by(erule pmf.rel_mono_strong)(simp add: pred_pmf_def) lemma rel_pmf_restrict_relpE [elim?]: assumes "rel_pmf (R \ P \ Q) x y" obtains "rel_pmf R x y" "pred_pmf P x" "pred_pmf Q y" proof show "rel_pmf R x y" using assms by(auto elim!: pmf.rel_mono_strong) have "pred_pmf (Domainp (R \ P \ Q)) x" using assms by(fold pmf.Domainp_rel) blast then show "pred_pmf P x" by(rule pmf_pred_mono_strong)(blast dest!: restrict_relp_DomainpD) have "pred_pmf (Domainp (R \ P \ Q)\\) y" using assms by(fold pmf.Domainp_rel)(auto simp only: pmf.rel_conversep Domainp_conversep) then show "pred_pmf Q y" by(rule pmf_pred_mono_strong)(auto dest!: restrict_relp_DomainpD) qed lemma rel_pmf_restrict_relp_iff: "rel_pmf (R \ P \ Q) x y \ rel_pmf R x y \ pred_pmf P x \ pred_pmf Q y" by(blast intro: rel_pmf_restrict_relpI elim: rel_pmf_restrict_relpE) lemma rel_pmf_OO_trans [trans]: "\ rel_pmf R p q; rel_pmf S q r \ \ rel_pmf (R OO S) p r" unfolding pmf.rel_compp by blast lemma pmf_pred_map [simp]: "pred_pmf P (map_pmf f p) = pred_pmf (P \ f) p" by(simp add: pred_pmf_def) lemma pred_pmf_bind [simp]: "pred_pmf P (bind_pmf p f) = pred_pmf (pred_pmf P \ f) p" by(simp add: pred_pmf_def) lemma pred_pmf_return [simp]: "pred_pmf P (return_pmf x) = P x" by(simp add: pred_pmf_def) lemma pred_pmf_of_set [simp]: "\ finite A; A \ {} \ \ pred_pmf P (pmf_of_set A) = Ball A P" by(simp add: pred_pmf_def) lemma pred_pmf_of_multiset [simp]: "M \ {#} \ pred_pmf P (pmf_of_multiset M) = Ball (set_mset M) P" by(simp add: pred_pmf_def) lemma pred_pmf_cond [simp]: "set_pmf p \ A \ {} \ pred_pmf P (cond_pmf p A) = pred_pmf (\x. x \ A \ P x) p" by(auto simp add: pred_pmf_def) lemma pred_pmf_pair [simp]: "pred_pmf P (pair_pmf p q) = pred_pmf (\x. pred_pmf (P \ Pair x) q) p" by(simp add: pred_pmf_def) lemma pred_pmf_join [simp]: "pred_pmf P (join_pmf p) = pred_pmf (pred_pmf P) p" by(simp add: pred_pmf_def) lemma pred_pmf_bernoulli [simp]: "\ 0 < p; p < 1 \ \ pred_pmf P (bernoulli_pmf p) = All P" by(simp add: pred_pmf_def) lemma pred_pmf_geometric [simp]: "\ 0 < p; p < 1 \ \ pred_pmf P (geometric_pmf p) = All P" by(simp add: pred_pmf_def set_pmf_geometric) lemma pred_pmf_poisson [simp]: "0 < rate \ pred_pmf P (poisson_pmf rate) = All P" by(simp add: pred_pmf_def) lemma pmf_rel_map_restrict_relp: shows pmf_rel_map_restrict_relp1: "rel_pmf (R \ P \ Q) (map_pmf f p) = rel_pmf (R \ f \ P \ f \ Q) p" and pmf_rel_map_restrict_relp2: "rel_pmf (R \ P \ Q) p (map_pmf g q) = rel_pmf ((\x. R x \ g) \ P \ Q \ g) p q" by(simp_all add: pmf.rel_map restrict_relp_def fun_eq_iff) lemma pred_pmf_conj [simp]: "pred_pmf (\x. P x \ Q x) = (\x. pred_pmf P x \ pred_pmf Q x)" by(auto simp add: pred_pmf_def) lemma pred_pmf_top [simp]: "pred_pmf (\_. True) = (\_. True)" by(simp add: pred_pmf_def) lemma rel_pmf_of_setI: assumes A: "A \ {}" "finite A" and B: "B \ {}" "finite B" and card: "\X. X \ A \ card B * card X \ card A * card {y\B. \x\X. R x y}" shows "rel_pmf R (pmf_of_set A) (pmf_of_set B)" apply(rule rel_pmf_measureI) using assms apply(clarsimp simp add: measure_pmf_of_set card_gt_0_iff field_simps of_nat_mult[symmetric] simp del: of_nat_mult) apply(subst mult.commute) apply(erule meta_allE) apply(erule meta_impE) prefer 2 apply(erule order_trans) apply(auto simp add: card_gt_0_iff intro: card_mono) done consts rel_witness_pmf :: "('a \ 'b \ bool) \ 'a pmf \ 'b pmf \ ('a \ 'b) pmf" specification (rel_witness_pmf) set_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \ set_pmf (rel_witness_pmf A xy) \ {(a, b). A a b}" map1_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \ map_pmf fst (rel_witness_pmf A xy) = fst xy" map2_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \ map_pmf snd (rel_witness_pmf A xy) = snd xy" apply(fold all_conj_distrib imp_conjR) apply(rule choice allI)+ apply(unfold pmf.in_rel) by blast lemmas set_rel_witness_pmf = set_rel_witness_pmf'[of _ "(x, y)" for x y, simplified] lemmas map1_rel_witness_pmf = map1_rel_witness_pmf'[of _ "(x, y)" for x y, simplified] lemmas map2_rel_witness_pmf = map2_rel_witness_pmf'[of _ "(x, y)" for x y, simplified] lemmas rel_witness_pmf = set_rel_witness_pmf map1_rel_witness_pmf map2_rel_witness_pmf lemma rel_witness_pmf1: assumes "rel_pmf A p q" shows "rel_pmf (\a (a', b). a = a' \ A a' b) p (rel_witness_pmf A (p, q))" using map1_rel_witness_pmf[OF assms, symmetric] unfolding pmf.rel_eq[symmetric] pmf.rel_map by(rule pmf.rel_mono_strong)(auto dest: set_rel_witness_pmf[OF assms, THEN subsetD]) lemma rel_witness_pmf2: assumes "rel_pmf A p q" shows "rel_pmf (\(a, b') b. b = b' \ A a b') (rel_witness_pmf A (p, q)) q" using map2_rel_witness_pmf[OF assms] unfolding pmf.rel_eq[symmetric] pmf.rel_map by(rule pmf.rel_mono_strong)(auto dest: set_rel_witness_pmf[OF assms, THEN subsetD]) lemma cond_pmf_of_set: assumes fin: "finite A" and nonempty: "A \ B \ {}" shows "cond_pmf (pmf_of_set A) B = pmf_of_set (A \ B)" (is "?lhs = ?rhs") proof(rule pmf_eqI) from nonempty have A: "A \ {}" by auto show "pmf ?lhs x = pmf ?rhs x" for x by(subst pmf_cond; clarsimp simp add: fin A nonempty measure_pmf_of_set split: split_indicator) qed lemma pair_pmf_of_set: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" shows "pair_pmf (pmf_of_set A) (pmf_of_set B) = pmf_of_set (A \ B)" by(rule pmf_eqI)(clarsimp simp add: pmf_pair assms split: split_indicator) lemma emeasure_cond_pmf: fixes p A defines "q \ cond_pmf p A" assumes "set_pmf p \ A \ {}" shows "emeasure (measure_pmf q) B = emeasure (measure_pmf p) (A \ B) / emeasure (measure_pmf p) A" proof - note [transfer_rule] = cond_pmf.transfer[OF assms(2), folded q_def] interpret pmf_as_measure . show ?thesis by transfer simp qed lemma measure_cond_pmf: "measure (measure_pmf (cond_pmf p A)) B = measure (measure_pmf p) (A \ B) / measure (measure_pmf p) A" if "set_pmf p \ A \ {}" using emeasure_cond_pmf[OF that, of B] that by(auto simp add: measure_pmf.emeasure_eq_measure measure_pmf_posI divide_ennreal) lemma emeasure_measure_pmf_zero_iff: "emeasure (measure_pmf p) s = 0 \ set_pmf p \ s = {}" (is "?lhs = ?rhs") proof - have "?lhs \ (AE x in measure_pmf p. x \ s)" by(subst AE_iff_measurable)(auto) also have "\ = ?rhs" by(auto simp add: AE_measure_pmf_iff) finally show ?thesis . qed subsection \Subprobability mass functions\ lemma ord_spmf_return_spmf1: "ord_spmf R (return_spmf x) p \ lossless_spmf p \ (\y\set_spmf p. R x y)" by(auto simp add: rel_pmf_return_pmf1 ord_option.simps in_set_spmf lossless_iff_set_pmf_None Ball_def) (metis option.exhaust) lemma ord_spmf_conv: "ord_spmf R = rel_spmf R OO ord_spmf (=)" apply(subst pmf.rel_compp[symmetric]) apply(rule arg_cong[where f="rel_pmf"]) apply(rule ext)+ apply(auto elim!: ord_option.cases option.rel_cases intro: option.rel_intros) done lemma ord_spmf_expand: "NO_MATCH (=) R \ ord_spmf R = rel_spmf R OO ord_spmf (=)" by(rule ord_spmf_conv) lemma ord_spmf_eqD_measure: "ord_spmf (=) p q \ measure (measure_spmf p) A \ measure (measure_spmf q) A" by(drule ord_spmf_eqD_measure_spmf)(simp add: le_measure measure_spmf.emeasure_eq_measure) lemma ord_spmf_measureD: assumes "ord_spmf R p q" shows "measure (measure_spmf p) A \ measure (measure_spmf q) {y. \x\A. R x y}" (is "?lhs \ ?rhs") proof - from assms obtain p' where *: "rel_spmf R p p'" and **: "ord_spmf (=) p' q" by(auto simp add: ord_spmf_expand) have "?lhs \ measure (measure_spmf p') {y. \x\A. R x y}" using * by(rule rel_spmf_measureD) also have "\ \ ?rhs" using ** by(rule ord_spmf_eqD_measure) finally show ?thesis . qed lemma ord_spmf_bind_pmfI1: "(\x. x \ set_pmf p \ ord_spmf R (f x) q) \ ord_spmf R (bind_pmf p f) q" apply(rewrite at "ord_spmf _ _ \" bind_return_pmf[symmetric, where f="\_ :: unit. q"]) apply(rule rel_pmf_bindI[where R="\x y. x \ set_pmf p"]) apply(simp_all add: rel_pmf_return_pmf2) done lemma ord_spmf_bind_spmfI1: "(\x. x \ set_spmf p \ ord_spmf R (f x) q) \ ord_spmf R (bind_spmf p f) q" unfolding bind_spmf_def by(rule ord_spmf_bind_pmfI1)(auto split: option.split simp add: in_set_spmf) lemma spmf_of_set_empty: "spmf_of_set {} = return_pmf None" by(simp add: spmf_of_set_def) lemma rel_spmf_of_setI: assumes card: "\X. X \ A \ card B * card X \ card A * card {y\B. \x\X. R x y}" and eq: "(finite A \ A \ {}) \ (finite B \ B \ {})" shows "rel_spmf R (spmf_of_set A) (spmf_of_set B)" using eq by(clarsimp simp add: spmf_of_set_def card rel_pmf_of_setI simp del: spmf_of_pmf_pmf_of_set cong: conj_cong) lemmas map_bind_spmf = map_spmf_bind_spmf lemma nn_integral_measure_spmf_conv_measure_pmf: assumes [measurable]: "f \ borel_measurable (count_space UNIV)" shows "nn_integral (measure_spmf p) f = nn_integral (restrict_space (measure_pmf p) (range Some)) (f \ the)" by(simp add: measure_spmf_def nn_integral_distr o_def) lemma nn_integral_spmf_neq_infinity: "(\\<^sup>+ x. spmf p x \count_space UNIV) \ \" using nn_integral_measure_spmf[where f="\_. 1", of p, symmetric] by simp lemma return_pmf_bind_option: "return_pmf (Option.bind x f) = bind_spmf (return_pmf x) (return_pmf \ f)" by(cases x) simp_all lemma rel_spmf_pos_distr: "rel_spmf A OO rel_spmf B \ rel_spmf (A OO B)" unfolding option.rel_compp pmf.rel_compp .. lemma rel_spmf_OO_trans [trans]: "\ rel_spmf R p q; rel_spmf S q r \ \ rel_spmf (R OO S) p r" by(rule rel_spmf_pos_distr[THEN predicate2D]) auto lemma map_spmf_eq_map_spmf_iff: "map_spmf f p = map_spmf g q \ rel_spmf (\x y. f x = g y) p q" by(simp add: spmf_rel_eq[symmetric] spmf_rel_map) lemma map_spmf_eq_map_spmfI: "rel_spmf (\x y. f x = g y) p q \ map_spmf f p = map_spmf g q" by(simp add: map_spmf_eq_map_spmf_iff) lemma spmf_rel_mono_strong: "\rel_spmf A f g; \x y. \ x \ set_spmf f; y \ set_spmf g; A x y \ \ B x y \ \ rel_spmf B f g" apply(erule pmf.rel_mono_strong) apply(erule option.rel_mono_strong) by(clarsimp simp add: in_set_spmf) lemma set_spmf_eq_empty: "set_spmf p = {} \ p = return_pmf None" by auto (metis restrict_spmf_empty restrict_spmf_trivial) lemma measure_pair_spmf_times: "measure (measure_spmf (pair_spmf p q)) (A \ B) = measure (measure_spmf p) A * measure (measure_spmf q) B" proof - have "emeasure (measure_spmf (pair_spmf p q)) (A \ B) = (\\<^sup>+ x. ennreal (spmf (pair_spmf p q) x) * indicator (A \ B) x \count_space UNIV)" by(simp add: nn_integral_spmf[symmetric] nn_integral_count_space_indicator) also have "\ = (\\<^sup>+ x. (\\<^sup>+ y. (ennreal (spmf p x) * indicator A x) * (ennreal (spmf q y) * indicator B y) \count_space UNIV) \count_space UNIV)" by(subst nn_integral_fst_count_space[symmetric])(auto intro!: nn_integral_cong split: split_indicator simp add: ennreal_mult) also have "\ = (\\<^sup>+ x. ennreal (spmf p x) * indicator A x * emeasure (measure_spmf q) B \count_space UNIV)" by(simp add: nn_integral_cmult nn_integral_spmf[symmetric] nn_integral_count_space_indicator) also have "\ = emeasure (measure_spmf p) A * emeasure (measure_spmf q) B" by(simp add: nn_integral_multc)(simp add: nn_integral_spmf[symmetric] nn_integral_count_space_indicator) finally show ?thesis by(simp add: measure_spmf.emeasure_eq_measure ennreal_mult[symmetric]) qed lemma lossless_spmfD_set_spmf_nonempty: "lossless_spmf p \ set_spmf p \ {}" using set_pmf_not_empty[of p] by(auto simp add: set_spmf_def bind_UNION lossless_iff_set_pmf_None) lemma set_spmf_return_pmf: "set_spmf (return_pmf x) = set_option x" by(cases x) simp_all lemma bind_spmf_pmf_assoc: "bind_spmf (bind_pmf p f) g = bind_pmf p (\x. bind_spmf (f x) g)" by(simp add: bind_spmf_def bind_assoc_pmf) lemma bind_spmf_of_set: "\ finite A; A \ {} \ \ bind_spmf (spmf_of_set A) f = bind_pmf (pmf_of_set A) f" by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set) lemma bind_spmf_map_pmf: "bind_spmf (map_pmf f p) g = bind_pmf p (\x. bind_spmf (return_pmf (f x)) g)" by(simp add: map_pmf_def bind_spmf_def bind_assoc_pmf) lemma rel_spmf_eqI [simp]: "rel_spmf (=) x x" by(simp add: option.rel_eq) lemma set_spmf_map_pmf: "set_spmf (map_pmf f p) = (\x\set_pmf p. set_option (f x))" (* Move up *) by(simp add: set_spmf_def bind_UNION) lemma ord_spmf_return_spmf [simp]: "ord_spmf (=) (return_spmf x) p \ p = return_spmf x" proof - have "p = return_spmf x \ ord_spmf (=) (return_spmf x) p" by simp thus ?thesis by (metis (no_types) ord_option_eq_simps(2) rel_pmf_return_pmf1 rel_pmf_return_pmf2 spmf.leq_antisym) qed declare set_bind_spmf [simp] set_spmf_return_pmf [simp] lemma bind_spmf_pmf_commute: "bind_spmf p (\x. bind_pmf q (f x)) = bind_pmf q (\y. bind_spmf p (\x. f x y))" unfolding bind_spmf_def by(subst bind_commute_pmf)(auto intro: bind_pmf_cong[OF refl] split: option.split) lemma return_pmf_map_option_conv_bind: "return_pmf (map_option f x) = bind_spmf (return_pmf x) (return_spmf \ f)" by(cases x) simp_all lemma lossless_return_pmf_iff [simp]: "lossless_spmf (return_pmf x) \ x \ None" by(cases x) simp_all lemma lossless_map_pmf: "lossless_spmf (map_pmf f p) \ (\x \ set_pmf p. f x \ None)" using image_iff by(fastforce simp add: lossless_iff_set_pmf_None) lemma bind_pmf_spmf_assoc: "g None = return_pmf None \ bind_pmf (bind_spmf p f) g = bind_spmf p (\x. bind_pmf (f x) g)" by(auto simp add: bind_spmf_def bind_assoc_pmf bind_return_pmf fun_eq_iff intro!: arg_cong2[where f=bind_pmf] split: option.split) abbreviation pred_spmf :: "('a \ bool) \ 'a spmf \ bool" where "pred_spmf P \ pred_pmf (pred_option P)" lemma pred_spmf_def: "pred_spmf P p \ (\x\set_spmf p. P x)" by(auto simp add: pred_pmf_def pred_option_def set_spmf_def) lemma spmf_pred_mono_strong: "\ pred_spmf P p; \a. \ a \ set_spmf p; P a \ \ P' a \ \ pred_spmf P' p" by(simp add: pred_spmf_def) lemma spmf_Domainp_rel: "Domainp (rel_spmf R) = pred_spmf (Domainp R)" by(simp add: pmf.Domainp_rel option.Domainp_rel) lemma rel_spmf_restrict_relpI [intro?]: "\ rel_spmf R p q; pred_spmf P p; pred_spmf Q q \ \ rel_spmf (R \ P \ Q) p q" by(erule spmf_rel_mono_strong)(simp add: pred_spmf_def) lemma rel_spmf_restrict_relpE [elim?]: assumes "rel_spmf (R \ P \ Q) x y" obtains "rel_spmf R x y" "pred_spmf P x" "pred_spmf Q y" proof show "rel_spmf R x y" using assms by(auto elim!: spmf_rel_mono_strong) have "pred_spmf (Domainp (R \ P \ Q)) x" using assms by(fold spmf_Domainp_rel) blast then show "pred_spmf P x" by(rule spmf_pred_mono_strong)(blast dest!: restrict_relp_DomainpD) have "pred_spmf (Domainp (R \ P \ Q)\\) y" using assms by(fold spmf_Domainp_rel)(auto simp only: spmf_rel_conversep Domainp_conversep) then show "pred_spmf Q y" by(rule spmf_pred_mono_strong)(auto dest!: restrict_relp_DomainpD) qed lemma rel_spmf_restrict_relp_iff: "rel_spmf (R \ P \ Q) x y \ rel_spmf R x y \ pred_spmf P x \ pred_spmf Q y" by(blast intro: rel_spmf_restrict_relpI elim: rel_spmf_restrict_relpE) lemma spmf_pred_map: "pred_spmf P (map_spmf f p) = pred_spmf (P \ f) p" by(simp) lemma pred_spmf_bind [simp]: "pred_spmf P (bind_spmf p f) = pred_spmf (pred_spmf P \ f) p" by(simp add: pred_spmf_def bind_UNION) lemma pred_spmf_return: "pred_spmf P (return_spmf x) = P x" by simp lemma pred_spmf_return_pmf_None: "pred_spmf P (return_pmf None)" by simp lemma pred_spmf_spmf_of_pmf [simp]: "pred_spmf P (spmf_of_pmf p) = pred_pmf P p" unfolding pred_spmf_def by(simp add: pred_pmf_def) lemma pred_spmf_of_set [simp]: "pred_spmf P (spmf_of_set A) = (finite A \ Ball A P)" by(auto simp add: pred_spmf_def set_spmf_of_set) lemma pred_spmf_assert_spmf [simp]: "pred_spmf P (assert_spmf b) = (b \ P ())" by(cases b) simp_all lemma pred_spmf_pair [simp]: "pred_spmf P (pair_spmf p q) = pred_spmf (\x. pred_spmf (P \ Pair x) q) p" by(simp add: pred_spmf_def) lemma set_spmf_try [simp]: "set_spmf (try_spmf p q) = set_spmf p \ (if lossless_spmf p then {} else set_spmf q)" by(auto simp add: try_spmf_def set_spmf_bind_pmf in_set_spmf lossless_iff_set_pmf_None split: option.splits)(metis option.collapse) lemma try_spmf_bind_out1: "(\x. lossless_spmf (f x)) \ bind_spmf (TRY p ELSE q) f = TRY (bind_spmf p f) ELSE (bind_spmf q f)" apply(clarsimp simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_return_pmf intro!: bind_pmf_cong[OF refl] split: option.split) apply(rewrite in "\ = _" bind_return_pmf'[symmetric]) apply(rule bind_pmf_cong[OF refl]) apply(clarsimp split: option.split simp add: lossless_iff_set_pmf_None) done lemma pred_spmf_try [simp]: "pred_spmf P (try_spmf p q) = (pred_spmf P p \ (\ lossless_spmf p \ pred_spmf P q))" by(auto simp add: pred_spmf_def) lemma pred_spmf_cond [simp]: "pred_spmf P (cond_spmf p A) = pred_spmf (\x. x \ A \ P x) p" by(auto simp add: pred_spmf_def) lemma spmf_rel_map_restrict_relp: shows spmf_rel_map_restrict_relp1: "rel_spmf (R \ P \ Q) (map_spmf f p) = rel_spmf (R \ f \ P \ f \ Q) p" and spmf_rel_map_restrict_relp2: "rel_spmf (R \ P \ Q) p (map_spmf g q) = rel_spmf ((\x. R x \ g) \ P \ Q \ g) p q" by(simp_all add: spmf_rel_map restrict_relp_def) lemma pred_spmf_conj: "pred_spmf (\x. P x \ Q x) = (\x. pred_spmf P x \ pred_spmf Q x)" by simp lemma spmf_of_pmf_parametric [transfer_rule]: includes lifting_syntax shows "(rel_pmf A ===> rel_spmf A) spmf_of_pmf spmf_of_pmf" unfolding spmf_of_pmf_def[abs_def] by transfer_prover lemma mono2mono_return_pmf[THEN spmf.mono2mono, simp, cont_intro]: (* Move to SPMF *) shows monotone_return_pmf: "monotone option_ord (ord_spmf (=)) return_pmf" by(rule monotoneI)(auto simp add: flat_ord_def) lemma mcont2mcont_return_pmf[THEN spmf.mcont2mcont, simp, cont_intro]: (* Move to SPMF *) shows mcont_return_pmf: "mcont (flat_lub None) option_ord lub_spmf (ord_spmf (=)) return_pmf" by(rule mcont_finite_chains[OF _ _ flat_interpretation[THEN ccpo] ccpo_spmf]) simp_all lemma pred_spmf_top: (* Move up *) "pred_spmf (\_. True) = (\_. True)" by(simp) lemma rel_spmf_restrict_relpI' [intro?]: "\ rel_spmf (\x y. P x \ Q y \ R x y) p q; pred_spmf P p; pred_spmf Q q \ \ rel_spmf (R \ P \ Q) p q" by(erule spmf_rel_mono_strong)(simp add: pred_spmf_def) lemma set_spmf_map_pmf_MATCH [simp]: assumes "NO_MATCH (map_option g) f" shows "set_spmf (map_pmf f p) = (\x\set_pmf p. set_option (f x))" by(rule set_spmf_map_pmf) lemma rel_spmf_bindI': "\ rel_spmf A p q; \x y. \ A x y; x \ set_spmf p; y \ set_spmf q \ \ rel_spmf B (f x) (g y) \ \ rel_spmf B (p \ f) (q \ g)" apply(rule rel_spmf_bindI[where R="\x y. A x y \ x \ set_spmf p \ y \ set_spmf q"]) apply(erule spmf_rel_mono_strong; simp) apply simp done definition rel_witness_spmf :: "('a \ 'b \ bool) \ 'a spmf \ 'b spmf \ ('a \ 'b) spmf" where "rel_witness_spmf A = map_pmf rel_witness_option \ rel_witness_pmf (rel_option A)" lemma assumes "rel_spmf A p q" shows rel_witness_spmf1: "rel_spmf (\a (a', b). a = a' \ A a' b) p (rel_witness_spmf A (p, q))" and rel_witness_spmf2: "rel_spmf (\(a, b') b. b = b' \ A a b') (rel_witness_spmf A (p, q)) q" by(auto simp add: pmf.rel_map rel_witness_spmf_def intro: pmf.rel_mono_strong[OF rel_witness_pmf1[OF assms]] rel_witness_option1 pmf.rel_mono_strong[OF rel_witness_pmf2[OF assms]] rel_witness_option2) lemma weight_assert_spmf [simp]: "weight_spmf (assert_spmf b) = indicator {True} b" by(simp split: split_indicator) definition enforce_spmf :: "('a \ bool) \ 'a spmf \ 'a spmf" where "enforce_spmf P = map_pmf (enforce_option P)" lemma enforce_spmf_parametric [transfer_rule]: includes lifting_syntax shows "((A ===> (=)) ===> rel_spmf A ===> rel_spmf A) enforce_spmf enforce_spmf" unfolding enforce_spmf_def by transfer_prover lemma enforce_return_spmf [simp]: "enforce_spmf P (return_spmf x) = (if P x then return_spmf x else return_pmf None)" by(simp add: enforce_spmf_def) lemma enforce_return_pmf_None [simp]: "enforce_spmf P (return_pmf None) = return_pmf None" by(simp add: enforce_spmf_def) lemma enforce_map_spmf: "enforce_spmf P (map_spmf f p) = map_spmf f (enforce_spmf (P \ f) p)" by(simp add: enforce_spmf_def pmf.map_comp o_def enforce_map_option) lemma enforce_bind_spmf [simp]: "enforce_spmf P (bind_spmf p f) = bind_spmf p (enforce_spmf P \ f)" by(auto simp add: enforce_spmf_def bind_spmf_def map_bind_pmf intro!: bind_pmf_cong split: option.split) lemma set_enforce_spmf [simp]: "set_spmf (enforce_spmf P p) = {a \ set_spmf p. P a}" by(auto simp add: enforce_spmf_def in_set_spmf) lemma enforce_spmf_alt_def: "enforce_spmf P p = bind_spmf p (\a. bind_spmf (assert_spmf (P a)) (\_ :: unit. return_spmf a))" by(auto simp add: enforce_spmf_def assert_spmf_def map_pmf_def bind_spmf_def bind_return_pmf intro!: bind_pmf_cong split: option.split) lemma bind_enforce_spmf [simp]: "bind_spmf (enforce_spmf P p) f = bind_spmf p (\x. if P x then f x else return_pmf None)" by(auto simp add: enforce_spmf_alt_def assert_spmf_def intro!: bind_spmf_cong) lemma weight_enforce_spmf: "weight_spmf (enforce_spmf P p) = weight_spmf p - measure (measure_spmf p) {x. \ P x}" (is "?lhs = ?rhs") proof - have "?lhs = LINT x|measure_spmf p. indicator {x. P x} x" by(auto simp add: enforce_spmf_alt_def weight_bind_spmf o_def simp del: Bochner_Integration.integral_indicator intro!: Bochner_Integration.integral_cong split: split_indicator) also have "\ = ?rhs" by(subst measure_spmf.finite_measure_Diff[symmetric])(auto simp add: space_measure_spmf intro!: arg_cong2[where f=measure]) finally show ?thesis . qed lemma lossless_enforce_spmf [simp]: "lossless_spmf (enforce_spmf P p) \ lossless_spmf p \ set_spmf p \ {x. P x}" by(auto simp add: enforce_spmf_alt_def) lemma enforce_spmf_top [simp]: "enforce_spmf \ = id" by(simp add: enforce_spmf_def) lemma enforce_spmf_K_True [simp]: "enforce_spmf (\_. True) p = p" using enforce_spmf_top[THEN fun_cong, of p] by(simp add: top_fun_def) lemma enforce_spmf_bot [simp]: "enforce_spmf \ = (\_. return_pmf None)" by(simp add: enforce_spmf_def fun_eq_iff) lemma enforce_spmf_K_False [simp]: "enforce_spmf (\_. False) p = return_pmf None" using enforce_spmf_bot[THEN fun_cong, of p] by(simp add: bot_fun_def) lemma enforce_pred_id_spmf: "enforce_spmf P p = p" if "pred_spmf P p" proof - have "enforce_spmf P p = map_pmf id p" using that by(auto simp add: enforce_spmf_def enforce_pred_id_option simp del: map_pmf_id intro!: pmf.map_cong_pred[OF refl] elim!: pmf_pred_mono_strong) then show ?thesis by simp qed lemma map_the_spmf_of_pmf [simp]: "map_pmf the (spmf_of_pmf p) = p" by(simp add: spmf_of_pmf_def pmf.map_comp o_def) lemma bind_bind_conv_pair_spmf: "bind_spmf p (\x. bind_spmf q (f x)) = bind_spmf (pair_spmf p q) (\(x, y). f x y)" by(simp add: pair_spmf_alt_def) lemma cond_spmf_spmf_of_set: "cond_spmf (spmf_of_set A) B = spmf_of_set (A \ B)" if "finite A" by(rule spmf_eqI)(auto simp add: spmf_of_set measure_spmf_of_set that split: split_indicator) lemma pair_spmf_of_set: "pair_spmf (spmf_of_set A) (spmf_of_set B) = spmf_of_set (A \ B)" by(rule spmf_eqI)(clarsimp simp add: spmf_of_set card_cartesian_product split: split_indicator) lemma emeasure_cond_spmf: "emeasure (measure_spmf (cond_spmf p A)) B = emeasure (measure_spmf p) (A \ B) / emeasure (measure_spmf p) A" apply(clarsimp simp add: cond_spmf_def emeasure_measure_spmf_conv_measure_pmf emeasure_measure_pmf_zero_iff set_pmf_Int_Some split!: if_split) apply blast apply(subst (asm) emeasure_cond_pmf) by(auto simp add: set_pmf_Int_Some image_Int) lemma measure_cond_spmf: "measure (measure_spmf (cond_spmf p A)) B = measure (measure_spmf p) (A \ B) / measure (measure_spmf p) A" apply(clarsimp simp add: cond_spmf_def measure_measure_spmf_conv_measure_pmf measure_pmf_zero_iff set_pmf_Int_Some split!: if_split) apply(subst (asm) measure_cond_pmf) by(auto simp add: image_Int set_pmf_Int_Some) lemma lossless_cond_spmf [simp]: "lossless_spmf (cond_spmf p A) \ set_spmf p \ A \ {}" by(clarsimp simp add: cond_spmf_def lossless_iff_set_pmf_None set_pmf_Int_Some) lemma measure_spmf_eq_density: "measure_spmf p = density (count_space UNIV) (spmf p)" by(rule measure_eqI)(simp_all add: emeasure_density nn_integral_spmf[symmetric] nn_integral_count_space_indicator) lemma integral_measure_spmf: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes A: "finite A" shows "(\a. a \ set_spmf M \ f a \ 0 \ a \ A) \ (LINT x|measure_spmf M. f x) = (\a\A. spmf M a *\<^sub>R f a)" unfolding measure_spmf_eq_density apply (simp add: integral_density) apply (subst lebesgue_integral_count_space_finite_support) by (auto intro!: finite_subset[OF _ \finite A\] sum.mono_neutral_left simp: spmf_eq_0_set_spmf) lemma image_set_spmf_eq: "f ` set_spmf p = g ` set_spmf q" if "ASSUMPTION (map_spmf f p = map_spmf g q)" using that[unfolded ASSUMPTION_def, THEN arg_cong[where f=set_spmf]] by simp lemma map_spmf_const: "map_spmf (\_. x) p = scale_spmf (weight_spmf p) (return_spmf x)" by(simp add: map_spmf_conv_bind_spmf bind_spmf_const) lemma cond_return_pmf [simp]: "cond_pmf (return_pmf x) A = return_pmf x" if "x \ A" using that by(intro pmf_eqI)(auto simp add: pmf_cond split: split_indicator) lemma cond_return_spmf [simp]: "cond_spmf (return_spmf x) A = (if x \ A then return_spmf x else return_pmf None)" by(simp add: cond_spmf_def) lemma measure_range_Some_eq_weight: "measure (measure_pmf p) (range Some) = weight_spmf p" by (simp add: measure_measure_spmf_conv_measure_pmf space_measure_spmf) lemma restrict_spmf_eq_return_pmf_None [simp]: "restrict_spmf p A = return_pmf None \ set_spmf p \ A = {}" by(auto 4 3 simp add: restrict_spmf_def map_pmf_eq_return_pmf_iff bind_UNION in_set_spmf bind_eq_None_conv option.the_def dest: bspec split: if_split_asm option.split_asm) definition mk_lossless :: "'a spmf \ 'a spmf" where "mk_lossless p = scale_spmf (inverse (weight_spmf p)) p" lemma mk_lossless_idem [simp]: "mk_lossless (mk_lossless p) = mk_lossless p" by(simp add: mk_lossless_def weight_scale_spmf min_def max_def inverse_eq_divide) lemma mk_lossless_return [simp]: "mk_lossless (return_pmf x) = return_pmf x" by(cases x)(simp_all add: mk_lossless_def) lemma mk_lossless_map [simp]: "mk_lossless (map_spmf f p) = map_spmf f (mk_lossless p)" by(simp add: mk_lossless_def map_scale_spmf) lemma spmf_mk_lossless [simp]: "spmf (mk_lossless p) x = spmf p x / weight_spmf p" by(simp add: mk_lossless_def spmf_scale_spmf inverse_eq_divide max_def) lemma set_spmf_mk_lossless [simp]: "set_spmf (mk_lossless p) = set_spmf p" by(simp add: mk_lossless_def set_scale_spmf measure_spmf_zero_iff zero_less_measure_iff) lemma mk_lossless_lossless [simp]: "lossless_spmf p \ mk_lossless p = p" by(simp add: mk_lossless_def lossless_weight_spmfD) lemma mk_lossless_eq_return_pmf_None [simp]: "mk_lossless p = return_pmf None \ p = return_pmf None" proof - have aux: "weight_spmf p = 0 \ spmf p i = 0" for i by(rule antisym, rule order_trans[OF spmf_le_weight]) (auto intro!: order_trans[OF spmf_le_weight]) have[simp]: " spmf (scale_spmf (inverse (weight_spmf p)) p) = spmf (return_pmf None) \ spmf p i = 0" for i by(drule fun_cong[where x=i]) (auto simp add: aux spmf_scale_spmf max_def) show ?thesis by(auto simp add: mk_lossless_def intro: spmf_eqI) qed lemma return_pmf_None_eq_mk_lossless [simp]: "return_pmf None = mk_lossless p \ p = return_pmf None" by(metis mk_lossless_eq_return_pmf_None) lemma mk_lossless_spmf_of_set [simp]: "mk_lossless (spmf_of_set A) = spmf_of_set A" by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set) lemma weight_mk_lossless: "weight_spmf (mk_lossless p) = (if p = return_pmf None then 0 else 1)" by(simp add: mk_lossless_def weight_scale_spmf min_def max_def inverse_eq_divide weight_spmf_eq_0) lemma mk_lossless_parametric [transfer_rule]: includes lifting_syntax shows "(rel_spmf A ===> rel_spmf A) mk_lossless mk_lossless" by(simp add: mk_lossless_def rel_fun_def rel_spmf_weightD rel_spmf_scaleI) lemma rel_spmf_mk_losslessI: "rel_spmf A p q \ rel_spmf A (mk_lossless p) (mk_lossless q)" by(rule mk_lossless_parametric[THEN rel_funD]) lemma rel_spmf_restrict_spmfI: "rel_spmf (\x y. (x \ A \ y \ B \ R x y) \ x \ A \ y \ B) p q \ rel_spmf R (restrict_spmf p A) (restrict_spmf q B)" by(auto simp add: restrict_spmf_def pmf.rel_map elim!: option.rel_cases pmf.rel_mono_strong) lemma cond_spmf_alt: "cond_spmf p A = mk_lossless (restrict_spmf p A)" proof(cases "set_spmf p \ A = {}") case True then show ?thesis by(simp add: cond_spmf_def measure_spmf_zero_iff) next case False show ?thesis by(rule spmf_eqI)(simp add: False cond_spmf_def pmf_cond set_pmf_Int_Some image_iff measure_measure_spmf_conv_measure_pmf[symmetric] spmf_scale_spmf max_def inverse_eq_divide) qed lemma cond_spmf_bind: "cond_spmf (bind_spmf p f) A = mk_lossless (p \ (\x. f x \ A))" by(simp add: cond_spmf_alt restrict_bind_spmf scale_bind_spmf) lemma cond_spmf_UNIV [simp]: "cond_spmf p UNIV = mk_lossless p" by(clarsimp simp add: cond_spmf_alt) lemma cond_pmf_singleton: "cond_pmf p A = return_pmf x" if "set_pmf p \ A = {x}" proof - have[simp]: "set_pmf p \ A = {x} \ x \ A \ measure_pmf.prob p A = pmf p x" by(auto simp add: measure_pmf_single[symmetric] AE_measure_pmf_iff intro!: measure_pmf.finite_measure_eq_AE) have "pmf (cond_pmf p A) i = pmf (return_pmf x) i" for i using that by(auto simp add: pmf_cond measure_pmf_zero_iff pmf_eq_0_set_pmf split: split_indicator) then show ?thesis by(rule pmf_eqI) qed definition cond_spmf_fst :: "('a \ 'b) spmf \ 'a \ 'b spmf" where "cond_spmf_fst p a = map_spmf snd (cond_spmf p ({a} \ UNIV))" lemma cond_spmf_fst_return_spmf [simp]: "cond_spmf_fst (return_spmf (x, y)) x = return_spmf y" by(simp add: cond_spmf_fst_def) lemma cond_spmf_fst_map_Pair [simp]: "cond_spmf_fst (map_spmf (Pair x) p) x = mk_lossless p" by(clarsimp simp add: cond_spmf_fst_def spmf.map_comp o_def) lemma cond_spmf_fst_map_Pair' [simp]: "cond_spmf_fst (map_spmf (\y. (x, f y)) p) x = map_spmf f (mk_lossless p)" by(subst spmf.map_comp[where f="Pair x", symmetric, unfolded o_def]) simp lemma cond_spmf_fst_eq_return_None [simp]: "cond_spmf_fst p x = return_pmf None \ x \ fst ` set_spmf p" by(auto 4 4 simp add: cond_spmf_fst_def map_pmf_eq_return_pmf_iff in_set_spmf[symmetric] dest: bspec[where x="Some _"] intro: ccontr rev_image_eqI) lemma cond_spmf_fst_map_Pair1: "cond_spmf_fst (map_spmf (\x. (f x, g x)) p) (f x) = return_spmf (g (inv_into (set_spmf p) f (f x)))" if "x \ set_spmf p" "inj_on f (set_spmf p)" proof - let ?foo="\y. map_option (\x. (f x, g x)) -` Some ` ({f y} \ UNIV)" have[simp]: "y \ set_spmf p \ f x = f y \ set_pmf p \ (?foo y) \ {}" for y by(auto simp add: vimage_def image_def in_set_spmf) have[simp]: "y \ set_spmf p \ f x = f y \ map_spmf snd (map_spmf (\x. (f x, g x)) (cond_pmf p (?foo y))) = return_spmf (g x)" for y using that by(subst cond_pmf_singleton[where x="Some x"]) (auto simp add: in_set_spmf elim: inj_onD) show ?thesis using that by(auto simp add: cond_spmf_fst_def cond_spmf_def) (erule notE, subst cond_map_pmf, simp_all) qed lemma lossless_cond_spmf_fst [simp]: "lossless_spmf (cond_spmf_fst p x) \ x \ fst ` set_spmf p" by(auto simp add: cond_spmf_fst_def intro: rev_image_eqI) lemma cond_spmf_fst_inverse: "bind_spmf (map_spmf fst p) (\x. map_spmf (Pair x) (cond_spmf_fst p x)) = p" (is "?lhs = ?rhs") proof(rule spmf_eqI) fix i :: "'a \ 'b" have *: "({x} \ UNIV \ (Pair x \ snd) -` {i}) = (if x = fst i then {i} else {})" for x by(cases i)auto have "spmf ?lhs i = LINT x|measure_spmf (map_spmf fst p). spmf (map_spmf (Pair x \ snd) (cond_spmf p ({x} \ UNIV))) i" by(auto simp add: spmf_bind spmf.map_comp[symmetric] cond_spmf_fst_def intro!: integral_cong_AE) also have "\ = LINT x|measure_spmf (map_spmf fst p). measure (measure_spmf (cond_spmf p ({x} \ UNIV))) ((Pair x \ snd) -` {i})" by(rule integral_cong_AE)(auto simp add: spmf_map) also have "\ = LINT x|measure_spmf (map_spmf fst p). measure (measure_spmf p) ({x} \ UNIV \ (Pair x \ snd) -` {i}) / measure (measure_spmf p) ({x} \ UNIV)" by(rule integral_cong_AE; clarsimp simp add: measure_cond_spmf) also have "\ = spmf (map_spmf fst p) (fst i) * spmf p i / measure (measure_spmf p) ({fst i} \ UNIV)" by(simp add: * if_distrib[where f="measure (measure_spmf _)"] cong: if_cong) (subst integral_measure_spmf[where A="{fst i}"]; auto split: if_split_asm simp add: spmf_conv_measure_spmf) also have "\ = spmf p i" by(clarsimp simp add: spmf_map vimage_fst)(metis (no_types, lifting) Int_insert_left_if1 in_set_spmf_iff_spmf insertI1 insert_UNIV insert_absorb insert_not_empty measure_spmf_zero_iff mem_Sigma_iff prod.collapse) finally show "spmf ?lhs i = spmf ?rhs i" . qed subsubsection \Embedding of @{typ "'a option"} into @{typ "'a spmf"}\ text \This theoretically follows from the embedding between @{typ "_ id"} into @{typ "_ prob"} and the isomorphism between @{typ "(_, _ prob) optionT"} and @{typ "_ spmf"}, but we would only get the monomorphic version via this connection. So we do it directly. \ lemma bind_option_spmf_monad [simp]: "monad.bind_option (return_pmf None) x = bind_spmf (return_pmf x)" by(cases x)(simp_all add: fun_eq_iff) locale option_to_spmf begin text \ We have to get the embedding into the lifting package such that we can use the parametrisation of transfer rules. \ definition the_pmf :: "'a pmf \ 'a" where "the_pmf p = (THE x. p = return_pmf x)" lemma the_pmf_return [simp]: "the_pmf (return_pmf x) = x" by(simp add: the_pmf_def) lemma type_definition_option_spmf: "type_definition return_pmf the_pmf {x. \y :: 'a option. x = return_pmf y}" by unfold_locales(auto) context begin private setup_lifting type_definition_option_spmf abbreviation cr_spmf_option where "cr_spmf_option \ cr_option" abbreviation pcr_spmf_option where "pcr_spmf_option \ pcr_option" lemmas Quotient_spmf_option = Quotient_option and cr_spmf_option_def = cr_option_def and pcr_spmf_option_bi_unique = option.bi_unique and Domainp_pcr_spmf_option = option.domain and Domainp_pcr_spmf_option_eq = option.domain_eq and Domainp_pcr_spmf_option_par = option.domain_par and Domainp_pcr_spmf_option_left_total = option.domain_par_left_total and pcr_spmf_option_left_unique = option.left_unique and pcr_spmf_option_cr_eq = option.pcr_cr_eq and pcr_spmf_option_return_pmf_transfer = option.rep_transfer and pcr_spmf_option_right_total = option.right_total and pcr_spmf_option_right_unique = option.right_unique and pcr_spmf_option_def = pcr_option_def bundle spmf_option_lifting = [[Lifting.lifting_restore_internal "Misc_CryptHOL.option.lifting"]] end context includes lifting_syntax begin lemma return_option_spmf_transfer [transfer_parametric return_spmf_parametric, transfer_rule]: "((=) ===> cr_spmf_option) return_spmf Some" by(rule rel_funI)(simp add: cr_spmf_option_def) lemma map_option_spmf_transfer [transfer_parametric map_spmf_parametric, transfer_rule]: "(((=) ===> (=)) ===> cr_spmf_option ===> cr_spmf_option) map_spmf map_option" unfolding rel_fun_eq by(auto simp add: rel_fun_def cr_spmf_option_def) lemma fail_option_spmf_transfer [transfer_parametric return_spmf_None_parametric, transfer_rule]: "cr_spmf_option (return_pmf None) None" by(simp add: cr_spmf_option_def) lemma bind_option_spmf_transfer [transfer_parametric bind_spmf_parametric, transfer_rule]: "(cr_spmf_option ===> ((=) ===> cr_spmf_option) ===> cr_spmf_option) bind_spmf Option.bind" apply(clarsimp simp add: rel_fun_def cr_spmf_option_def) subgoal for x f g by(cases x; simp) done lemma set_option_spmf_transfer [transfer_parametric set_spmf_parametric, transfer_rule]: "(cr_spmf_option ===> rel_set (=)) set_spmf set_option" by(clarsimp simp add: rel_fun_def cr_spmf_option_def rel_set_eq) lemma rel_option_spmf_transfer [transfer_parametric rel_spmf_parametric, transfer_rule]: "(((=) ===> (=) ===> (=)) ===> cr_spmf_option ===> cr_spmf_option ===> (=)) rel_spmf rel_option" unfolding rel_fun_eq by(simp add: rel_fun_def cr_spmf_option_def) end end locale option_le_spmf begin text \ Embedding where only successful computations in the option monad are related to Dirac spmf. \ definition cr_option_le_spmf :: "'a option \ 'a spmf \ bool" where "cr_option_le_spmf x p \ ord_spmf (=) (return_pmf x) p" context includes lifting_syntax begin lemma return_option_le_spmf_transfer [transfer_rule]: "((=) ===> cr_option_le_spmf) (\x. x) return_pmf" by(rule rel_funI)(simp add: cr_option_le_spmf_def ord_option_reflI) lemma map_option_le_spmf_transfer [transfer_rule]: "(((=) ===> (=)) ===> cr_option_le_spmf ===> cr_option_le_spmf) map_option map_spmf" unfolding rel_fun_eq apply(clarsimp simp add: rel_fun_def cr_option_le_spmf_def rel_pmf_return_pmf1 ord_option_map1 ord_option_map2) subgoal for f x p y by(cases x; simp add: ord_option_reflI) done lemma bind_option_le_spmf_transfer [transfer_rule]: "(cr_option_le_spmf ===> ((=) ===> cr_option_le_spmf) ===> cr_option_le_spmf) Option.bind bind_spmf" apply(clarsimp simp add: rel_fun_def cr_option_le_spmf_def) subgoal for x p f g by(cases x; auto 4 3 simp add: rel_pmf_return_pmf1 set_pmf_bind_spmf) done end end interpretation rel_spmf_characterisation by unfold_locales(rule rel_pmf_measureI) lemma if_distrib_bind_spmf1 [if_distribs]: "bind_spmf (if b then x else y) f = (if b then bind_spmf x f else bind_spmf y f)" by simp lemma if_distrib_bind_spmf2 [if_distribs]: "bind_spmf x (\y. if b then f y else g y) = (if b then bind_spmf x f else bind_spmf x g)" by simp lemma rel_spmf_if_distrib [if_distribs]: "rel_spmf R (if b then x else y) (if b then x' else y') \ (b \ rel_spmf R x x') \ (\ b \ rel_spmf R y y')" by(simp) lemma if_distrib_map_spmf [if_distribs]: "map_spmf f (if b then p else q) = (if b then map_spmf f p else map_spmf f q)" by simp lemma if_distrib_restrict_spmf1 [if_distribs]: "restrict_spmf (if b then p else q) A = (if b then restrict_spmf p A else restrict_spmf q A)" by simp end diff --git a/thys/DiscretePricing/CRR_Model.thy b/thys/DiscretePricing/CRR_Model.thy --- a/thys/DiscretePricing/CRR_Model.thy +++ b/thys/DiscretePricing/CRR_Model.thy @@ -1,3252 +1,3208 @@ (* Title: CRR_Model.thy Author: Mnacho Echenim, Univ. Grenoble Alpes *) section \The Cox Ross Rubinstein model\ text \This section defines the Cox-Ross-Rubinstein model of a financial market, and charcterizes a risk-neutral probability space for this market. This, together with the proof that every derivative is attainable, permits to obtain a formula to explicitely compute the fair price of any derivative.\ theory CRR_Model imports Fair_Price begin locale CRR_hyps = prob_grw + rsk_free_asset + fixes stk assumes stocks: "stocks Mkt = {stk, risk_free_asset}" and stk_price: "prices Mkt stk = geom_proc" and S0_positive: "0 < init" and down_positive: "0 < d" and down_lt_up: "d < u" and psgt: "0 < p" and pslt: "p < 1" locale CRR_market = CRR_hyps + fixes G assumes stock_filtration:"G = stoch_proc_filt M geom_proc borel" subsection \Preliminary results on the market\ lemma (in CRR_market) case_asset: assumes "asset \ stocks Mkt" shows "asset = stk \ asset = risk_free_asset" proof (rule ccontr) assume "\ (asset = stk \ asset = risk_free_asset)" hence "asset \ stk \ asset \ risk_free_asset" by simp moreover have "asset \ {stk, risk_free_asset}" using assms stocks by simp ultimately show False by auto qed lemma (in CRR_market) assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows bernoulli_gen_filtration: "filtration N G" and bernoulli_sigma_finite: "\n. sigma_finite_subalgebra N (G n)" proof - show "filtration N G" proof - have "disc_filtr M (stoch_proc_filt M geom_proc borel)" proof (rule stoch_proc_filt_disc_filtr) fix i show "random_variable borel (geom_proc i)" by (simp add: geom_rand_walk_borel_measurable) qed hence "filtration M G" using stock_filtration by (simp add: filtration_def disc_filtr_def) have "filt_equiv nat_filtration M N" using pslt psgt by (simp add: assms bernoulli_stream_equiv) hence "sets N = sets M" unfolding filt_equiv_def by simp thus ?thesis unfolding filtration_def by (metis filtration_def \Filtration.filtration M G\ sets_eq_imp_space_eq subalgebra_def) qed show "\n. sigma_finite_subalgebra N (G n)" using assms unfolding subalgebra_def using filtration_def subalgebra_sigma_finite by (metis \Filtration.filtration N G\ bernoulli_stream_def prob_space.prob_space_stream_space prob_space.subalgebra_sigma_finite prob_space_measure_pmf) qed sublocale CRR_market \ rfr_disc_equity_market _ G proof (unfold_locales) show "disc_filtr M G \ sets (G \) = {{}, space M}" proof show "sets (G \) = {{}, space M}" using infinite_cts_filtration.stoch_proc_filt_triv_init stock_filtration geometric_process geom_rand_walk_borel_adapted by (meson infinite_coin_toss_space_axioms infinite_cts_filtration_axioms.intro infinite_cts_filtration_def init_triv_filt_def) show "disc_filtr M G" by (metis Filtration.filtration_def bernoulli bernoulli_gen_filtration disc_filtr_def psgt pslt) qed show "\asset\stocks Mkt. borel_adapt_stoch_proc G (prices Mkt asset)" proof - have "borel_adapt_stoch_proc G (prices Mkt stk)" using stk_price stock_filtration stoch_proc_filt_adapt by (simp add: stoch_proc_filt_adapt geom_rand_walk_borel_measurable) moreover have "borel_adapt_stoch_proc G (prices Mkt risk_free_asset)" using \disc_filtr M G \ sets (G \) = {{}, space M}\ disc_filtr_prob_space.disc_rfr_proc_borel_adapted disc_filtr_prob_space.intro disc_filtr_prob_space_axioms.intro prob_space_axioms rf_price by fastforce moreover have "disc_filtr_prob_space M G" proof (unfold_locales) show "disc_filtr M G" by (simp add: \disc_filtr M G \ sets (G \) = {{}, space M}\) qed ultimately show ?thesis using stocks by force qed qed lemma (in CRR_market) two_stocks: shows "stk \ risk_free_asset" proof (rule ccontr) assume "\stk \ risk_free_asset" hence "disc_rfr_proc r = prices Mkt stk" using rf_price by simp also have "... = geom_proc" using stk_price by simp finally have eqf: "disc_rfr_proc r = geom_proc" . hence "\w. disc_rfr_proc r 0 w = geom_proc 0 w" by simp hence "1 = init" using geometric_process by simp have eqfs: "\w. disc_rfr_proc r (Suc 0) w = geom_proc (Suc 0) w" using eqf by simp hence "disc_rfr_proc r (Suc 0) (sconst True) = geom_proc (Suc 0) (sconst True)" by simp hence "1+r = u" using geometric_process \1 = init\ by simp have "disc_rfr_proc r (Suc 0) (sconst False) = geom_proc (Suc 0) (sconst False)" using eqfs by simp hence "1+r = d" using geometric_process \1 = init\ by simp show False using \1+r = u\ \1+r = d\ down_lt_up by simp qed lemma (in CRR_market) stock_pf_vp_expand: assumes "stock_portfolio Mkt pf" shows "val_process Mkt pf n w = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" proof - have "val_process Mkt pf n w =(sum (\x. ((prices Mkt) x n w) * (pf x (Suc n) w)) (stocks Mkt))" proof (rule subset_val_process') show "finite (stocks Mkt)" using stocks by auto show "support_set pf \ stocks Mkt" using assms unfolding stock_portfolio_def by simp qed also have "... = (\x\ {stk, risk_free_asset}. ((prices Mkt) x n w) * (pf x (Suc n) w))" using stocks by simp also have "... = prices Mkt stk n w * pf stk (Suc n) w + (\ x\ {risk_free_asset}. ((prices Mkt) x n w) * (pf x (Suc n) w))" by (simp add:two_stocks) also have "... = prices Mkt stk n w * pf stk (Suc n) w + prices Mkt risk_free_asset n w * pf risk_free_asset (Suc n) w" by simp also have "... = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using rf_price stk_price by simp finally show ?thesis . qed lemma (in CRR_market) stock_pf_uvp_expand: assumes "stock_portfolio Mkt pf" shows "cls_val_process Mkt pf (Suc n) w = geom_proc (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" proof - have "cls_val_process Mkt pf (Suc n) w =(sum (\x. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w)) (stocks Mkt))" proof (rule subset_cls_val_process') show "finite (stocks Mkt)" using stocks by auto show "support_set pf \ stocks Mkt" using assms unfolding stock_portfolio_def by simp qed also have "... = (\x\ {stk, risk_free_asset}. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w))" using stocks by simp also have "... = prices Mkt stk (Suc n) w * pf stk (Suc n) w + (\ x\ {risk_free_asset}. ((prices Mkt) x (Suc n) w) * (pf x (Suc n) w))" by (simp add:two_stocks) also have "... = prices Mkt stk (Suc n) w * pf stk (Suc n) w + prices Mkt risk_free_asset (Suc n) w * pf risk_free_asset (Suc n) w" by simp also have "... = geom_proc (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using rf_price stk_price by simp finally show ?thesis . qed lemma (in CRR_market) pos_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "d < 1+r" and "0 < pf stk (Suc n) (spick w n False)" and "val_process Mkt pf n (spick w n False) \ 0" shows "cls_val_process Mkt pf (Suc n) (spick w n False) < 0" proof - define wnf where "wnf = spick w n False" have "cls_val_process Mkt pf (Suc n) (spick w n False) = geom_proc (Suc n) wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms by (simp add:stock_pf_uvp_expand) also have "... = d * geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using geometric_process spickI[of n w False] by simp also have "... = d * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" by simp also have "... < (1+r) * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms geom_rand_walk_strictly_positive S0_positive down_positive down_lt_up by simp also have "... = (1+r) * (geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf)" by (simp add: distrib_left) also have "... = (1+r) * val_process Mkt pf n wnf" using stock_pf_vp_expand assms by simp also have "... \ 0" proof - have "0 < 1+r" using assms down_positive by simp moreover have "val_process Mkt pf n wnf \ 0" using assms unfolding wnf_def by simp ultimately show "(1+r) * (val_process Mkt pf n wnf) \ 0" unfolding wnf_def using less_eq_real_def[of 0 "1+r"] mult_nonneg_nonpos[of "1+r" "val_process Mkt pf n (spick w n False)"] by simp qed finally show ?thesis . qed lemma (in CRR_market) neg_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "1+r < u" and "pf stk (Suc n) (spick w n True) < 0" and "val_process Mkt pf n (spick w n True) \ 0" shows "cls_val_process Mkt pf (Suc n) (spick w n True) < 0" proof - define wnf where "wnf = spick w n True" have "cls_val_process Mkt pf (Suc n) (spick w n True) = geom_proc (Suc n) wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms by (simp add:stock_pf_uvp_expand) also have "... = u * geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r (Suc n) wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using geometric_process spickI[of n w True] by simp also have "... = u * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" by simp also have "... < (1+r) * geom_proc n wnf * pf stk (Suc n) wnf + (1+r) * disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf" unfolding wnf_def using assms geom_rand_walk_strictly_positive S0_positive down_positive down_lt_up by simp also have "... = (1+r) * (geom_proc n wnf * pf stk (Suc n) wnf + disc_rfr_proc r n wnf * pf risk_free_asset (Suc n) wnf)" by (simp add: distrib_left) also have "... = (1+r) * val_process Mkt pf n wnf" using stock_pf_vp_expand assms by simp also have "... \ 0" proof - have "0 < 1+r" using acceptable_rate by simp moreover have "val_process Mkt pf n wnf \ 0" using assms unfolding wnf_def by simp ultimately show "(1+r) * (val_process Mkt pf n wnf) \ 0" unfolding wnf_def using less_eq_real_def[of 0 "1+r"] mult_nonneg_nonpos[of "1+r" "val_process Mkt pf n (spick w n True)"] by simp qed finally show ?thesis . qed lemma (in CRR_market) zero_pf_neg_uvp: assumes "stock_portfolio Mkt pf" and "pf stk (Suc n) w = 0" and "pf risk_free_asset (Suc n) w \ 0" and "val_process Mkt pf n w \ 0" shows "cls_val_process Mkt pf (Suc n) w < 0" proof - have "cls_val_process Mkt pf (Suc n) w = S (Suc n) w * pf stk (Suc n) w + disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using assms by (simp add:stock_pf_uvp_expand) also have "... = disc_rfr_proc r (Suc n) w * pf risk_free_asset (Suc n) w" using assms by simp also have "... = (1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" by simp also have "... < 0" proof - have "0 < 1+r" using acceptable_rate by simp moreover have "0 < disc_rfr_proc r n w" using acceptable_rate by (simp add: disc_rfr_proc_positive) ultimately have "0 < (1+r) * disc_rfr_proc r n w" by simp have 1: "0< pf risk_free_asset (Suc n) w \ 0 <(1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" proof (intro impI) assume "0 < pf risk_free_asset (Suc n) w" thus "0 < (1 + r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using \0 < (1+r) * disc_rfr_proc r n w\ by simp qed have 2: "pf risk_free_asset (Suc n) w < 0 \ (1+r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w < 0" proof (intro impI) assume "pf risk_free_asset (Suc n) w < 0" thus "(1 + r) * disc_rfr_proc r n w * pf risk_free_asset (Suc n) w < 0" using \0 < (1+r) * disc_rfr_proc r n w\ by (simp add:mult_pos_neg) qed have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = geom_proc n w * pf stk (Suc n) w + disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using assms by (simp add:stock_pf_vp_expand) also have "... = disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" using assms by simp finally have "0\ disc_rfr_proc r n w * pf risk_free_asset (Suc n) w" . have "0< pf risk_free_asset (Suc n) w \ pf risk_free_asset (Suc n) w < 0" using assms by linarith thus ?thesis using "2" \0 < disc_rfr_proc r n w\ \disc_rfr_proc r n w * pf risk_free_asset (Suc n) w \ 0\ mult_pos_pos by fastforce qed finally show ?thesis . qed lemma (in CRR_market) neg_pf_exists: assumes "stock_portfolio Mkt pf" and "trading_strategy pf" and "1+r < u" and "d < 1+r" and "val_process Mkt pf n w \ 0" and "pf stk (Suc n) w \ 0 \ pf risk_free_asset (Suc n) w \ 0" shows "\y. cls_val_process Mkt pf (Suc n) y < 0" proof - have "borel_predict_stoch_proc G (pf stk)" proof (rule inc_predict_support_trading_strat') show "trading_strategy pf" using assms by simp show "stk \ support_set pf \ {stk}" by simp qed hence "pf stk (Suc n) \ borel_measurable (G n)" unfolding predict_stoch_proc_def by simp have "val_process Mkt pf n \ borel_measurable (G n)" proof - have "borel_adapt_stoch_proc G (val_process Mkt pf)" using assms using support_adapt_def ats_val_process_adapted readable unfolding stock_portfolio_def by blast thus ?thesis unfolding adapt_stoch_proc_def by simp qed define wn where "wn = pseudo_proj_True n w" show ?thesis proof (cases "pf stk (Suc n) w \ 0") case True show ?thesis proof (cases "pf stk (Suc n) w > 0") case True have "0 0 < pf stk (Suc n) w\ by simp also have "... = pf stk (Suc n) wn" unfolding wn_def using \pf stk (Suc n) \ borel_measurable (G n)\ stoch_proc_subalg_nat_filt[of geom_proc] geometric_process nat_filtration_info stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = pf stk (Suc n) (spick wn n False)" using \pf stk (Suc n) \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed moreover have "0 \ val_process Mkt pf n (spick wn n False)" proof - have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = val_process Mkt pf n wn" unfolding wn_def using \val_process Mkt pf n \ borel_measurable (G n)\ nat_filtration_info stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = val_process Mkt pf n (spick wn n False)" using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed ultimately have "cls_val_process Mkt pf (Suc n) (spick wn n False) < 0" using assms by (simp add:pos_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto next case False have "0 >pf stk (Suc n) (spick wn n True)" proof - have "0 > pf stk (Suc n) w" using \\ 0 < pf stk (Suc n) w\ \pf stk (Suc n) w \ 0\ by simp also have "pf stk (Suc n) w = pf stk (Suc n) wn" unfolding wn_def using \pf stk (Suc n) \ borel_measurable (G n)\ nat_filtration_info stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis comp_apply geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = pf stk (Suc n) (spick wn n True)" using \pf stk (Suc n) \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed moreover have "0 \ val_process Mkt pf n (spick wn n True)" proof - have "0 \ val_process Mkt pf n w" using assms by simp also have "val_process Mkt pf n w = val_process Mkt pf n wn" unfolding wn_def using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) also have "... = val_process Mkt pf n (spick wn n True)" using \val_process Mkt pf n \ borel_measurable (G n)\ comp_def nat_filtration_info pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) finally show ?thesis . qed ultimately have "cls_val_process Mkt pf (Suc n) (spick wn n True) < 0" using assms by (simp add:neg_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto qed next case False hence "pf risk_free_asset (Suc n) w \ 0" using assms by simp hence "cls_val_process Mkt pf (Suc n) w < 0" using False assms by (auto simp add:zero_pf_neg_uvp) thus "\y. cls_val_process Mkt pf (Suc n) y < 0" by auto qed qed lemma (in CRR_market) non_zero_components: assumes "val_process Mkt pf n y \ 0" and "stock_portfolio Mkt pf" shows "pf stk (Suc n) y \ 0 \ pf risk_free_asset (Suc n) y \ 0" proof (rule ccontr) assume "\(pf stk (Suc n) y \ 0 \ pf risk_free_asset (Suc n) y \ 0)" hence "pf stk (Suc n) y = 0" "pf risk_free_asset (Suc n) y = 0" by auto have "val_process Mkt pf n y = geom_proc n y * pf stk (Suc n) y + disc_rfr_proc r n y * pf risk_free_asset (Suc n) y" using \stock_portfolio Mkt pf\ stock_pf_vp_expand[of pf n] by simp also have "... = 0" using \pf stk (Suc n) y = 0\ \pf risk_free_asset (Suc n) y = 0\ by simp finally have "val_process Mkt pf n y = 0" . moreover have "val_process Mkt pf n y \ 0" using assms by simp ultimately show False by simp qed lemma (in CRR_market) neg_pf_Suc: assumes "stock_portfolio Mkt pf" and "trading_strategy pf" and "self_financing Mkt pf" and "1+r < u" and "d < 1+r" and "cls_val_process Mkt pf n w < 0" shows "n \ m \ \y. cls_val_process Mkt pf m y < 0" proof (induct m) case 0 assume "n \ 0" hence "n=0" by simp thus "\y. cls_val_process Mkt pf 0 y < 0" using assms by auto next case (Suc m) assume "n \ Suc m" thus "\y. cls_val_process Mkt pf (Suc m) y < 0" proof (cases "n < Suc m") case False hence "n = Suc m" using \n \ Suc m\ by simp thus "\y. cls_val_process Mkt pf (Suc m) y < 0" using assms by auto next case True hence "n \ m" by simp hence "\y. cls_val_process Mkt pf m y < 0" using Suc by simp from this obtain y where "cls_val_process Mkt pf m y < 0" by auto hence "val_process Mkt pf m y < 0" using assms by (simp add:self_financingE) hence "val_process Mkt pf m y \ 0" by simp have "val_process Mkt pf m y \ 0" using \val_process Mkt pf m y < 0\ by simp hence "pf stk (Suc m) y \ 0 \ pf risk_free_asset (Suc m) y \ 0" using assms non_zero_components by simp thus "\y. cls_val_process Mkt pf (Suc m) y < 0" using neg_pf_exists[of pf m y] assms \val_process Mkt pf m y \ 0\ by simp qed qed lemma (in CRR_market) viable_if: assumes "1+r < u" and "d < 1+r" shows "viable_market Mkt" unfolding viable_market_def proof (rule ccontr) assume "\(\p. stock_portfolio Mkt p \ \ arbitrage_process Mkt p)" hence "\p. stock_portfolio Mkt p \ arbitrage_process Mkt p" by simp from this obtain pf where "stock_portfolio Mkt pf" and "arbitrage_process Mkt pf" by auto have "(\ m. (self_financing Mkt pf) \ (trading_strategy pf) \ (\w \ space M. cls_val_process Mkt pf 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt pf m w) \ 0 < \

(w in M. cls_val_process Mkt pf m w > 0))" using \arbitrage_process Mkt pf\ using arbitrage_processE by simp from this obtain m where "self_financing Mkt pf" and "(trading_strategy pf)" and "(\w \ space M. cls_val_process Mkt pf 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt pf m w)" and "0 < \

(w in M. cls_val_process Mkt pf m w > 0)" by auto have "{w\ space M. cls_val_process Mkt pf m w > 0} \ {}" using \0 < \

(w in M. cls_val_process Mkt pf m w > 0)\ by force hence "\w\ space M. cls_val_process Mkt pf m w > 0" by auto from this obtain y where "y\ space M" and "cls_val_process Mkt pf m y > 0" by auto define A where "A = {n::nat. n \ m \ cls_val_process Mkt pf n y > 0}" have "finite A" unfolding A_def by auto have "m \ A" using \cls_val_process Mkt pf m y > 0\ unfolding A_def by simp hence "A \ {}" by auto hence "Min A \ A" using \finite A\ by simp have "Min A \ m" using \finite A\ \m\ A\ by simp have "0 < Min A" proof - have "cls_val_process Mkt pf 0 y = 0" using \y\ space M\ \\w \ space M. cls_val_process Mkt pf 0 w = 0\ by simp hence "0\ A" unfolding A_def by simp moreover have "0 \ Min A" by simp ultimately show ?thesis using \Min A \ A\ neq0_conv by fastforce qed hence "\l. Suc l = Min A" using Suc_diff_1 by blast from this obtain l where "Suc l = Min A" by auto have "cls_val_process Mkt pf l y \ 0" proof - have "l < Min A" using \Suc l = Min A\ by simp hence "l\ A" using \finite A\ \A \ {}\ by auto moreover have "l \ m" using \Suc l = Min A\ \m\ A\ \finite A\ \A \ {}\ \l < Min A\ by auto ultimately show ?thesis unfolding A_def by auto qed hence "val_process Mkt pf l y \ 0" using \self_financing Mkt pf\ by (simp add:self_financingE) moreover have "pf stk (Suc l) y \ 0 \ pf risk_free_asset (Suc l) y \ 0" proof (rule ccontr) assume "\(pf stk (Suc l) y \ 0 \ pf risk_free_asset (Suc l) y \ 0)" hence "pf stk (Suc l) y = 0" "pf risk_free_asset (Suc l) y = 0" by auto have "cls_val_process Mkt pf (Min A) y = geom_proc (Suc l) y * pf stk (Suc l) y + disc_rfr_proc r (Suc l) y * pf risk_free_asset (Suc l) y" using \stock_portfolio Mkt pf\ \Suc l = Min A\ stock_pf_uvp_expand[of pf l] by simp also have "... = 0" using \pf stk (Suc l) y = 0\ \pf risk_free_asset (Suc l) y = 0\ by simp finally have "cls_val_process Mkt pf (Min A) y = 0" . moreover have "cls_val_process Mkt pf (Min A) y > 0" using \Min A \ A\ unfolding A_def by simp ultimately show False by simp qed ultimately have "\z. cls_val_process Mkt pf (Suc l) z < 0" using assms \stock_portfolio Mkt pf\ \trading_strategy pf\ by (simp add:neg_pf_exists) from this obtain z where "cls_val_process Mkt pf (Suc l) z < 0" by auto hence "\x'. cls_val_process Mkt pf m x' < 0" using neg_pf_Suc assms \trading_strategy pf\ \self_financing Mkt pf\ \Suc l = Min A\ \Min A \ m\ \stock_portfolio Mkt pf\ by simp from this obtain x' where "cls_val_process Mkt pf m x' < 0" by auto have "x'\ space M" using bernoulli_stream_space bernoulli by auto hence "x'\ {w\ space M. \0 \ cls_val_process Mkt pf m w}" using \cls_val_process Mkt pf m x' < 0\ by auto from \AE w in M. 0 \ cls_val_process Mkt pf m w\ obtain N where "{w\ space M. \0 \ cls_val_process Mkt pf m w} \ N" and "emeasure M N = 0" and "N\ sets M" using AE_E by auto have "{w\ space M. (stake m w = stake m x')} \ N" proof fix x assume "x \ {w \ space M. stake m w = stake m x'}" hence "x\ space M" and "stake m x = stake m x'" by auto have "cls_val_process Mkt pf m \ borel_measurable (G m)" proof - have "borel_adapt_stoch_proc G (cls_val_process Mkt pf)" using \trading_strategy pf\ \stock_portfolio Mkt pf\ by (meson support_adapt_def readable stock_portfolio_def subsetCE cls_val_process_adapted) thus ?thesis unfolding adapt_stoch_proc_def by simp qed hence "cls_val_process Mkt pf m x' = cls_val_process Mkt pf m x" using \stake m x = stake m x'\ borel_measurable_stake[of "cls_val_process Mkt pf m" m x x'] pseudo_proj_True_stake_image spickI stoch_proc_subalg_nat_filt[of geom_proc] geometric_process stock_filtration by (metis geom_rand_walk_borel_adapted measurable_from_subalg) hence "cls_val_process Mkt pf m x < 0" using \cls_val_process Mkt pf m x' < 0\ by simp thus "x\ N" using \{w\ space M. \0 \ cls_val_process Mkt pf m w} \ N\ \x\ space M\ \cls_val_process Mkt pf (Suc l) z < 0\ by auto qed moreover have "emeasure M {w\ space M. (stake m w = stake m x')} \ 0" using bernoulli_stream_pref_prob_neq_zero psgt pslt by simp ultimately show False using \emeasure M N = 0\ \N \ events\ emeasure_eq_0 by blast qed lemma (in CRR_market) viable_only_if_d: assumes "viable_market Mkt" shows "d < 1+r" proof (rule ccontr) assume "\ d < 1+r" hence "1+r \ d" by simp define arb_pf where "arb_pf = (\ (x::'a) (n::nat) w. 0::real)(stk:= (\ n w. 1), risk_free_asset := (\ n w. - geom_proc 0 w))" have "support_set arb_pf = {stk, risk_free_asset}" proof show "support_set arb_pf \ {stk, risk_free_asset}" by (simp add: arb_pf_def subset_iff support_set_def) have "stk\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks by simp moreover have "risk_free_asset\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks geometric_process S0_positive by simp ultimately show "{stk, risk_free_asset}\ support_set arb_pf" by simp qed hence "stock_portfolio Mkt arb_pf" using stocks by (simp add: portfolio_def stock_portfolio_def) have "arbitrage_process Mkt arb_pf" proof (rule arbitrage_processI, intro exI conjI) show "self_financing Mkt arb_pf" unfolding arb_pf_def using \support_set arb_pf = {stk, risk_free_asset}\ by (simp add: static_portfolio_self_financing) show "trading_strategy arb_pf" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio arb_pf" unfolding portfolio_def using \support_set arb_pf = {stk, risk_free_asset}\ by simp fix asset assume "asset\ support_set arb_pf" show "borel_predict_stoch_proc G (arb_pf asset)" proof (cases "asset = stk") case True hence "arb_pf asset = (\ n w. 1)" unfolding arb_pf_def by (simp add: two_stocks) show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. 1)\ by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. 1)\ by simp qed qed next case False hence "arb_pf asset = (\ n w. - geom_proc 0 w)" using \support_set arb_pf = {stk, risk_free_asset}\ \asset \ support_set arb_pf\ unfolding arb_pf_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. - geom_proc 0 w)\ geometric_process by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. - geom_proc 0 w)\ geometric_process by simp qed qed qed qed show "\w\space M. cls_val_process Mkt arb_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt arb_pf 0 w = geom_proc 0 w * arb_pf stk (Suc 0) w + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_vp_expand \stock_portfolio Mkt arb_pf\ using \self_financing Mkt arb_pf\ self_financingE by fastforce also have "... = geom_proc 0 w * (1) + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = geom_proc 0 w + arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc 0 w - geom_proc 0 w" unfolding arb_pf_def by simp also have "... = 0" by simp finally show "cls_val_process Mkt arb_pf 0 w = 0" . qed have dev: "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" proof (intro ballI) fix w assume "w\ space M" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w * arb_pf stk (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_uvp_expand \stock_portfolio Mkt arb_pf\ by simp also have "... = geom_proc (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = geom_proc (Suc 0) w + (1+r) * arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" by (simp add:arb_pf_def) finally show "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" . qed have iniT: "\w\ space M. snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w > 0" proof (intro ballI impI) fix w assume "w\ space M" and "snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = u * geom_proc 0 w - (1+r) * geom_proc 0 w" using \snth w 0\ geometric_process by simp also have "... = (u - (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... > 0" using S0_positive \1 + r \ d\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w > 0" . qed have iniF: "\w\ space M. \snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (intro ballI impI) fix w assume "w\ space M" and "\snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w - (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = d * geom_proc 0 w - (1+r) * geom_proc 0 w" using \\snth w 0\ geometric_process by simp also have "... = (d - (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... \ 0" using S0_positive \1 + r \ d\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" . qed have "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof fix w assume "w\ space M" show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (cases "snth w 0") case True thus ?thesis using \w\ space M\ iniT by auto next case False thus ?thesis using \w\ space M\ iniF by simp qed qed thus "AE w in M. 0 \ cls_val_process Mkt arb_pf (Suc 0) w" by simp show "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" proof - have "cls_val_process Mkt arb_pf (Suc 0) \ borel_measurable M" using borel_adapt_stoch_proc_borel_measurable cls_val_process_adapted \trading_strategy arb_pf\ \stock_portfolio Mkt arb_pf\ using support_adapt_def readable unfolding stock_portfolio_def by blast hence set_event:"{w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w} \ sets M" using borel_measurable_iff_greater by blast have "\n. emeasure M {w \ space M. w !! n} = ennreal p" using bernoulli p_gt_0 p_lt_1 bernoulli_stream_component_probability[of M p] by auto hence "emeasure M {w \ space M. w !! 0} = ennreal p" by blast moreover have "{w \ space M. w !! 0} \ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" proof fix w assume "w\ {w \ space M. w !! 0}" hence "w \ space M" and "w !! 0" by auto note wprops = this hence "0 < cls_val_process Mkt arb_pf 1 w" using iniT by simp thus "w\ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using wprops by simp qed ultimately have "p \ emeasure M {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using emeasure_mono set_event by fastforce hence "p \ prob {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" by (simp add: emeasure_eq_measure) thus "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" using psgt by simp qed qed thus False using assms unfolding viable_market_def using \stock_portfolio Mkt arb_pf\ by simp qed lemma (in CRR_market) viable_only_if_u: assumes "viable_market Mkt" shows "1+r < u" proof (rule ccontr) assume "\ 1+r < u" hence "u \ 1+r" by simp define arb_pf where "arb_pf = (\ (x::'a) (n::nat) w. 0::real)(stk:= (\ n w. -1), risk_free_asset := (\ n w. geom_proc 0 w))" have "support_set arb_pf = {stk, risk_free_asset}" proof show "support_set arb_pf \ {stk, risk_free_asset}" by (simp add: arb_pf_def subset_iff support_set_def) have "stk\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks by simp moreover have "risk_free_asset\ support_set arb_pf" unfolding arb_pf_def support_set_def using two_stocks geometric_process S0_positive by simp ultimately show "{stk, risk_free_asset}\ support_set arb_pf" by simp qed hence "stock_portfolio Mkt arb_pf" using stocks by (simp add: portfolio_def stock_portfolio_def) have "arbitrage_process Mkt arb_pf" proof (rule arbitrage_processI, intro exI conjI) show "self_financing Mkt arb_pf" unfolding arb_pf_def using \support_set arb_pf = {stk, risk_free_asset}\ by (simp add: static_portfolio_self_financing) show "trading_strategy arb_pf" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio arb_pf" unfolding portfolio_def using \support_set arb_pf = {stk, risk_free_asset}\ by simp fix asset assume "asset\ support_set arb_pf" show "borel_predict_stoch_proc G (arb_pf asset)" proof (cases "asset = stk") case True hence "arb_pf asset = (\ n w. -1)" unfolding arb_pf_def by (simp add: two_stocks) show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. -1)\ by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. -1)\ by simp qed qed next case False hence "arb_pf asset = (\ n w. geom_proc 0 w)" using \support_set arb_pf = {stk, risk_free_asset}\ \asset \ support_set arb_pf\ unfolding arb_pf_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "arb_pf asset 0 \ borel_measurable (G 0)" using \arb_pf asset = (\ n w. geom_proc 0 w)\ geometric_process by simp show "\n. arb_pf asset (Suc n) \ borel_measurable (G n)" proof fix n show "arb_pf asset (Suc n) \ borel_measurable (G n)" using \arb_pf asset = (\ n w. geom_proc 0 w)\ geometric_process by simp qed qed qed qed show "\w\space M. cls_val_process Mkt arb_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt arb_pf 0 w = geom_proc 0 w * arb_pf stk (Suc 0) w + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_vp_expand \stock_portfolio Mkt arb_pf\ using \self_financing Mkt arb_pf\ self_financingE by fastforce also have "... = geom_proc 0 w * (-1) + disc_rfr_proc r 0 w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = -geom_proc 0 w + arb_pf risk_free_asset (Suc 0) w" by simp also have "... = geom_proc 0 w - geom_proc 0 w" unfolding arb_pf_def by simp also have "... = 0" by simp finally show "cls_val_process Mkt arb_pf 0 w = 0" . qed have dev: "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" proof (intro ballI) fix w assume "w\ space M" have "cls_val_process Mkt arb_pf (Suc 0) w = geom_proc (Suc 0) w * arb_pf stk (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" using stock_pf_uvp_expand \stock_portfolio Mkt arb_pf\ by simp also have "... = -geom_proc (Suc 0) w + disc_rfr_proc r (Suc 0) w * arb_pf risk_free_asset (Suc 0) w" by (simp add: arb_pf_def two_stocks) also have "... = -geom_proc (Suc 0) w + (1+r) * arb_pf risk_free_asset (Suc 0) w" by simp also have "... = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" by (simp add:arb_pf_def) finally show "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" . qed have iniT: "\w\ space M. snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (intro ballI impI) fix w assume "w\ space M" and "snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = - u * geom_proc 0 w + (1+r) * geom_proc 0 w" using \snth w 0\ geometric_process by simp also have "... = (-u + (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... \ 0" using S0_positive \u\ 1 + r\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" . qed have iniF: "\w\ space M. \snth w 0 \ cls_val_process Mkt arb_pf (Suc 0) w > 0" proof (intro ballI impI) fix w assume "w\ space M" and "\snth w 0" have "cls_val_process Mkt arb_pf (Suc 0) w = -geom_proc (Suc 0) w + (1+r) * geom_proc 0 w" using dev \w\ space M\ by simp also have "... = -d * geom_proc 0 w + (1+r) * geom_proc 0 w" using \\snth w 0\ geometric_process by simp also have "... = (-d + (1+r)) * geom_proc 0 w" by (simp add: left_diff_distrib) also have "... > 0" using S0_positive \u <= 1 + r\ down_lt_up geometric_process by auto finally show "cls_val_process Mkt arb_pf (Suc 0) w > 0" . qed have "\w\ space M. cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof fix w assume "w\ space M" show "cls_val_process Mkt arb_pf (Suc 0) w \ 0" proof (cases "snth w 0") case True thus ?thesis using \w\ space M\ iniT by simp next case False thus ?thesis using \w\ space M\ iniF by auto qed qed thus "AE w in M. 0 \ cls_val_process Mkt arb_pf (Suc 0) w" by simp show "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" proof - have "cls_val_process Mkt arb_pf (Suc 0) \ borel_measurable M" using borel_adapt_stoch_proc_borel_measurable cls_val_process_adapted \trading_strategy arb_pf\ \stock_portfolio Mkt arb_pf\ using support_adapt_def readable unfolding stock_portfolio_def by blast hence set_event:"{w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w} \ sets M" using borel_measurable_iff_greater by blast have "\n. emeasure M {w \ space M. \w !! n} = ennreal (1-p)" using bernoulli p_gt_0 p_lt_1 bernoulli_stream_component_probability_compl[of M p] by auto hence "emeasure M {w \ space M. \w !! 0} = ennreal (1-p)" by blast moreover have "{w \ space M. \w !! 0} \ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" proof fix w assume "w\ {w \ space M. \w !! 0}" hence "w \ space M" and "\w !! 0" by auto note wprops = this hence "0 < cls_val_process Mkt arb_pf 1 w" using iniF by simp thus "w\ {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using wprops by simp qed ultimately have "1-p \ emeasure M {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" using emeasure_mono set_event by fastforce hence "1-p \ prob {w \ space M. 0 < cls_val_process Mkt arb_pf 1 w}" by (simp add: emeasure_eq_measure) thus "0 < prob {w \ space M. 0 < cls_val_process Mkt arb_pf (Suc 0) w}" using pslt by simp qed qed thus False using assms unfolding viable_market_def using \stock_portfolio Mkt arb_pf\ by simp qed lemma (in CRR_market) viable_iff: shows "viable_market Mkt \ (d < 1+r \ 1+r < u)" using viable_if viable_only_if_d viable_only_if_u by auto subsection \Risk-neutral probability space for the geometric random walk\ lemma (in CRR_market) stock_price_borel_measurable: shows "borel_adapt_stoch_proc G (prices Mkt stk)" proof - have "borel_adapt_stoch_proc (stoch_proc_filt M geom_proc borel) (prices Mkt stk)" by (simp add: geom_rand_walk_borel_measurable stk_price stoch_proc_filt_adapt) thus ?thesis by (simp add:stock_filtration) qed lemma (in CRR_market) risk_free_asset_martingale: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "martingale N G (discounted_value r (prices Mkt risk_free_asset))" proof - have "filtration N G" by (simp add: assms bernoulli_gen_filtration) moreover have "\n. sigma_finite_subalgebra N (G n)" by (simp add: assms bernoulli_sigma_finite) moreover have "finite_measure N" using assms bernoulli_stream_def prob_space.prob_space_stream_space prob_space_def prob_space_measure_pmf by auto moreover have "discounted_value r (prices Mkt risk_free_asset) = (\ n w. 1)" using discounted_rfr by auto ultimately show ?thesis using finite_measure.constant_martingale by simp qed lemma (in infinite_coin_toss_space) nat_filtration_from_eq_sets: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "sets (infinite_coin_toss_space.nat_filtration N n) = sets (nat_filtration n)" proof - have "sigma_sets (space (bernoulli_stream q)) {pseudo_proj_True n -` B \ space N |B. B \ sets (bernoulli_stream q)} = sigma_sets (space (bernoulli_stream p)) {pseudo_proj_True n -` B \ space M |B. B \ sets (bernoulli_stream p)}" proof - have "sets N = events" by (metis assms(1) bernoulli_stream_def infinite_coin_toss_space_axioms infinite_coin_toss_space_def sets_measure_pmf sets_stream_space_cong) then show ?thesis using assms(1) bernoulli_stream_space infinite_coin_toss_space_axioms infinite_coin_toss_space_def by auto qed thus ?thesis using infinite_coin_toss_space.nat_filtration_sets using assms(1) assms(2) assms(3) infinite_coin_toss_space_axioms infinite_coin_toss_space_def by auto qed lemma (in CRR_market) geom_proc_integrable: assumes "N = bernoulli_stream q" and "0 \ q" and "q \ 1" shows "integrable N (geom_proc n)" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by unfold_locales show "geom_proc n \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using geometric_process prob_grw.geom_rand_walk_borel_adapted[of q N geom_proc u d init] by (metis \infinite_coin_toss_space q N\ geom_rand_walk_pseudo_proj_True infinite_coin_toss_space.nat_filtration_borel_measurable_characterization prob_grw.geom_rand_walk_borel_measurable prob_grw_axioms prob_grw_def) qed lemma (in CRR_market) CRR_infinite_cts_filtration: shows "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) lemma (in CRR_market) proj_stoch_proc_geom_disc_fct: shows "disc_fct (proj_stoch_proc geom_proc n)" unfolding disc_fct_def using CRR_infinite_cts_filtration by (simp add: countable_finite geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_set_finite_range) lemma (in CRR_market) proj_stoch_proc_geom_rng: assumes "N = bernoulli_stream q" shows "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" proof - have "random_variable (stream_space borel) (proj_stoch_proc geom_proc n)" using CRR_infinite_cts_filtration using geom_rand_walk_borel_adapted nat_discrete_filtration proj_stoch_measurable_if_adapted by blast then show ?thesis using assms(1) bernoulli bernoulli_stream_def by auto qed lemma (in CRR_market) proj_stoch_proc_geom_open_set: shows "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume "r\ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof show "infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast show "range (proj_stoch_proc geom_proc n) \ infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r = {r}" using infinite_cts_filtration.stream_space_single_preimage \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast qed qed lemma (in CRR_market) bernoulli_AE_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof (rule finite_measure.charact_cond_exp') have "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) show "finite_measure N" using assms by (simp add: bernoulli_stream_def prob_space.finite_measure prob_space.prob_space_stream_space prob_space_measure_pmf) show "disc_fct (proj_stoch_proc geom_proc n)" using proj_stoch_proc_geom_disc_fct by simp show "integrable N X" using assms by simp show "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" using assms proj_stoch_proc_geom_rng by simp show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" using proj_stoch_proc_geom_open_set by simp qed lemma (in CRR_market) geom_proc_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_AE_cond_exp) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in CRR_market) expl_cond_eq_sets: assumes "N = bernoulli_stream q" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule expl_cond_exp_borel) show "proj_stoch_proc geom_proc n \ space N \ space (stream_space borel)" proof - have "random_variable (stream_space borel) (proj_stoch_proc geom_proc n)" using CRR_infinite_cts_filtration geom_rand_walk_borel_adapted proj_stoch_measurable_if_adapted nat_discrete_filtration by blast then show ?thesis by (simp add: assms(1) bernoulli bernoulli_stream_space measurable_def) qed show "disc_fct (proj_stoch_proc geom_proc n)" unfolding disc_fct_def using CRR_infinite_cts_filtration by (simp add: countable_finite geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_set_finite_range) show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume "r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof show "infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast show "range (proj_stoch_proc geom_proc n) \ infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r = {r}" using infinite_cts_filtration.stream_space_single_preimage \r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)\ geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by blast qed qed qed lemma (in CRR_market) bernoulli_real_cond_exp_AE: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof - have "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_AE_cond_exp by simp show "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" proof - have "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by simp moreover have "subalgebra (infinite_coin_toss_space.nat_filtration N n) (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_coin_toss_space.stoch_proc_subalg_nat_filt[of q N geom_proc n] infinite_cts_filtration.stoch_proc_filt_gen[of q N] by (metis \infinite_coin_toss_space q N\ infinite_cts_filtration_axioms.intro infinite_cts_filtration_def prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) ultimately show ?thesis using measurable_from_subalg by blast qed show "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" proof - have "expl_cond_expect N (proj_stoch_proc geom_proc n) X \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by (simp add: expl_cond_eq_sets assms) moreover have "subalgebra (infinite_coin_toss_space.nat_filtration N n) (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_coin_toss_space.stoch_proc_subalg_nat_filt[of q N geom_proc n] infinite_cts_filtration.stoch_proc_filt_gen[of q N] by (metis \infinite_coin_toss_space q N\ infinite_cts_filtration_axioms.intro infinite_cts_filtration_def prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) ultimately show ?thesis using measurable_from_subalg by blast qed show "0 < q" and "q < 1" using assms by auto qed thus ?thesis by simp qed lemma (in CRR_market) geom_proc_real_cond_exp_AE: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_real_cond_exp_AE) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in CRR_market) geom_proc_stoch_proc_filt: assumes "N= bernoulli_stream q" and "0 < q" and "q < 1" shows "stoch_proc_filt N geom_proc borel n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" proof (rule infinite_cts_filtration.stoch_proc_filt_gen) show "infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)" unfolding infinite_cts_filtration_def proof show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "infinite_cts_filtration_axioms N (infinite_coin_toss_space.nat_filtration N)" using infinite_cts_filtration_axioms_def by blast qed show "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) geom_proc" using \infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)\ prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def using infinite_cts_filtration_def by auto qed lemma (in CRR_market) bernoulli_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "integrable N X" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" proof - have aeq: "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_AE_cond_exp by simp have "\w. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) X w = expl_cond_expect N (proj_stoch_proc geom_proc n) X w" using assms bernoulli_real_cond_exp_AE by simp moreover have "stoch_proc_filt N geom_proc borel n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" using assms geom_proc_stoch_proc_filt by simp ultimately show ?thesis by simp qed lemma (in CRR_market) stock_cond_exp: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) (geom_proc (Suc n)) w = expl_cond_expect N (proj_stoch_proc geom_proc n) (geom_proc (Suc n)) w" proof (rule bernoulli_cond_exp) show "integrable N (geom_proc (Suc n))" using assms geom_proc_integrable[of N q "Suc n"] by simp qed (auto simp add: assms) lemma (in prob_space) discount_factor_real_cond_exp: assumes "integrable M X" and "subalgebra M G" and "-1 < r" shows "AE w in M. real_cond_exp M G (\x. discount_factor r n x * X x) w = discount_factor r n w * (real_cond_exp M G X) w" proof (rule sigma_finite_subalgebra.real_cond_exp_mult) show "sigma_finite_subalgebra M G" using assms subalgebra_sigma_finite by simp show "discount_factor r n \ borel_measurable G" by (simp add: discount_factor_borel_measurable) show "random_variable borel X" using assms by simp show "integrable M (\x. discount_factor r n x * X x)" using assms discounted_integrable[of M "\n. X"] unfolding discounted_value_def by simp qed lemma (in prob_space) discounted_value_real_cond_exp: assumes "integrable M X" and "-1 < r" and "subalgebra M G" shows "AE w in M. real_cond_exp M G ((discounted_value r (\ m. X)) n) w = discounted_value r (\m. (real_cond_exp M G X)) n w" using assms unfolding discounted_value_def init_triv_filt_def filtration_def by (simp add: assms discount_factor_real_cond_exp) lemma (in CRR_market) assumes "q = (1 + r - d)/(u -d)" and "viable_market Mkt" shows gt_param: "0 < q" and lt_param: "q < 1" and risk_neutral_param: "u * q + d * (1 - q) = 1 + r" proof - show "0 < q" using down_lt_up viable_only_if_d assms by simp show "q < 1" using down_lt_up viable_only_if_u assms by simp show "u * q + d * (1 - q) = 1 + r" proof - have "1 - q = 1 - (1 + r - d) / (u - d)" using assms by simp also have "... = (u - d)/(u - d) - (1 + r - d) / (u - d)" using down_lt_up by simp also have "... = (u - d - (1 + r - d))/(u-d)" using diff_divide_distrib[of "u - d" "1 + r -d" "u -d"] by simp also have "... = (u - 1 - r)/(u-d)" by simp finally have "1 - q = (u - 1 - r)/(u -d)" . hence "u * q + d * (1 - q) = u * (1 + r - d)/(u - d) + d * (u - 1 - r)/(u - d)" using assms by simp also have "... = (u * (1 + r - d) + d * (u - 1 - r))/(u - d)" using add_divide_distrib[of "u * (1 + r - d)"] by simp also have "... = (u * (1 + r) - u * d + d * u - d * (1 + r))/(u - d)" by (simp add: diff_diff_add right_diff_distrib') also have "... = (u * (1+r) - d * (1+r))/(u - d)" by simp also have "... = ((u - d) * (1+r))/(u - d)" by (simp add: left_diff_distrib) also have "... = 1 + r" using down_lt_up by simp finally show ?thesis . qed qed lemma (in CRR_market) bernoulli_expl_cond_expect_adapt: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) f\ borel_measurable (G n)" proof - have "sets N = sets M" using assms by (simp add: bernoulli bernoulli_stream_def sets_stream_space_cong) have icf: "infinite_cts_filtration p M nat_filtration" by (unfold_locales, simp) have "G n = stoch_proc_filt M geom_proc borel n" using stock_filtration by simp also have "... = fct_gen_subalgebra M (stream_space borel) (proj_stoch_proc geom_proc n)" proof (rule infinite_cts_filtration.stoch_proc_filt_gen) show "infinite_cts_filtration p M nat_filtration" using icf . show "borel_adapt_stoch_proc nat_filtration geom_proc" using geom_rand_walk_borel_adapted . qed also have "... = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" by (rule fct_gen_subalgebra_eq_sets, (simp add: \sets N = sets M\)) finally have "G n = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)" . moreover have "expl_cond_expect N (proj_stoch_proc geom_proc n) f \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" by (simp add: expl_cond_eq_sets assms) ultimately show ?thesis by simp qed lemma (in CRR_market) real_cond_exp_discount_stock: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" proof - have qlt: "0 < q" and qgt: "q < 1" using assms by auto have "G n = (fct_gen_subalgebra M (stream_space borel) (proj_stoch_proc geom_proc n))" using stock_filtration infinite_cts_filtration.stoch_proc_filt_gen[of p M nat_filtration geom_proc n] geometric_process geom_rand_walk_borel_adapted CRR_infinite_cts_filtration by simp also have "... = (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule fct_gen_subalgebra_eq_sets) show "events = sets N" using assms qlt qgt by (simp add: bernoulli bernoulli_stream_def sets_stream_space_cong) qed finally have "G n = (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" . hence "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (prices Mkt stk) (Suc n)) w" by simp moreover have "AE w in N. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (prices Mkt stk) (Suc n)) w = real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n)) w" proof - have "\w. (discounted_value r (prices Mkt stk) (Suc n)) w = (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n)) w" proof fix w show "discounted_value r (prices Mkt stk) (Suc n) w = discounted_value r (\m. prices Mkt stk (Suc n)) (Suc n) w" by (simp add: discounted_value_def) qed hence "(discounted_value r (prices Mkt stk) (Suc n)) = (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n))" by auto thus ?thesis by simp qed moreover have "AE w in N. (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) (discounted_value r (\m. (prices Mkt stk) (Suc n)) (Suc n))) w = discounted_value r (\m. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (Suc n) w" proof (rule prob_space.discounted_value_real_cond_exp) show "-1 < r" using acceptable_rate by simp show "integrable N (prices Mkt stk (Suc n))" using stk_price geom_proc_integrable assms qlt qgt by simp show "subalgebra N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n))" proof (rule fct_gen_subalgebra_is_subalgebra) show "proj_stoch_proc geom_proc n \ N \\<^sub>M stream_space borel" proof - have "proj_stoch_proc geom_proc n \ measurable M (stream_space borel)" proof (rule proj_stoch_measurable_if_adapted) show "borel_adapt_stoch_proc nat_filtration geom_proc" using geometric_process geom_rand_walk_borel_adapted by simp show "filtration M nat_filtration" using CRR_infinite_cts_filtration by (simp add: nat_discrete_filtration) qed thus ?thesis using assms bernoulli_stream_equiv filt_equiv_measurable qlt qgt psgt pslt by blast qed qed show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) qed moreover have "AE w in N. discounted_value r (\m. real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (Suc n) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" proof (rule discounted_AE_cong) have "AEeq N (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (\w. q * (prices Mkt stk) (Suc n) (pseudo_proj_True n w) + (1 - q) * (prices Mkt stk) (Suc n) (pseudo_proj_False n w))" proof (rule infinite_cts_filtration.f_borel_Suc_real_cond_exp) show icf: "infinite_cts_filtration q N (infinite_coin_toss_space.nat_filtration N)" unfolding infinite_cts_filtration_def proof show "infinite_coin_toss_space q N" using assms qlt qgt by (simp add: infinite_coin_toss_space_def) show "infinite_cts_filtration_axioms N (infinite_coin_toss_space.nat_filtration N)" using infinite_cts_filtration_axioms_def by blast qed have badapt: "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) (prices Mkt stk)" using stk_price prob_grw.geom_rand_walk_borel_adapted[of q N geom_proc] unfolding adapt_stoch_proc_def by (metis (full_types) borel_measurable_integrable geom_proc_integrable geom_rand_walk_pseudo_proj_True icf infinite_coin_toss_space.nat_filtration_borel_measurable_characterization infinite_coin_toss_space_def infinite_cts_filtration_def) show "prices Mkt stk (Suc n) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (Suc n))" using badapt unfolding adapt_stoch_proc_def by simp show "proj_stoch_proc geom_proc n \ infinite_coin_toss_space.nat_filtration N n \\<^sub>M stream_space borel" proof (rule proj_stoch_adapted_if_adapted) show "filtration N (infinite_coin_toss_space.nat_filtration N)" using icf using infinite_coin_toss_space.nat_discrete_filtration infinite_cts_filtration_def by blast show "borel_adapt_stoch_proc (infinite_coin_toss_space.nat_filtration N) geom_proc" using badapt stk_price by simp qed show "set_discriminating n (proj_stoch_proc geom_proc n) (stream_space borel)" unfolding set_discriminating_def proof (intro allI impI) fix w assume "proj_stoch_proc geom_proc n w \ proj_stoch_proc geom_proc n (pseudo_proj_True n w)" hence False using CRR_infinite_cts_filtration by (metis \proj_stoch_proc geom_proc n w \ proj_stoch_proc geom_proc n (pseudo_proj_True n w)\ geom_rand_walk_borel_adapted infinite_cts_filtration.proj_stoch_proj_invariant) thus "\A\sets (stream_space borel). (proj_stoch_proc geom_proc n w \ A) = (proj_stoch_proc geom_proc n (pseudo_proj_True n w) \ A)" by simp qed show "\w. proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (infinite_coin_toss_space.nat_filtration N n)" proof fix w show "proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (infinite_coin_toss_space.nat_filtration N n)" using \proj_stoch_proc geom_proc n \ infinite_coin_toss_space.nat_filtration N n \\<^sub>M stream_space borel\ using assms geom_rand_walk_borel_adapted nat_filtration_from_eq_sets qlt qgt infinite_cts_filtration.proj_stoch_singleton_set CRR_infinite_cts_filtration by blast qed show "\r\range (proj_stoch_proc geom_proc n) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" proof fix r assume asm: "r \ range (proj_stoch_proc geom_proc n) \ space (stream_space borel)" define A where "A = infinite_cts_filtration.stream_space_single (proj_stoch_proc geom_proc n) r" have "A \ sets (stream_space borel)" using infinite_cts_filtration.stream_space_single_set unfolding A_def using badapt icf stk_price asm by blast moreover have "range (proj_stoch_proc geom_proc n) \ A = {r}" unfolding A_def using badapt icf stk_price infinite_cts_filtration.stream_space_single_preimage asm by blast ultimately show "\A\sets (stream_space borel). range (proj_stoch_proc geom_proc n) \ A = {r}" by auto qed show "\y z. proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n \ prices Mkt stk (Suc n) y = prices Mkt stk (Suc n) z" proof (intro allI impI) fix y z assume "proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n" hence "geom_proc n y = geom_proc n z" using proj_stoch_proc_component(2)[of n n] proof - show ?thesis by (metis \\w f. n \ n \ proj_stoch_proc f n w !! n = f n w\ \proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n\ order_refl) qed hence "geom_proc (Suc n) y = geom_proc (Suc n) z" using geometric_process by (simp add: \proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n\) thus "prices Mkt stk (Suc n) y = prices Mkt stk (Suc n) z" using stk_price by simp qed show "0 < q" and "q < 1" using assms by auto qed moreover have "\w. q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = (q * u + (1 - q) * d) * prices Mkt stk n w" proof fix w have "q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = q * geom_proc (Suc n) (pseudo_proj_True n w) + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" by (simp add:stk_price) also have "... = q * u * geom_proc n (pseudo_proj_True n w) + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" using geometric_process unfolding pseudo_proj_True_def by simp also have "... = q * u * geom_proc n w + (1-q) * geom_proc (Suc n) (pseudo_proj_False n w)" by (metis geom_rand_walk_pseudo_proj_True o_apply) also have "... = q * u * geom_proc n w + (1-q) * d * geom_proc n (pseudo_proj_False n w)" using geometric_process unfolding pseudo_proj_False_def by simp also have "... = q * u * geom_proc n w + (1-q) * d * geom_proc n w" by (metis geom_rand_walk_pseudo_proj_False o_apply) also have "... = (q * u + (1 - q) * d) * geom_proc n w" by (simp add: distrib_right) finally show "q * prices Mkt stk (Suc n) (pseudo_proj_True n w) + (1 - q) * prices Mkt stk (Suc n) (pseudo_proj_False n w) = (q * u + (1 - q) * d) * prices Mkt stk n w" using stk_price by simp qed ultimately show "AEeq N (real_cond_exp N (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc n)) ((prices Mkt stk) (Suc n))) (\w. (q * u + (1 - q) * d) * prices Mkt stk n w)" by simp qed ultimately show ?thesis by auto qed lemma (in CRR_market) risky_asset_martingale_only_if: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "martingale N G (discounted_value r (prices Mkt stk))" shows "q = (1 + r - d) / (u - d)" proof - have "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = discounted_value r (prices Mkt stk) 0 w" using assms unfolding martingale_def by simp hence "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = prices Mkt stk 0 w" by (simp add: discounted_init) moreover have "AE w in N. real_cond_exp N (G 0) (discounted_value r (prices Mkt stk) (Suc 0)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk 0 w) (Suc 0) w" using assms real_cond_exp_discount_stock by simp ultimately have "AE w in N. discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk 0 w) (Suc 0) w = prices Mkt stk 0 w" by auto hence "AE w in N. discounted_value r (\m w. (q * u + (1 - q) * d) * init) (Suc 0) w = (\w. init) w" using stk_price geometric_process by simp hence "AE w in N. discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init = (\w. init) w" unfolding discounted_value_def by simp hence "AE w in N. (1+r) * discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init = (1+r) * (\w. init) w" by auto hence prev: "AE w in N. discount_factor r 0 w * (q * u + (1 - q) * d) * init = (1+r) * (\w. init) w" using discount_factor_times_rfr[of r 0] acceptable_rate proof - have "\s. (1 + r) * discount_factor r (Suc 0) (s::bool stream) = discount_factor r 0 s" by (metis (no_types) \\w. - 1 < r \ (1 + r) * discount_factor r (Suc 0) w = discount_factor r 0 w\ acceptable_rate) then show ?thesis using \AEeq N (\w. (1 + r) * discount_factor r (Suc 0) w * (q * u + (1 - q) * d) * init) (\w. (1 + r) * init)\ by presburger qed hence "\w. (\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init) w = (\w. (1+r) * init) w" proof - have "(\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" proof (rule borel_measurable_times)+ show "(\x. init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp show "(\x. q * u + (1 - q) * d) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp show "discount_factor r 0 \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" using discount_factor_nonrandom[of r 0 "infinite_coin_toss_space.nat_filtration N 0"] by simp qed moreover have "(\w. (1 + r) * init) \ borel_measurable (infinite_coin_toss_space.nat_filtration N 0)" by simp moreover have "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) ultimately show ?thesis using prev infinite_coin_toss_space.nat_filtration_AE_eq[of q N "(\w. discount_factor r 0 w * (q * u + (1 - q) * d) * init)" "(\w. (1 + r) * init)" 0] assms by (simp add: discount_factor_init) qed hence "(q * u + (1 - q) * d) * init = (1+r) * init" by (simp add: discount_factor_init) hence "q * u + (1 - q) * d = 1+r" using S0_positive by simp hence "q * u + d - q * d = 1+r" by (simp add: left_diff_distrib) hence "q * (u - d) = 1 + r - d" by (metis (no_types, opaque_lifting) add.commute add.left_commute add_diff_cancel_left' add_uminus_conv_diff left_diff_distrib mult.commute) thus "q = (1 + r - d) / (u - d)" using down_lt_up by (metis add.commute add.right_neutral diff_add_cancel nonzero_eq_divide_eq order_less_irrefl) qed locale CRR_market_viable = CRR_market + assumes CRR_viable: "viable_market Mkt" lemma (in CRR_market_viable) real_cond_exp_discount_stock_q_const: assumes "N = bernoulli_stream q" and "q = (1+r-d) / (u-d)" shows "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" proof - have qlt: "0 < q" and qgt: "q < 1" using assms gt_param lt_param CRR_viable by auto have "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (q * u + (1 - q) * d) * prices Mkt stk n w) (Suc n) w" using assms real_cond_exp_discount_stock[of N q] qlt qgt by simp moreover have "\w. (q * u + (1 - q) * d) * prices Mkt stk n w = (1+r) * prices Mkt stk n w" using risk_neutral_param assms CRR_viable by (simp add: mult.commute) ultimately have "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w" by simp moreover have "\w\ space N. discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w = discounted_value r (\m w. prices Mkt stk n w) n w" using acceptable_rate by (simp add:discounted_mult_times_rfr) moreover hence "\w\ space N. discounted_value r (\m w. (1+r) * prices Mkt stk n w) (Suc n) w = discounted_value r (prices Mkt stk) n w" using acceptable_rate by (simp add:discounted_value_def) ultimately show "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" by simp qed lemma (in CRR_market_viable) risky_asset_martingale_if: assumes "N = bernoulli_stream q" and "q = (1 + r - d) / (u - d)" shows "martingale N G (discounted_value r (prices Mkt stk))" proof (rule disc_martingale_charact) have qlt: "0 < q" and qgt: "q < 1" using assms gt_param lt_param CRR_viable by auto show "\n. integrable N (discounted_value r (prices Mkt stk) n)" proof fix n show "integrable N (discounted_value r (prices Mkt stk) n)" proof (rule discounted_integrable) show "space N = space M" using assms by (simp add: bernoulli bernoulli_stream_space) show "integrable N (prices Mkt stk n)" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms qlt qgt by (simp add: infinite_coin_toss_space_def) show "prices Mkt stk n \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using geom_rand_walk_borel_adapted stk_price nat_filtration_from_eq_sets unfolding adapt_stoch_proc_def by (metis \infinite_coin_toss_space q N\ borel_measurable_integrable geom_proc_integrable geom_rand_walk_pseudo_proj_True infinite_coin_toss_space.nat_filtration_borel_measurable_characterization infinite_coin_toss_space_def) qed show "-1 < r" using acceptable_rate by simp qed qed show "filtration N G" using qlt qgt by (simp add: bernoulli_gen_filtration assms) show "\n. sigma_finite_subalgebra N (G n)" using qlt qgt by (simp add: assms bernoulli_sigma_finite) show "\m. discounted_value r (prices Mkt stk) m \ borel_measurable (G m)" proof fix m have "discounted_value r (\ma. prices Mkt stk m) m \ borel_measurable (G m)" proof (rule discounted_measurable) show "prices Mkt stk m \ borel_measurable (G m)" using stock_price_borel_measurable unfolding adapt_stoch_proc_def by simp qed thus "discounted_value r (prices Mkt stk) m \ borel_measurable (G m)" by (metis (mono_tags, lifting) discounted_value_def measurable_cong) qed show "\n. AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" proof fix n show "AE w in N. real_cond_exp N (G n) (discounted_value r (prices Mkt stk) (Suc n)) w = discounted_value r (prices Mkt stk) n w" using assms real_cond_exp_discount_stock_q_const by simp qed qed lemma (in CRR_market_viable) risk_neutral_iff': assumes "N = bernoulli_stream q" and "0 \ q" and "q \ 1" and "filt_equiv nat_filtration M N" shows "rfr_disc_equity_market.risk_neutral_prob G Mkt r N \ q= (1 + r - d) / (u - d)" proof have "0 < q" and "q < 1" using assms filt_equiv_sgt filt_equiv_slt psgt pslt by auto note qprops = this have dem: "rfr_disc_equity_market M G Mkt r risk_free_asset" by unfold_locales { assume "rfr_disc_equity_market.risk_neutral_prob G Mkt r N" hence "(prob_space N) \ (\ asset \ stocks Mkt. martingale N G (discounted_value r (prices Mkt asset)))" using rfr_disc_equity_market.risk_neutral_prob_def[of M G Mkt] dem by simp hence "martingale N G (discounted_value r (prices Mkt stk))" using stocks by simp thus "q = (1 + r - d) / (u - d)" using assms risky_asset_martingale_only_if[of N q] qprops by simp } { assume "q = (1 + r - d) / (u - d)" hence "martingale N G (discounted_value r (prices Mkt stk))" using risky_asset_martingale_if[of N q] assms by simp moreover have "martingale N G (discounted_value r (prices Mkt risk_free_asset))" using risk_free_asset_martingale assms qprops by simp ultimately show "rfr_disc_equity_market.risk_neutral_prob G Mkt r N" using stocks using assms(1) bernoulli_stream_def dem prob_space.prob_space_stream_space prob_space_measure_pmf rfr_disc_equity_market.risk_neutral_prob_def by fastforce } qed lemma (in CRR_market_viable) risk_neutral_iff: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "rfr_disc_equity_market.risk_neutral_prob G Mkt r N \ q= (1 + r - d) / (u - d)" using bernoulli_stream_equiv assms risk_neutral_iff' psgt pslt by auto subsection \Existence of a replicating portfolio\ fun (in CRR_market) rn_rev_price where "rn_rev_price N der matur 0 w = der w" | "rn_rev_price N der matur (Suc n) w = discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w" lemma (in CRR_market) stock_filtration_eq: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "G n = stoch_proc_filt N geom_proc borel n" proof - have "G n= stoch_proc_filt M geom_proc borel n" using stock_filtration by simp also have "... = stoch_proc_filt N geom_proc borel n" proof (rule stoch_proc_filt_filt_equiv) show "filt_equiv nat_filtration M N" using assms bernoulli_stream_equiv psgt pslt by simp qed finally show ?thesis . qed lemma (in CRR_market) real_exp_eq: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "real_cond_exp N (stoch_proc_filt N geom_proc borel n) der w = expl_cond_expect N (proj_stoch_proc geom_proc n) der w" proof - have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel n) der w = expl_cond_expect N (proj_stoch_proc geom_proc n) der w" proof (rule bernoulli_cond_exp) show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto show "integrable N der" using \integrable N der\ . qed qed lemma (in CRR_market) rn_rev_price_rev_borel_adapt: assumes "cash_flow \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "(n \ matur) \ (rn_rev_price N cash_flow matur n) \ borel_measurable (G (matur - n))" proof (induct n) case 0 thus ?case using assms by simp next case (Suc n) have "rn_rev_price N cash_flow matur (Suc n) = (\w. discount_factor r (Suc 0) w * (expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N cash_flow matur n)) w)" using rn_rev_price.simps(2) by blast also have "... \ borel_measurable (G (matur - Suc n))" proof (rule borel_measurable_times) show "discount_factor r (Suc 0) \ borel_measurable (G (matur - Suc n))" by (simp add:discount_factor_borel_measurable) show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N cash_flow matur n) \ borel_measurable (G (matur - Suc n))" using assms by (simp add: bernoulli_expl_cond_expect_adapt) qed finally show "rn_rev_price N cash_flow matur (Suc n) \ borel_measurable (G (matur - Suc n))" . qed lemma (in infinite_coin_toss_space) bernoulli_discounted_integrable: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (nat_filtration n)" and "-1 < r" shows "integrable N (discounted_value r (\m. der) m)" proof - have "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N n)" using assms filt_equiv_filtration by (simp add: assms(1) measurable_def nat_filtration_from_eq_sets nat_filtration_space) qed thus ?thesis using discounted_integrable assms by (metis \prob_space N\ prob_space.discounted_integrable) qed lemma (in CRR_market) rn_rev_expl_cond_expect: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "n \ matur \ rn_rev_price N der matur n w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) w" proof (induct n arbitrary: w) case 0 have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast have "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed have "rn_rev_price N der matur 0 w = der w" by simp also have "... = expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) w" proof (rule nat_filtration_AE_eq) show "der \ borel_measurable (nat_filtration matur)" using \der \ borel_measurable (nat_filtration matur)\ . have "(discounted_value r (\m. der) 0) = der" unfolding discounted_value_def discount_factor_def by simp moreover have "AEeq N (real_cond_exp N (G matur) der) der" proof (rule sigma_finite_subalgebra.real_cond_exp_F_meas) show "der \ borel_measurable (G matur)" using assms by simp show "integrable N der" using \integrable N der\ . show "sigma_finite_subalgebra N (G matur)" using bernoulli_sigma_finite using assms by simp qed moreover have "\w. real_cond_exp N (stoch_proc_filt N geom_proc borel matur) der w = expl_cond_expect N (proj_stoch_proc geom_proc matur) der w" using assms real_exp_eq by simp ultimately have eqn: "AEeq N der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" using stock_filtration_eq assms by auto have "stoch_proc_filt M geom_proc borel matur = stoch_proc_filt N geom_proc borel matur" using bernoulli_stream_equiv[of N q] assms psgt pslt by (simp add: stoch_proc_filt_filt_equiv) also have "stoch_proc_filt N geom_proc borel matur = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur)" using assms geom_proc_stoch_proc_filt by simp finally have "stoch_proc_filt M geom_proc borel matur = fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur)" . moreover have "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (fct_gen_subalgebra N (stream_space borel) (proj_stoch_proc geom_proc matur))" proof (rule expl_cond_exp_borel) show "proj_stoch_proc geom_proc matur \ space N \ space (stream_space borel)" using assms proj_stoch_proc_geom_rng by (simp add: measurable_def) show "disc_fct (proj_stoch_proc geom_proc matur)" using proj_stoch_proc_geom_disc_fct by simp show "\r\range (proj_stoch_proc geom_proc matur) \ space (stream_space borel). \A\sets (stream_space borel). range (proj_stoch_proc geom_proc matur) \ A = {r}" using proj_stoch_proc_geom_open_set by simp qed ultimately show ebm: "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (nat_filtration matur)" by (metis geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt) show "AEeq M der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" proof (rule filt_equiv_borel_AE_eq_iff[THEN iffD2]) show "filt_equiv nat_filtration M N" using assms bernoulli_stream_equiv psgt pslt by simp show "der \ borel_measurable (nat_filtration matur)" using \der \ borel_measurable (nat_filtration matur)\ . show "AEeq N der (expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0))" using eqn . show "expl_cond_expect N (proj_stoch_proc geom_proc matur) (discounted_value r (\m. der) 0) \ borel_measurable (nat_filtration matur)" using ebm . show "prob_space N" using assms by (simp add: bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) show "prob_space M" by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) qed show "0 < p" "p < 1" using psgt pslt by auto qed also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - 0)) (discounted_value r (\m. der) 0) w" by simp finally show "rn_rev_price N der matur 0 w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - 0)) (discounted_value r (\m. der) 0) w" . next case (Suc n) have "rn_rev_price N der matur (Suc n) w = discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w" by simp also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w" proof - have "expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (rn_rev_price N der matur n) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w" proof (rule real_exp_eq[symmetric]) show "rn_rev_price N der matur n \ borel_measurable (G (matur - n))" using assms rn_rev_price_rev_borel_adapt Suc by simp show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto qed thus ?thesis by simp qed also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "rn_rev_price N der matur n \ borel_measurable N" proof - have "rn_rev_price N der matur n \ borel_measurable (G (matur - n))" by (metis (full_types) Suc.prems Suc_leD assms(1) assms(2) assms(3) assms(4) rn_rev_price_rev_borel_adapt) then show ?thesis by (metis (no_types) assms(2) bernoulli bernoulli_stream_def filtration_measurable measurable_cong_sets sets_measure_pmf sets_stream_space_cong) qed show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" using Suc.hyps Suc.prems Suc_leD \rn_rev_price N der matur n \ borel_measurable N\ by presburger show "AEeq N (rn_rev_price N der matur n) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n))" using Suc by auto qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (rn_rev_price N der matur n) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) qed thus ?thesis by simp qed also have "... = discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n))) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" by simp show "expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n) \ borel_measurable N" by (metis assms(2) assms(3) assms(4) bernoulli bernoulli_expl_cond_expect_adapt bernoulli_stream_def filtration_measurable measurable_cong_sets sets_measure_pmf sets_stream_space_cong) show "AEeq N (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))" proof - have "discounted_value r (\m. der) n \ borel_measurable (G matur)" using assms discounted_measurable[of der] by simp hence "\w. (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) w = (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w" using real_exp_eq[of _ matur N q "matur-n"] assms by simp thus ?thesis by simp qed qed show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (expl_cond_expect N (proj_stoch_proc geom_proc (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) qed thus ?thesis by simp qed also have "... = real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" proof (rule infinite_coin_toss_space.nat_filtration_AE_eq) show "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) show "(\a. discount_factor r (Suc 0) a * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) a) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" proof - have "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) \ borel_measurable (infinite_coin_toss_space.nat_filtration N (matur - Suc n))" by (metis assms(2) assms(3) assms(4) borel_measurable_cond_exp infinite_coin_toss_space.intro infinite_coin_toss_space.stoch_proc_subalg_nat_filt linear measurable_from_subalg not_less prob_grw.geom_rand_walk_borel_adapted prob_grw_axioms prob_grw_def) thus ?thesis using discounted_measurable[of "real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))"] unfolding discounted_value_def by simp qed show "0 < q" "q < 1" using assms by auto show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n)))" proof- have "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n)) w) (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n) w)" proof - have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - n)) (discounted_value r (\m. der) n))) (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n))" proof (rule sigma_finite_subalgebra.real_cond_exp_nested_subalg) show "sigma_finite_subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite stock_filtration_eq by auto show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce show "subalgebra (stoch_proc_filt N geom_proc borel (matur - n)) (stoch_proc_filt N geom_proc borel (matur - Suc n))" proof - have "init_triv_filt M (stoch_proc_filt M geom_proc borel)" using infinite_cts_filtration.stoch_proc_filt_triv_init using info_filtration stock_filtration by auto moreover have "matur - (Suc n) \ matur - n" by simp ultimately show ?thesis unfolding init_triv_filt_def filtration_def using assms(2) assms(3) assms(4) stock_filtration stock_filtration_eq by auto qed show "integrable N (discounted_value r (\m. der) n) " using bernoulli_discounted_integrable[of N q der matur r n] acceptable_rate assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast qed thus ?thesis by auto qed moreover have "AEeq N (\w. discount_factor r (Suc 0) w * real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n) w) (\w. discount_factor r (Suc 0) w * (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n) w)" proof - have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) n)) (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n)" proof (rule prob_space.discounted_value_real_cond_exp) show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "-1 < r" using acceptable_rate . show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce qed thus ?thesis by auto qed moreover have "\w. (\w. discount_factor r (Suc 0) w * (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) n) w) w = (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) (Suc n)) w" unfolding discounted_value_def discount_factor_def by simp moreover have "AEeq N (real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) (discounted_value r (\m. der) (Suc n))) (discounted_value r (\m. real_cond_exp N (stoch_proc_filt N geom_proc borel (matur - Suc n)) der) (Suc n))" proof (rule prob_space.discounted_value_real_cond_exp) show "prob_space N" using assms by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "der \ borel_measurable (nat_filtration matur)" using assms using geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "integrable N der" proof (rule infinite_coin_toss_space.nat_filtration_borel_measurable_integrable) show "infinite_coin_toss_space q N" using assms by (simp add: infinite_coin_toss_space_def) show "der \ borel_measurable (infinite_coin_toss_space.nat_filtration N matur)" by (metis \der \ borel_measurable (nat_filtration matur)\ \infinite_coin_toss_space q N\ assms(2) assms(3) assms(4) infinite_coin_toss_space.nat_filtration_space measurable_from_subalg nat_filtration_from_eq_sets nat_filtration_space subalgebra_def subset_eq) qed show "-1 < r" using acceptable_rate . show "subalgebra N (stoch_proc_filt N geom_proc borel (matur - Suc n))" using assms(2) assms(3) assms(4) bernoulli_sigma_finite sigma_finite_subalgebra.subalg stock_filtration_eq by fastforce qed ultimately show ?thesis by auto qed qed also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" proof (rule real_exp_eq) show "discounted_value r (\m. der) (Suc n) \ borel_measurable (G matur)" using assms discounted_measurable[of der] by simp show "N = bernoulli_stream q" "0 < q" "q < 1" using assms by auto qed finally show "rn_rev_price N der matur (Suc n) w = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc n)) (discounted_value r (\m. der) (Suc n)) w" . qed definition (in CRR_market) rn_price where "rn_price N der matur n w = expl_cond_expect N (proj_stoch_proc geom_proc n) (discounted_value r (\m. der) (matur - n)) w" definition (in CRR_market) rn_price_ind where "rn_price_ind N der matur n w = rn_rev_price N der matur (matur - n) w" lemma (in CRR_market) rn_price_eq: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" and "n \ matur" shows "rn_price N der matur n w = rn_price_ind N der matur n w" using rn_rev_expl_cond_expect unfolding rn_price_def rn_price_ind_def by (simp add: assms) lemma (in CRR_market) geom_proc_filt_info: fixes f::"bool stream \ 'b::{t0_space}" assumes "f \ borel_measurable (G n)" shows "f w = f (pseudo_proj_True n w)" proof - have "subalgebra (nat_filtration n) (G n)" using stoch_proc_subalg_nat_filt[of geom_proc n] geometric_process stock_filtration geom_rand_walk_borel_adapted by simp hence "f\ borel_measurable (nat_filtration n)" using assms by (simp add: measurable_from_subalg) thus ?thesis using nat_filtration_info[of f n] by (metis comp_apply) qed lemma (in CRR_market) geom_proc_filt_info': fixes f::"bool stream \ 'b::{t0_space}" assumes "f \ borel_measurable (G n)" shows "f w = f (pseudo_proj_False n w)" proof - have "subalgebra (nat_filtration n) (G n)" using stoch_proc_subalg_nat_filt[of geom_proc n] geometric_process stock_filtration geom_rand_walk_borel_adapted by simp hence "f\ borel_measurable (nat_filtration n)" using assms by (simp add: measurable_from_subalg) thus ?thesis using nat_filtration_info'[of f n] by (metis comp_apply) qed lemma (in CRR_market) rn_price_borel_adapt: assumes "cash_flow \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" and "n \ matur" shows "(rn_price N cash_flow matur n) \ borel_measurable (G n)" proof - show "(rn_price N cash_flow matur n) \ borel_measurable (G n)" using assms rn_rev_price_rev_borel_adapt[of cash_flow matur N q "matur - n"] rn_price_eq rn_price_ind_def by (smt add.right_neutral cancel_comm_monoid_add_class.diff_cancel diff_commute diff_le_self increasing_measurable_info measurable_cong nat_le_linear ordered_cancel_comm_monoid_diff_class.add_diff_inverse) qed definition (in CRR_market) delta_price where "delta_price N cash_flow T = (\ n w. if (Suc n \ T) then (rn_price N cash_flow T (Suc n) (pseudo_proj_True n w) - rn_price N cash_flow T (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) else 0)" lemma (in CRR_market) delta_price_eq: assumes "Suc n \ T" shows "delta_price N cash_flow T n w = (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d))" proof - have "(geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) = geom_proc n w * (u - d)" by (simp add: geom_rand_walk_diff_induct) then show ?thesis unfolding delta_price_def using assms spick_eq_pseudo_proj_True spick_eq_pseudo_proj_False by simp qed lemma (in CRR_market) geom_proc_spick: shows "geom_proc (Suc n) (spick w n x) = (if x then u else d) * geom_proc n w" proof - have "geom_proc (Suc n) (spick w n x) = geom_rand_walk u d init (Suc n) (spick w n x)" using geometric_process by simp also have "... = (case (spick w n x) !! n of True \ u | False \ d) * geom_rand_walk u d init n (spick w n x)" by simp also have "... = (case x of True \ u | False \ d) * geom_rand_walk u d init n (spick w n x)" unfolding spick_def by simp also have "... = (if x then u else d) * geom_rand_walk u d init n (spick w n x)" by simp also have "... = (if x then u else d) * geom_rand_walk u d init n w" by (metis comp_def geom_rand_walk_pseudo_proj_True geometric_process pseudo_proj_True_stake_image spickI) finally show ?thesis using geometric_process by simp qed lemma (in CRR_market) spick_red_geom: shows "(\w. spick w n x) \ measurable (fct_gen_subalgebra M borel (geom_proc n)) (fct_gen_subalgebra M borel (geom_proc (Suc n)))" unfolding measurable_def proof (intro CollectI conjI) show "(\w. spick w n x) \ space (fct_gen_subalgebra M borel (geom_proc n)) \ space (fct_gen_subalgebra M borel (geom_proc (Suc n)))" by (simp add: bernoulli bernoulli_stream_space fct_gen_subalgebra_space) show "\y\sets (fct_gen_subalgebra M borel (geom_proc (Suc n))). (\w. spick w n x) -` y \ space (fct_gen_subalgebra M borel (geom_proc n)) \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof fix A assume A: "A \ sets (fct_gen_subalgebra M borel (geom_proc (Suc n)))" show "(\w. spick w n x) -` A \ space (fct_gen_subalgebra M borel (geom_proc n)) \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof - define sp where "sp = (\w. spick w n x)" have "A \ {(geom_proc (Suc n)) -` B \ space M |B. B \ sets borel}" using A by (simp add:fct_gen_subalgebra_sigma_sets) from this obtain C where "C\ sets borel" and "A = (geom_proc (Suc n)) -`C \ space M" by auto hence "A = (geom_proc (Suc n)) -`C" using bernoulli bernoulli_stream_space by simp hence "sp -`A = sp -` (geom_proc (Suc n)) -`C" by simp also have "... = (geom_proc (Suc n) \ sp) -` C" by auto also have "... = (\w. (if x then u else d) * geom_proc n w) -` C" using geom_proc_spick sp_def by auto also have "... \ sets (fct_gen_subalgebra M borel (geom_proc n))" proof (cases x) case True hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. u * geom_proc n w) -` C" by simp moreover have "(\w. u * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ by (metis bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space measurable_sets) next case False hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. d * geom_proc n w) -` C" by simp moreover have "(\w. d * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ by (metis bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space measurable_sets) qed finally show ?thesis unfolding sp_def by (simp add: bernoulli bernoulli_stream_space fct_gen_subalgebra_space) qed qed qed lemma (in CRR_market) geom_spick_Suc: assumes "A \ {(geom_proc (Suc n)) -` B |B. B \ sets borel}" shows "(\w. spick w n x) -`A \ {geom_proc n -`B | B. B\ sets borel}" proof - have "sets (fct_gen_subalgebra M borel (geom_proc n)) = {geom_proc n -` B \space M |B. B \ sets borel}" by (simp add: fct_gen_subalgebra_sigma_sets) also have "... = {geom_proc n -` B |B. B \ sets borel}" using bernoulli bernoulli_stream_space by simp finally have sf: "sets (fct_gen_subalgebra M borel (geom_proc n)) = {geom_proc n -` B |B. B \ sets borel}" . define sp where "sp = (\w. spick w n x)" from assms(1) obtain C where "C\ sets borel" and "A = (geom_proc (Suc n)) -`C" by auto hence "A = (geom_proc (Suc n)) -`C" using bernoulli bernoulli_stream_space by simp hence "sp -`A = sp -` (geom_proc (Suc n)) -`C" by simp also have "... = (geom_proc (Suc n) \ sp) -` C" by auto also have "... = (\w. (if x then u else d) * geom_proc n w) -` C" using geom_proc_spick sp_def by auto also have "... \ {geom_proc n -`B | B. B\ sets borel}" proof (cases x) case True hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. u * geom_proc n w) -` C" by simp moreover have "(\w. u * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ sf by (simp add: bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space in_borel_measurable_borel) next case False hence "(\w. (if x then u else d) * geom_proc n w) -` C = (\w. d * geom_proc n w) -` C" by simp moreover have "(\w. d * geom_proc n w) \ borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" proof - have "geom_proc n \borel_measurable (fct_gen_subalgebra M borel (geom_proc n))" using fct_gen_subalgebra_fct_measurable by (metis (no_types, lifting) geom_rand_walk_borel_measurable measurable_def mem_Collect_eq) thus ?thesis by simp qed ultimately show ?thesis using \C\ sets borel\ sf by (simp add: bernoulli bernoulli_stream_preimage fct_gen_subalgebra_space in_borel_measurable_borel) qed finally show ?thesis unfolding sp_def . qed lemma (in CRR_market) geom_spick_lt: assumes "m< n" shows "geom_proc m (spick w n x) = geom_proc m w" proof - have "geom_proc m (spick w n x) = geom_proc m (pseudo_proj_True m (spick w n x))" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_True m w)" using assms by (metis less_imp_le_nat pseudo_proj_True_def pseudo_proj_True_prefix spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) finally show ?thesis . qed lemma (in CRR_market) geom_spick_eq: shows "geom_proc m (spick w m x) = geom_proc m w" proof (cases x) case True have "geom_proc m (spick w m x) = geom_proc m (pseudo_proj_True m (spick w m x))" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_True m w)" using True by (metis pseudo_proj_True_def spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_True by (metis comp_apply) finally show ?thesis . next case False have "geom_proc m (spick w m x) = geom_proc m (pseudo_proj_False m (spick w m x))" using geom_rand_walk_pseudo_proj_False by (metis comp_apply) also have "... = geom_proc m (pseudo_proj_False m w)" using False by (metis pseudo_proj_False_def spickI) also have "... = geom_proc m w" using geom_rand_walk_pseudo_proj_False by (metis comp_apply) finally show ?thesis . qed lemma (in CRR_market) spick_red_geom_filt: shows "(\w. spick w n x) \ measurable (G n) (G (Suc n))" unfolding measurable_def proof (intro CollectI conjI) show "(\w. spick w n x) \ space (G n) \ space (G (Suc n))" using stock_filtration by (simp add: bernoulli bernoulli_stream_space stoch_proc_filt_space) show "\y\sets (G (Suc n)). (\w. spick w n x) -` y \ space (G n) \ sets (G n)" proof fix B assume "B\ sets (G (Suc n))" hence "B\ (sigma_sets (space M) (\ i\ {m. m\ (Suc n)}. {(geom_proc i -`A) \ (space M) | A. A\ sets borel }))" using stock_filtration stoch_proc_filt_sets geometric_process proof - have "\n. sigma_sets (space M) (\n\{na. na \ n}. {geom_proc n -` R \ space M |R. R \ sets borel}) = sets (G n)" by (simp add: geom_rand_walk_borel_measurable stoch_proc_filt_sets stock_filtration) then show ?thesis using \B \ sets (G (Suc n))\ by blast qed hence "(\w. spick w n x) -` B \ sets (G n)" proof (induct rule:sigma_sets.induct) { fix C assume "C \ (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel})" hence "\m \ Suc n. C\ {geom_proc m -` A \ space M |A. A \ sets borel}" by auto from this obtain m where "m\ Suc n" and "C\ {geom_proc m -` A \ space M |A. A \ sets borel}" by auto note Cprops = this from this obtain D where "C = geom_proc m -` D\ space M" and "D\ sets borel" by auto hence "C = geom_proc m -`D" using bernoulli bernoulli_stream_space by simp have "C\ {geom_proc m -` A |A. A \ sets borel}" using bernoulli bernoulli_stream_space Cprops by simp show "(\w. spick w n x) -` C \ sets (G n)" proof (cases "m \ n") case True have "(\w. spick w n x) -` C = (\w. spick w n x) -` geom_proc m -`D" using \C = geom_proc m -`D\ by simp also have "... = (geom_proc m \ (\w. spick w n x)) -`D" by auto also have "... = geom_proc m -`D" using geom_spick_lt geom_spick_eq \m\n\ using le_eq_less_or_eq by auto also have "... \ sets (G n)" using stock_filtration geometric_process \D\ sets borel\ by (metis (no_types, lifting) True adapt_stoch_proc_def bernoulli bernoulli_stream_preimage geom_rand_walk_borel_measurable increasing_measurable_info measurable_sets stoch_proc_filt_adapt stoch_proc_filt_space) finally show "(\w. spick w n x) -` C \ sets (G n)" . next case False hence "m = Suc n" using \m \ Suc n\ by simp hence "(\w. spick w n x) -` C \ {geom_proc n -` B |B. B \ sets borel}" using \C\ {geom_proc m -` A |A. A \ sets borel}\ geom_spick_Suc by simp also have "... \ sets (G n)" proof - have "{geom_proc n -` B |B. B \ sets borel} \ {geom_proc n -` B \ space M |B. B \ sets borel}" using bernoulli bernoulli_stream_space by simp also have "... \ (\i\{m. m \ n}. {geom_proc i -` A \ space M |A. A \ sets borel})" by auto also have "... \ sigma_sets (space M) (\i\{m. m \ n}. {geom_proc i -` A \ space M |A. A \ sets borel})" by (rule sigma_sets_superset_generator) also have "... = sets (G n)" using stock_filtration geometric_process stoch_proc_filt_sets[of n geom_proc M borel] geom_rand_walk_borel_measurable by blast finally show ?thesis . qed finally show ?thesis . qed } show "(\w. spick w n x) -` {} \ sets (G n)" by simp { fix C assume "C \ sigma_sets (space M) (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel})" and "(\w. spick w n x) -` C \ sets (G n)" hence "(\w. spick w n x) -` (space M - C) = (\w. spick w n x) -` (space M) - (\w. spick w n x) -` C" by (simp add: vimage_Diff) also have "... = space M - (\w. spick w n x) -` C" using bernoulli bernoulli_stream_space by simp also have "... \ sets (G n)" using \(\w. spick w n x) -` C \ sets (G n)\ by (metis algebra.compl_sets disc_filtr_def discrete_filtration sets.sigma_algebra_axioms sigma_algebra_def subalgebra_def) finally show "(\w. spick w n x) -` (space M - C) \ sets (G n)" . } { fix C::"nat \ bool stream set" assume "(\i. C i \ sigma_sets (space M) (\i\{m. m \ Suc n}. {geom_proc i -` A \ space M |A. A \ sets borel}))" and "(\i. (\w. spick w n x) -` C i \ sets (G n))" hence "(\w. spick w n x) -` \(C ` UNIV) = (\ i\ UNIV. (\w. spick w n x) -` (C i))" by blast also have "... \ sets (G n)" using \\i. (\w. spick w n x) -` C i \ sets (G n)\ by simp finally show "(\w. spick w n x) -` \(C ` UNIV) \ sets (G n)" . } qed thus "(\w. spick w n x) -` B \ space (G n) \ sets (G n)" using stock_filtration stoch_proc_filt_space bernoulli bernoulli_stream_space by simp qed qed lemma (in CRR_market) delta_price_adapted: fixes cash_flow::"bool stream \ real" assumes "cash_flow \ borel_measurable (G T)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "borel_adapt_stoch_proc G (delta_price N cash_flow T)" unfolding adapt_stoch_proc_def proof fix n show "delta_price N cash_flow T n \ borel_measurable (G n)" proof (cases "Suc n \ T") case True hence deleq: "\w. delta_price N cash_flow T n w = (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d))" using delta_price_eq by simp have "(\w. rn_price N cash_flow T (Suc n) (spick w n True)) \ borel_measurable (G n)" proof - have "rn_price N cash_flow T (Suc n) \ borel_measurable (G (Suc n))" using rn_price_borel_adapt assms using True by blast moreover have "(\w. spick w n True) \ G n \\<^sub>M G (Suc n)" using spick_red_geom_filt by simp ultimately show ?thesis by simp qed moreover have "(\w. rn_price N cash_flow T (Suc n) (spick w n False)) \ borel_measurable (G n)" proof - have "rn_price N cash_flow T (Suc n) \ borel_measurable (G (Suc n))" using rn_price_borel_adapt assms using True by blast moreover have "(\w. spick w n False) \ G n \\<^sub>M G (Suc n)" using spick_red_geom_filt by simp ultimately show ?thesis by simp qed ultimately have "(\w. rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False)) \ borel_measurable (G n)" by simp moreover have "(\w. (geom_proc n w) * (u - d)) \ borel_measurable (G n)" proof - have "geom_proc n \ borel_measurable (G n)" using stock_filtration by (metis adapt_stoch_proc_def stk_price stock_price_borel_measurable) thus ?thesis by simp qed ultimately have "(\w. (rn_price N cash_flow T (Suc n) (spick w n True) - rn_price N cash_flow T (Suc n) (spick w n False))/ ((geom_proc n w) * (u - d)))\ borel_measurable (G n)" by simp thus ?thesis using deleq by presburger next case False thus ?thesis unfolding delta_price_def by simp qed qed fun (in CRR_market) delta_predict where "delta_predict N der matur 0 = (\w. delta_price N der matur 0 w)" | "delta_predict N der matur (Suc n) = (\w. delta_price N der matur n w)" lemma (in CRR_market) delta_predict_predict: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "borel_predict_stoch_proc G (delta_predict N der matur)" unfolding predict_stoch_proc_def proof (intro conjI) show "delta_predict N der matur 0 \ borel_measurable (G 0)" using delta_price_adapted[of der matur N q] assms unfolding adapt_stoch_proc_def by force show "\n. delta_predict N der matur (Suc n) \ borel_measurable (G n)" proof fix n show "delta_predict N der matur (Suc n) \ borel_measurable (G n)" using delta_price_adapted[of der matur N q] assms unfolding adapt_stoch_proc_def by force qed qed definition (in CRR_market) delta_pf where "delta_pf N der matur = qty_single stk (delta_predict N der matur)" lemma (in CRR_market) delta_pf_support: shows "support_set (delta_pf N der matur) \ {stk}" unfolding delta_pf_def using single_comp_support[of stk "delta_predict N der matur"] by simp definition (in CRR_market) self_fin_delta_pf where "self_fin_delta_pf N der matur v0 = self_finance Mkt v0 (delta_pf N der matur) risk_free_asset" lemma (in disc_equity_market) self_finance_trading_strat: assumes "trading_strategy pf" and "portfolio pf" and "borel_adapt_stoch_proc F (prices Mkt asset)" and "support_adapt Mkt pf" shows "trading_strategy (self_finance Mkt v pf asset)" unfolding self_finance_def proof (rule sum_trading_strat) show "trading_strategy pf" using assms by simp show "trading_strategy (qty_single asset (remaining_qty Mkt v pf asset))" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio (qty_single asset (remaining_qty Mkt v pf asset))" by (simp add: self_finance_def single_comp_portfolio) show "\a. a \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) \ borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" proof (cases "support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {}") case False hence eqasset: "support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {asset}" using single_comp_support by fastforce fix a assume "a\ support_set (qty_single asset (remaining_qty Mkt v pf asset))" hence "a = asset" using eqasset by simp hence "qty_single asset (remaining_qty Mkt v pf asset) a = (remaining_qty Mkt v pf asset)" unfolding qty_single_def by simp moreover have "borel_predict_stoch_proc F (remaining_qty Mkt v pf asset)" proof (rule remaining_qty_predict) show "trading_strategy pf" using assms by simp show "borel_adapt_stoch_proc F (prices Mkt asset)" using assms by simp show "support_adapt Mkt pf" using assms by simp qed ultimately show "borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" by simp next case True thus "\a. a \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) \ support_set (qty_single asset (remaining_qty Mkt v pf asset)) = {} \ borel_predict_stoch_proc F (qty_single asset (remaining_qty Mkt v pf asset) a)" by simp qed qed qed lemma (in CRR_market) self_fin_delta_pf_trad_strat: assumes "der\ borel_measurable (G matur)" and "N = bernoulli_stream q" and "0 < q" and "q < 1" shows "trading_strategy (self_fin_delta_pf N der matur v0)" unfolding self_fin_delta_pf_def proof (rule self_finance_trading_strat) show "trading_strategy (delta_pf N der matur)" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio (delta_pf N der matur)" unfolding portfolio_def using delta_pf_support by (meson finite.emptyI finite_insert infinite_super) show "\asset. asset \ support_set (delta_pf N der matur) \ borel_predict_stoch_proc G (delta_pf N der matur asset)" proof (cases "support_set (delta_pf N der matur) = {}") case False fix asset assume "asset \ support_set (delta_pf N der matur)" hence "asset = stk" using False delta_pf_support by auto hence "delta_pf N der matur asset = delta_predict N der matur" unfolding delta_pf_def qty_single_def by simp thus "borel_predict_stoch_proc G (delta_pf N der matur asset)" using delta_predict_predict assms by simp next case True thus "\asset. asset \ support_set (delta_pf N der matur) \ support_set (delta_pf N der matur) = {} \ borel_predict_stoch_proc G (delta_pf N der matur asset)" by simp qed qed show "portfolio (delta_pf N der matur)" using delta_pf_support unfolding portfolio_def by (meson finite.emptyI finite_insert infinite_super) show "borel_adapt_stoch_proc G (prices Mkt risk_free_asset)" using rf_price disc_rfr_proc_borel_adapted by simp show "support_adapt Mkt (delta_pf N der matur)" unfolding support_adapt_def proof show "\asset. asset \ support_set (delta_pf N der matur) \ borel_adapt_stoch_proc G (prices Mkt asset)" proof (cases "support_set (delta_pf N der matur) = {}") case False fix asset assume "asset \ support_set (delta_pf N der matur)" hence "asset = stk" using False delta_pf_support by auto hence "prices Mkt asset = geom_proc" using stk_price by simp thus "borel_adapt_stoch_proc G (prices Mkt asset)" using \asset = stk\ stock_price_borel_measurable by auto next case True thus "\asset. asset \ support_set (delta_pf N der matur) \ borel_adapt_stoch_proc G (prices Mkt asset)" by simp qed qed qed definition (in CRR_market) delta_hedging where "delta_hedging N der matur = self_fin_delta_pf N der matur (prob_space.expectation N (discounted_value r (\m. der) matur))" lemma (in CRR_market) geom_proc_eq_snth: shows "(\m. m \ Suc n \ geom_proc m x = geom_proc m y) \ (\m. m \ n \ snth x m = snth y m)" proof (induct n ) case 0 assume asm: "(\m. m \Suc 0 \ geom_proc m x = geom_proc m y)" and "m\ 0" hence "m = 0" by simp have "geom_proc (Suc 0) x = geom_proc (Suc 0) y" using asm by simp have "snth x 0 = snth y 0" proof (rule ccontr) assume "snth x 0 \ snth y 0" show False proof (cases "snth x 0") case True hence "\ snth y 0" using \snth x 0 \ snth y 0\ by simp have "geom_proc (Suc 0) x = u * init" using geometric_process True by simp moreover have "geom_proc (Suc 0) y = d * init" using geometric_process \\ snth y 0\ by simp ultimately have "geom_proc (Suc 0) x \ geom_proc (Suc 0) y" using S0_positive down_lt_up by simp thus ?thesis using \geom_proc (Suc 0) x = geom_proc (Suc 0) y\ by simp next case False hence "snth y 0" using \snth x 0 \ snth y 0\ by simp have "geom_proc (Suc 0) x = d * init" using geometric_process False by simp moreover have "geom_proc (Suc 0) y = u * init" using geometric_process \snth y 0\ by simp ultimately have "geom_proc (Suc 0) x \ geom_proc (Suc 0) y" using S0_positive down_lt_up by simp thus ?thesis using \geom_proc (Suc 0) x = geom_proc (Suc 0) y\ by simp qed qed thus "\m. (\m. m \ Suc 0 \ geom_proc m x = geom_proc m y) \ m \ 0 \ x !! m = y !! m" by simp next case (Suc n) assume fst: "(\m. (\m. m \ Suc n \ geom_proc m x = geom_proc m y) \ m \ n \ x !! m = y !! m)" and scd: "(\m. m \ Suc (Suc n) \ geom_proc m x = geom_proc m y)" and "m \ Suc n" show "x !! m = y !! m" proof (cases "m \ n") case True thus ?thesis using fst scd by simp next case False hence "m = Suc n" using \m\ Suc n\ by simp have "geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y" using scd by simp show ?thesis proof (rule ccontr) assume "x !! m \ y !! m" thus False proof (cases "x !! m") case True hence "\ y !! m" using \x !! m \ y !! m\ by simp have "geom_proc (Suc (Suc n)) x = u * geom_proc (Suc n) x" using geometric_process True \m = Suc n\ by simp also have "... = u * geom_proc (Suc n) y" using scd \m = Suc n\ by simp finally have "geom_proc (Suc (Suc n)) x = u * geom_proc (Suc n) y" . moreover have "geom_proc (Suc (Suc n)) y = d * geom_proc (Suc n) y" using geometric_process \m = Suc n\ \\ y !! m\ by simp ultimately have "geom_proc (Suc (Suc n)) x \ geom_proc (Suc (Suc n)) y" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl mult_cancel_right) thus ?thesis using \geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y\ by simp next case False hence "y !! m" using \x !! m \ y !! m\ by simp have "geom_proc (Suc (Suc n)) x = d * geom_proc (Suc n) x" using geometric_process False \m = Suc n\ by simp also have "... = d * geom_proc (Suc n) y" using scd \m = Suc n\ by simp finally have "geom_proc (Suc (Suc n)) x = d * geom_proc (Suc n) y" . moreover have "geom_proc (Suc (Suc n)) y = u * geom_proc (Suc n) y" using geometric_process \m = Suc n\ \y !! m\ by simp ultimately have "geom_proc (Suc (Suc n)) x \ geom_proc (Suc (Suc n)) y" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl mult_cancel_right) thus ?thesis using \geom_proc (Suc (Suc n)) x = geom_proc (Suc (Suc n)) y\ by simp qed qed qed qed lemma (in CRR_market) geom_proc_eq_pseudo_proj_True: shows "(\m. m \ n \ geom_proc m x = geom_proc m y) \ (pseudo_proj_True (n) x = pseudo_proj_True (n) y)" -proof - - assume a1: "\m. m \ n \ geom_proc m x = geom_proc m y" - obtain nn :: "bool stream \ bool stream \ nat \ nat" where - "\x1 x2 x3. (\v4 geom_proc v4 x1) = (nn x1 x2 x3 < Suc (Suc x3) \ geom_proc (nn x1 x2 x3) x2 \ geom_proc (nn x1 x2 x3) x1)" - by moura - then have f2: "\n s sa na. (nn sa s n < Suc (Suc n) \ geom_proc (nn sa s n) s \ geom_proc (nn sa s n) sa \ \ na < Suc n) \ s !! na = sa !! na" - by (meson geom_proc_eq_snth less_Suc_eq_le) - obtain nna :: "bool stream \ bool stream \ nat \ nat" where - f3: "\x0 x1 x2. (\v3. Suc v3 < Suc x2 \ x1 !! v3 \ x0 !! v3) = (Suc (nna x0 x1 x2) < Suc x2 \ x1 !! nna x0 x1 x2 \ x0 !! nna x0 x1 x2)" - by moura - obtain nnb :: "nat \ nat" where - f4: "\x0. (\v2. x0 = Suc v2) = (x0 = Suc (nnb x0))" - by moura - moreover - { assume "\ nn y x (nnb n) < Suc (Suc (nnb n)) \ geom_proc (nn y x (nnb n)) x = geom_proc (nn y x (nnb n)) y" - moreover - { assume "\ nna y x n < Suc (nnb n)" - then have "\ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" - using f4 by (metis (no_types) Suc_le_D Suc_le_lessD less_Suc_eq_le) } - ultimately have "pseudo_proj_True n x = pseudo_proj_True n y \ \ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" -using f2 by meson } - ultimately have "pseudo_proj_True n x = pseudo_proj_True n y \ \ Suc (nna y x n) < Suc n \ x !! nna y x n = y !! nna y x n" - using a1 Suc_le_D less_Suc_eq_le by presburger - then show ?thesis - using f3 by (meson less_Suc_eq_le pseudo_proj_True_snth') -qed + by (meson geom_proc_eq_snth le_trans order.refl pseudo_proj_True_snth') lemma (in CRR_market) proj_stoch_eq_pseudo_proj_True: assumes "proj_stoch_proc geom_proc m x = proj_stoch_proc geom_proc m y" shows "pseudo_proj_True m x = pseudo_proj_True m y" proof - have "\ k \ m. geom_proc k x = geom_proc k y" proof (intro allI impI) fix k assume "k \ m" thus "geom_proc k x = geom_proc k y" using proj_stoch_proc_eq_snth[of geom_proc m x y k] assms by simp qed thus ?thesis using geom_proc_eq_pseudo_proj_True[of m x y] by auto qed lemma (in CRR_market_viable) rn_rev_price_cond_expect: assumes "N = bernoulli_stream q" and "0 borel_measurable (G matur)" and "Suc n \ matur" shows "expl_cond_expect N (proj_stoch_proc geom_proc n) (rn_rev_price N der matur (matur - Suc n)) w= (q * rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True n w) + (1 - q) * rn_rev_price N der matur (matur - Suc n) (pseudo_proj_False n w))" proof (rule infinite_cts_filtration.f_borel_Suc_expl_cond_expect) show "infinite_cts_filtration q N nat_filtration" using assms pslt psgt bernoulli_nat_filtration by simp show "rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))" using rn_rev_price_rev_borel_adapt[of der matur N q "Suc n"] assms stock_filtration stoch_proc_subalg_nat_filt[of geom_proc] geom_rand_walk_borel_adapted by (metis add_diff_cancel_right' diff_le_self measurable_from_subalg ordered_cancel_comm_monoid_diff_class.add_diff_inverse rn_rev_price_rev_borel_adapt) show "proj_stoch_proc geom_proc n \ nat_filtration n \\<^sub>M stream_space borel" using proj_stoch_adapted_if_adapted[of M nat_filtration geom_proc borel n] pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted nat_discrete_filtration by blast show "set_discriminating n (proj_stoch_proc geom_proc n) (stream_space borel)" using infinite_cts_filtration.proj_stoch_set_discriminating pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted by simp show "proj_stoch_proc geom_proc n -` {proj_stoch_proc geom_proc n w} \ sets (nat_filtration n)" using infinite_cts_filtration.proj_stoch_singleton_set pslt psgt bernoulli_nat_filtration[of M p] bernoulli geom_rand_walk_borel_adapted by simp show "\y z. proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n \ rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) z" proof (intro allI impI) fix y z assume as:"proj_stoch_proc geom_proc n y = proj_stoch_proc geom_proc n z \ y !! n = z !! n" hence "pseudo_proj_True n y = pseudo_proj_True n z" using proj_stoch_eq_pseudo_proj_True[of n y z] by simp moreover have "snth y n = snth z n" using as by simp ultimately have "pseudo_proj_True (Suc n) y = pseudo_proj_True (Suc n) z" - proof - - have f1: "\n s sa. (\na. Suc na \ n \ s !! na \ sa !! na) \ pseudo_proj_True n s = pseudo_proj_True n sa" - by (meson pseudo_proj_True_snth') - obtain nn :: "bool stream \ bool stream \ nat \ nat" where - "\x0 x1 x2. (\v3. Suc v3 \ x2 \ x1 !! v3 \ x0 !! v3) = (Suc (nn x0 x1 x2) \ x2 \ x1 !! nn x0 x1 x2 \ x0 !! nn x0 x1 x2)" - by moura - then have f2: "\n s sa. Suc (nn sa s n) \ n \ s !! nn sa s n \ sa !! nn sa s n \ pseudo_proj_True n s = pseudo_proj_True n sa" - using f1 by presburger - have f3: "stake n y = stake n (pseudo_proj_True n z)" - by (metis \pseudo_proj_True n y = pseudo_proj_True n z\ pseudo_proj_True_stake) - { assume "stake (Suc n) z \ stake (Suc n) (pseudo_proj_True (Suc n) y)" - then have "stake n y @ [y !! n] \ stake n z @ [z !! n]" - by (metis (no_types) pseudo_proj_True_stake stake_Suc) - then have "stake (Suc n) z = stake (Suc n) (pseudo_proj_True (Suc n) y)" - using f3 by (simp add: \y !! n = z !! n\ pseudo_proj_True_stake) } - then have "\ Suc (nn z y (Suc n)) \ Suc n \ y !! nn z y (Suc n) = z !! nn z y (Suc n)" - by (metis (no_types) pseudo_proj_True_stake stake_snth) - then show ?thesis - using f2 by blast - qed + by (metis (full_types) pseudo_proj_True_def pseudo_proj_True_same_img stake_Suc) have "rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True (Suc n) y)" using nat_filtration_info[of "rn_rev_price N der matur (matur - Suc n)" "Suc n"] rn_rev_price_rev_borel_adapt[of der matur N q] by (metis \rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))\ o_apply) also have "... = rn_rev_price N der matur (matur - Suc n) (pseudo_proj_True (Suc n) z)" using \pseudo_proj_True (Suc n) y = pseudo_proj_True (Suc n) z\ by simp also have "... = rn_rev_price N der matur (matur - Suc n) z" using nat_filtration_info[of "rn_rev_price N der matur (matur - Suc n)" "Suc n"] rn_rev_price_rev_borel_adapt[of der matur N q] by (metis \rn_rev_price N der matur (matur - Suc n) \ borel_measurable (nat_filtration (Suc n))\ o_apply) finally show "rn_rev_price N der matur (matur - Suc n) y = rn_rev_price N der matur (matur - Suc n) z" . qed show "0 < q" and "q < 1" using assms by auto qed lemma (in CRR_market_viable) rn_price_eq_ind: assumes "N = bernoulli_stream q" and "n < matur" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" shows "(1+r) * rn_price N der matur n w = q * rn_price N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price N der matur (Suc n) (pseudo_proj_False n w)" proof - define V where "V = rn_price N der matur" let ?m = "matur - Suc n" have "matur -n = Suc ?m" by (simp add: assms Suc_diff_Suc Suc_le_lessD) have "(1+r) * V n w = (1+r) * rn_price_ind N der matur n w" using rn_price_eq assms unfolding V_def by simp also have "... = (1+r) * rn_rev_price N der matur (Suc ?m) w" using \matur -n = Suc ?m\ unfolding rn_price_ind_def by simp also have "... = (1+r) * discount_factor r (Suc 0) w * expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc ?m)) (rn_rev_price N der matur ?m) w" by simp also have "... = expl_cond_expect N (proj_stoch_proc geom_proc (matur - Suc ?m)) (rn_rev_price N der matur ?m) w" unfolding discount_factor_def using acceptable_rate by auto also have "... = expl_cond_expect N (proj_stoch_proc geom_proc n) (rn_rev_price N der matur ?m) w" using \matur -n = Suc ?m\ by simp also have "... = (q * rn_rev_price N der matur ?m (pseudo_proj_True n w) + (1 - q) * rn_rev_price N der matur ?m (pseudo_proj_False n w))" using rn_rev_price_cond_expect[of N q der matur n w] assms by simp also have "... = q * rn_price_ind N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price_ind N der matur (Suc n) (pseudo_proj_False n w)" unfolding rn_price_ind_def by simp also have "... = q * rn_price N der matur (Suc n) (pseudo_proj_True n w) + (1 - q) * rn_price N der matur (Suc n) (pseudo_proj_False n w)" using rn_price_eq assms by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" unfolding V_def by simp finally have "(1+r) * V n w = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" . thus ?thesis unfolding V_def by simp qed lemma self_finance_updated_suc_suc: assumes "portfolio pf" and "\n. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc n)) w = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w - val_process Mkt pf (Suc n) w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc n)) w = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * remaining_qty Mkt v pf asset (Suc (Suc n)) w" using assms by (simp add: self_finance_updated) also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((remaining_qty Mkt v pf asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) / (prices Mkt asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" using assms by (metis nonzero_mult_div_cancel_left) also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + prices Mkt asset (Suc (Suc n)) w * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) + cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w)" using add_divide_distrib[symmetric, of "prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w" "prices Mkt asset (Suc n) w"] by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * ((prices Mkt asset (Suc n) w) * (remaining_qty Mkt v pf asset (Suc n) w) + cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)" by simp also have "... = cls_val_process Mkt pf (Suc (Suc n)) w + (prices Mkt asset (Suc (Suc n)) w / (prices Mkt asset (Suc n) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w - val_process Mkt pf (Suc n) w)" using self_finance_updated[of Mkt asset n w pf v] assms by auto finally show ?thesis . qed lemma self_finance_updated_suc_0: assumes "portfolio pf" and "\n w. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc 0) w = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * (val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc 0) w = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * remaining_qty Mkt v pf asset (Suc 0) w" using assms by (simp add: self_finance_updated) also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((remaining_qty Mkt v pf asset 0 w) + (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) / (prices Mkt asset 0 w) + (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" using assms by (metis nonzero_mult_div_cancel_left) also have "... = cls_val_process Mkt pf (Suc 0) w + prices Mkt asset (Suc 0) w * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w)" using add_divide_distrib[symmetric, of "prices Mkt asset 0 w * remaining_qty Mkt v pf asset 0 w" "prices Mkt asset 0 w"] by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + v - val_process Mkt pf 0 w)" by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * ((prices Mkt asset 0 w) * (remaining_qty Mkt v pf asset 0 w) + val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" using self_finance_init[of Mkt asset pf v w] assms by simp also have "... = cls_val_process Mkt pf (Suc 0) w + (prices Mkt asset (Suc 0) w / (prices Mkt asset 0 w)) * (val_process Mkt (self_finance Mkt v pf asset) 0 w - val_process Mkt pf 0 w)" by simp finally show ?thesis . qed lemma self_finance_updated_ind: assumes "portfolio pf" and "\n w. prices Mkt asset n w \ 0" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (prices Mkt asset (Suc n) w / (prices Mkt asset n w)) * (val_process Mkt (self_finance Mkt v pf asset) n w - val_process Mkt pf n w)" proof (cases "n = 0") case True thus ?thesis using assms self_finance_updated_suc_0 by simp next case False hence "\m. n = Suc m" by (simp add: not0_implies_Suc) from this obtain m where "n = Suc m" by auto hence "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc (Suc m)) w" by simp also have "... = cls_val_process Mkt pf (Suc (Suc m)) w + (prices Mkt asset (Suc (Suc m)) w / (prices Mkt asset (Suc m) w)) * (cls_val_process Mkt (self_finance Mkt v pf asset) (Suc m) w - val_process Mkt pf (Suc m) w)" using assms self_finance_updated_suc_suc[of pf] by simp also have "... = cls_val_process Mkt pf (Suc (Suc m)) w + (prices Mkt asset (Suc (Suc m)) w / (prices Mkt asset (Suc m) w)) * (val_process Mkt (self_finance Mkt v pf asset) (Suc m) w - val_process Mkt pf (Suc m) w)" using assms self_finance_charact unfolding self_financing_def by (simp add: self_finance_succ self_finance_updated) also have "... = cls_val_process Mkt pf (Suc n) w + (prices Mkt asset (Suc n) w / (prices Mkt asset n w)) * (val_process Mkt (self_finance Mkt v pf asset) n w - val_process Mkt pf n w)" using \n = Suc m\ by simp finally show ?thesis . qed lemma (in rfr_disc_equity_market) self_finance_risk_free_update_ind: assumes "portfolio pf" shows "cls_val_process Mkt (self_finance Mkt v pf risk_free_asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (1 + r) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" proof - have "cls_val_process Mkt (self_finance Mkt v pf risk_free_asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + (prices Mkt risk_free_asset (Suc n) w / (prices Mkt risk_free_asset n w)) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" proof (rule self_finance_updated_ind, (simp add: assms), intro allI) fix n w show "prices Mkt risk_free_asset n w \ 0" using positive by (metis less_irrefl) qed also have "... = cls_val_process Mkt pf (Suc n) w + (1+r) * (val_process Mkt (self_finance Mkt v pf risk_free_asset) n w - val_process Mkt pf n w)" using rf_price positive by (metis acceptable_rate disc_rfr_proc_Suc_div) finally show ?thesis . qed lemma (in CRR_market) delta_pf_portfolio: shows "portfolio (delta_pf N der matur)" unfolding delta_pf_def by (simp add: single_comp_portfolio) lemma (in CRR_market) delta_pf_updated: shows "cls_val_process Mkt (delta_pf N der matur) (Suc n) w = geom_proc (Suc n) w * delta_price N der matur n w" unfolding delta_pf_def using stk_price qty_single_updated[of Mkt] by simp lemma (in CRR_market) delta_pf_val_process: shows "val_process Mkt (delta_pf N der matur) n w = geom_proc n w * delta_price N der matur n w" unfolding delta_pf_def using stk_price qty_single_val_process[of Mkt] by simp lemma (in CRR_market) delta_hedging_cls_val_process: shows "cls_val_process Mkt (delta_hedging N der matur) (Suc n) w = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt (delta_hedging N der matur) n w - geom_proc n w * delta_price N der matur n w)" proof - define X where "X = delta_hedging N der matur" define init where "init = integral\<^sup>L N (discounted_value r (\m. der) matur)" have "cls_val_process Mkt X (Suc n) w = cls_val_process Mkt (delta_pf N der matur) (Suc n) w + (1 + r) * (val_process Mkt X n w - val_process Mkt (delta_pf N der matur) n w)" unfolding X_def delta_hedging_def self_fin_delta_pf_def init_def proof (rule self_finance_risk_free_update_ind) show "portfolio (delta_pf N der matur)" unfolding portfolio_def using delta_pf_support by (meson finite.simps infinite_super) qed also have "... = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt X n w - val_process Mkt (delta_pf N der matur) n w)" using delta_pf_updated by simp also have "... = geom_proc (Suc n) w * delta_price N der matur n w + (1 + r) * (val_process Mkt X n w - geom_proc n w * delta_price N der matur n w)" using delta_pf_val_process by simp finally show ?thesis unfolding X_def . qed lemma (in CRR_market_viable) delta_hedging_eq_derivative_price: fixes der::"bool stream \ real" and matur::nat assumes "N = bernoulli_stream ((1 + r - d) / (u - d))" and "der\ borel_measurable (G matur)" shows "\n w. n\ matur \ val_process Mkt (delta_hedging N der matur) n w = (rn_price N der matur) n w" unfolding delta_hedging_def proof - define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this define init where "init = (prob_space.expectation N (discounted_value r (\m. der) matur))" define X where "X = val_process Mkt (delta_hedging N der matur)" define V where "V = rn_price N der matur" define \ where "\ = delta_price N der matur" { fix n fix w have "n \ matur \ X n w = V n w" proof (induct n) case 0 have v0: "V 0 \ borel_measurable (G 0)" using assms rn_price_borel_adapt "0.prems" qprops unfolding V_def q_def by auto have "X 0 w= init" using self_finance_init[of Mkt risk_free_asset "delta_pf N der matur" "integral\<^sup>L N (discounted_value r (\m. der) matur)"] delta_pf_support unfolding X_def init_def delta_hedging_def self_fin_delta_pf_def init_def by (metis finite_insert infinite_imp_nonempty infinite_super less_irrefl portfolio_def positive) also have "... = V 0 w" proof - have "\x\space N. real_cond_exp N (G 0) (discounted_value r (\m. der) matur) x = integral\<^sup>L N (discounted_value r (\m. der) matur)" proof (rule prob_space.trivial_subalg_cond_expect_eq) show "prob_space N" using assms qprops unfolding q_def by (simp add: bernoulli bernoulli_stream_def prob_space.prob_space_stream_space prob_space_measure_pmf) have "init_triv_filt M (stoch_proc_filt M geom_proc borel)" proof (rule infinite_cts_filtration.stoch_proc_filt_triv_init) show "borel_adapt_stoch_proc nat_filtration geom_proc" using geom_rand_walk_borel_adapted by simp show "infinite_cts_filtration p M nat_filtration" using bernoulli_nat_filtration[of M p] bernoulli psgt pslt by simp qed hence "init_triv_filt N (stoch_proc_filt M geom_proc borel)" using assms qprops filt_equiv_triv_init[of nat_filtration N] stock_filtration bernoulli_stream_equiv[of N] psgt pslt unfolding q_def by simp thus "subalgebra N (G 0)" and "sets (G 0) = {{}, space N}" using stock_filtration unfolding init_triv_filt_def filtration_def bot_nat_def by auto show "integrable N (discounted_value r (\m. der) matur)" proof (rule bernoulli_discounted_integrable) show "der \ borel_measurable (nat_filtration matur)" using assms geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto qed (simp add: acceptable_rate) qed hence "integral\<^sup>L N (discounted_value r (\m. der) matur) = real_cond_exp N (G 0) (discounted_value r (\m. der) matur) w" using bernoulli_stream_space[of N q] by (simp add: assms(1) q_def) also have "... = real_cond_exp N (stoch_proc_filt M geom_proc borel 0) (discounted_value r (\m. der) matur) w" using stock_filtration by simp also have "... = real_cond_exp N (stoch_proc_filt N geom_proc borel 0) (discounted_value r (\m. der) matur) w" using stoch_proc_filt_filt_equiv[of nat_filtration M N geom_proc] bernoulli_stream_equiv[of N] q_def qprops assms pslt psgt by auto also have "... = expl_cond_expect N (proj_stoch_proc geom_proc 0) (discounted_value r (\m. der) matur) w" proof (rule bernoulli_cond_exp) show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto show "integrable N (discounted_value r (\m. der) matur)" proof (rule bernoulli_discounted_integrable) show "der \ borel_measurable (nat_filtration matur)" using assms geom_rand_walk_borel_adapted measurable_from_subalg stoch_proc_subalg_nat_filt stock_filtration by blast show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" "q < 1" using qprops by auto qed (simp add: acceptable_rate) qed finally show "init = V 0 w" unfolding init_def V_def rn_price_def by simp qed finally show "X 0 w = V 0 w" . next case (Suc n) hence "n < matur" by simp show ?case proof - have "X n w = V n w" using Suc by (simp add: Suc.hyps Suc.prems Suc_leD) have "0< 1+r" using acceptable_rate by simp let ?m = "matur - Suc n" have "matur -n = Suc ?m" by (simp add: Suc.prems Suc_diff_Suc Suc_le_lessD) have "(1+r) * V n w = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) *V (Suc n) (pseudo_proj_False n w)" using rn_price_eq_ind qprops assms Suc q_def V_def by simp show "X (Suc n) w = V (Suc n) w" proof (cases "snth w n") case True hence pseq: "pseudo_proj_True (Suc n) w = pseudo_proj_True (Suc n) (spick w n True)" by (metis (mono_tags, lifting) pseudo_proj_True_stake_image spickI stake_Suc) have "X (Suc n) w = cls_val_process Mkt (delta_hedging N der matur) (Suc n) w" unfolding X_def delta_hedging_def self_fin_delta_pf_def using delta_pf_portfolio unfolding self_financing_def by (metis less_irrefl positive self_finance_charact self_financingE) also have "... = geom_proc (Suc n) w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using delta_hedging_cls_val_process unfolding X_def \_def by simp also have "... = u * geom_proc n w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using True geometric_process by simp also have "... = u * geom_proc n w * \ n w + (1 + r) * X n w - (1+r) * geom_proc n w * \ n w" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * \ n w * u - geom_proc n w * \ n w * (1 + r)" by (simp add: mult.commute mult.left_commute) also have "... = (1+r)* X n w + geom_proc n w * \ n w * (u - (1 + r))" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * (V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) * (u - (1 + r))" using Suc V_def by (simp add: \_def delta_price_def geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + geom_proc n w * ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))) / (geom_proc n w * (u - d)) * (u - (1 + r))" proof - have "geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False) = geom_proc n w * (u - d)" by (simp add: geom_rand_walk_diff_induct) then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w)))* (u - (1 + r))/ (u-d)" proof - have "geom_proc n w \ 0" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl) then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))* (1 - q))" proof - have "1 - q = 1 - (1 + r - d)/(u -d)" unfolding q_def by simp also have "... = (u - d)/(u - d) - (1 + r - d)/(u -d)" using down_lt_up by simp also have "... = (u - d - (1 + r - d))/(u - d)" using diff_divide_distrib[of "u - d" "1 + r -d"] by simp also have "... = (u - (1+r))/(u-d)" by simp finally have "1 - q = (u - (1+r))/(u-d)" . thus ?thesis by simp qed also have "... = (1+r) * X n w + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" by (simp add: mult.commute right_diff_distrib) also have "... = (1+r) * V n w + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" using \X n w = V n w\ by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_False n w) + (1 - q) * V (Suc n) (pseudo_proj_True n w) - (1 - q) * V (Suc n) (pseudo_proj_False n w)" using assms Suc rn_price_eq_ind[of N q n matur der w] \n < matur\ qprops unfolding V_def q_def by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_True n w)" by simp also have "... = V (Suc n) (pseudo_proj_True n w)" using distrib_right[of q "1 - q" "V (Suc n) (pseudo_proj_True n w)"] by simp also have "... = V (Suc n) w" proof - have "V (Suc n) \ borel_measurable (G (Suc n))" unfolding V_def q_def proof (rule rn_price_borel_adapt) show "der \ borel_measurable (G matur)" using assms by simp show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" and "q < 1" using qprops by auto show "Suc n \ matur" using Suc by simp qed hence "V (Suc n) (pseudo_proj_True n w) = V (Suc n) (pseudo_proj_True (Suc n) (pseudo_proj_True n w))" using geom_proc_filt_info[of "V (Suc n)" "Suc n"] by simp also have "... = V (Suc n) (pseudo_proj_True (Suc n) w)" using True by (simp add: pseq spick_eq_pseudo_proj_True) also have "... = V (Suc n) w" using \V (Suc n) \ borel_measurable (G (Suc n))\ geom_proc_filt_info[of "V (Suc n)" "Suc n"] by simp finally show ?thesis . qed finally show "X (Suc n) w = V (Suc n) w" . next case False hence pseq: "pseudo_proj_True (Suc n) w = pseudo_proj_True (Suc n) (spick w n False)" using filtration by (metis (full_types) pseudo_proj_True_def spickI stake_Suc) have "X (Suc n) w = cls_val_process Mkt (delta_hedging N der matur) (Suc n) w" unfolding X_def delta_hedging_def self_fin_delta_pf_def using delta_pf_portfolio unfolding self_financing_def by (metis less_irrefl positive self_finance_charact self_financingE) also have "... = geom_proc (Suc n) w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using delta_hedging_cls_val_process unfolding X_def \_def by simp also have "... = d * geom_proc n w * \ n w + (1 + r) * (X n w - geom_proc n w * \ n w)" using False geometric_process by simp also have "... = d * geom_proc n w * \ n w + (1 + r) * X n w - (1+r) * geom_proc n w * \ n w" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * \ n w * d - geom_proc n w * \ n w * (1 + r)" by (simp add: mult.commute mult.left_commute) also have "... = (1+r)* X n w + geom_proc n w * \ n w * (d - (1 + r))" by (simp add: right_diff_distrib) also have "... = (1+r) * X n w + geom_proc n w * (V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))/ (geom_proc (Suc n) (spick w n True) - geom_proc (Suc n) (spick w n False)) * (d - (1 + r))" using Suc V_def by (simp add: \_def delta_price_def geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + geom_proc n w * ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))) / (geom_proc n w * (u - d)) * (d - (1 + r))" by (simp add: geom_rand_walk_diff_induct) also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w)))* (d - (1 + r))/ (u-d)" proof - have "geom_proc n w \ 0" by (metis S0_positive down_lt_up down_positive geom_rand_walk_strictly_positive less_irrefl) then show ?thesis by simp qed also have "... = (1+r) * X n w + ((V (Suc n) (pseudo_proj_True n w) - V (Suc n) (pseudo_proj_False n w))* (-q))" proof - have "0-q = 0-(1 + r - d)/(u -d)" unfolding q_def by simp also have "... = (d - (1 + r))/(u -d)" by (simp add: minus_divide_left) finally have "0 - q = (d - (1+r))/(u-d)" . thus ?thesis by simp qed also have "... = (1+r) * X n w + (- V (Suc n) (pseudo_proj_True n w) * q + V (Suc n) (pseudo_proj_False n w)* q)" by (metis (no_types, opaque_lifting) add.inverse_inverse distrib_right minus_mult_commute minus_real_def mult_minus_left) also have "... = (1+r) * X n w - q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" by simp also have "... = (1+r) * V n w -q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" using \X n w = V n w\ by simp also have "... = q * V (Suc n) (pseudo_proj_True n w) + (1 - q) * V (Suc n) (pseudo_proj_False n w) - q * V (Suc n) (pseudo_proj_True n w) + q * V (Suc n) (pseudo_proj_False n w)" using assms Suc rn_price_eq_ind[of N q n matur der w] \n < matur\ qprops unfolding V_def q_def by simp also have "... = (1-q) * V (Suc n) (pseudo_proj_False n w) + q * V (Suc n) (pseudo_proj_False n w)" by simp also have "... = V (Suc n) (pseudo_proj_False n w)" using distrib_right[of q "1 - q" "V (Suc n) (pseudo_proj_False n w)"] by simp also have "... = V (Suc n) w" proof - have "V (Suc n) \ borel_measurable (G (Suc n))" unfolding V_def q_def proof (rule rn_price_borel_adapt) show "der \ borel_measurable (G matur)" using assms by simp show "N = bernoulli_stream q" using assms unfolding q_def by simp show "0 < q" and "q < 1" using qprops by auto show "Suc n \ matur" using Suc by simp qed hence "V (Suc n) (pseudo_proj_False n w) = V (Suc n) (pseudo_proj_False (Suc n) (pseudo_proj_False n w))" using geom_proc_filt_info'[of "V (Suc n)" "Suc n"] by simp also have "... = V (Suc n) (pseudo_proj_False (Suc n) w)" using False spick_eq_pseudo_proj_False by (metis pseq pseudo_proj_True_imp_False) also have "... = V (Suc n) w" using \V (Suc n) \ borel_measurable (G (Suc n))\ geom_proc_filt_info'[of "V (Suc n)" "Suc n"] by simp finally show ?thesis . qed finally show "X (Suc n) w = V (Suc n) w" . qed qed qed } thus "\n w. n \ matur \ val_process Mkt (self_fin_delta_pf N der matur (integral\<^sup>L N (discounted_value r (\m. der) matur))) n w = rn_price N der matur n w" by (simp add: X_def init_def V_def delta_hedging_def) qed lemma (in CRR_market_viable) delta_hedging_same_cash_flow: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream ((1 + r - d) / (u - d))" shows "cls_val_process Mkt (delta_hedging N der matur) matur w = der w" proof - define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this have "cls_val_process Mkt (delta_hedging N der matur) matur w = val_process Mkt (delta_hedging N der matur) matur w" using self_financingE self_finance_charact unfolding delta_hedging_def self_fin_delta_pf_def by (metis delta_pf_portfolio mult_1s(1) mult_cancel_right not_real_square_gt_zero positive) also have "... = rn_price N der matur matur w" using delta_hedging_eq_derivative_price assms by simp also have "... = rn_rev_price N der matur 0 w" using rn_price_eq qprops assms unfolding rn_price_ind_def q_def by simp also have "... = der w" by simp finally show ?thesis . qed lemma (in CRR_market) delta_hedging_trading_strat: assumes "N = bernoulli_stream q" and "0 < q" and "q < 1" and "der \ borel_measurable (G matur)" shows "trading_strategy (delta_hedging N der matur)" unfolding delta_hedging_def by (simp add: assms self_fin_delta_pf_trad_strat) lemma (in CRR_market) delta_hedging_self_financing: shows "self_financing Mkt (delta_hedging N der matur)" unfolding delta_hedging_def self_fin_delta_pf_def proof (rule self_finance_charact) show "\n w. prices Mkt risk_free_asset (Suc n) w \ 0" using positive by (metis less_numeral_extra(3)) show "portfolio (delta_pf N der matur)" using delta_pf_portfolio . qed lemma (in CRR_market_viable) delta_hedging_replicating: assumes "der \ borel_measurable (G matur)" and "N = bernoulli_stream ((1 + r - d) / (u - d))" shows "replicating_portfolio (delta_hedging N der matur) der matur" unfolding replicating_portfolio_def proof (intro conjI) define q where "q = (1 + r - d) / (u - d)" have "0 < q" and "q < 1" unfolding q_def using assms gt_param lt_param CRR_viable by auto note qprops = this let ?X = "(delta_hedging N der matur)" show "trading_strategy ?X" using delta_hedging_trading_strat qprops assms unfolding q_def by simp show "self_financing Mkt ?X" using delta_hedging_self_financing . show "stock_portfolio Mkt (delta_hedging N der matur)" unfolding delta_hedging_def self_fin_delta_pf_def stock_portfolio_def portfolio_def using stocks delta_pf_support by (smt Un_insert_right delta_pf_portfolio insert_commute portfolio_def self_finance_def self_finance_portfolio single_comp_support subset_insertI2 subset_singleton_iff sum_support_set sup_bot.right_neutral) show "AEeq M (cls_val_process Mkt (delta_hedging N der matur) matur) der" using delta_hedging_same_cash_flow assms by simp qed definition (in disc_equity_market) complete_market where "complete_market \ (\matur. \ der\ borel_measurable (F matur). (\p. replicating_portfolio p der matur))" lemma (in CRR_market_viable) CRR_market_complete: shows "complete_market" unfolding complete_market_def proof (intro allI impI) fix matur::nat show "\ der \ borel_measurable (G matur). (\p. replicating_portfolio p der matur)" proof fix der::"bool stream\real" assume "der \ borel_measurable (G matur)" define N where "N = bernoulli_stream ((1 + r - d) / (u - d))" hence "replicating_portfolio (delta_hedging N der matur) der matur" using delta_hedging_replicating \der \ borel_measurable (G matur)\ by simp thus "\pf. replicating_portfolio pf der matur" by auto qed qed lemma subalgebras_filtration: assumes "filtration M F" and "\t. subalgebra (F t) (G t)" and "\ s t. s \ t \ subalgebra (G t) (G s)" shows "filtration M G" unfolding filtration_def proof (intro conjI allI impI) { fix t have "subalgebra (F t) (G t)" using assms by simp moreover have "subalgebra M (F t)" using assms unfolding filtration_def by simp ultimately show "subalgebra M (G t)" by (metis subalgebra_def subsetCE subsetI) } { fix s t::'b assume "s \ t" thus "subalgebra (G t) (G s)" using assms by simp } qed lemma subfilt_filt_equiv: assumes "filt_equiv F M N" and "\ t. subalgebra (F t) (G t)" and "\ s t. s \ t \ subalgebra (G t) (G s)" shows "filt_equiv G M N" unfolding filt_equiv_def proof (intro conjI) show "sets M = sets N" using assms unfolding filt_equiv_def by simp show "filtration M G" using assms subalgebras_filtration[of M F G] unfolding filt_equiv_def by simp show "\t A. A \ sets (G t) \ (emeasure M A = 0) = (emeasure N A = 0)" proof (intro allI ballI impI) fix t fix A assume "A\ sets (G t)" hence "A \ sets (F t)" using assms unfolding subalgebra_def by auto thus "(emeasure M A = 0) = (emeasure N A = 0)" using assms unfolding filt_equiv_def by simp qed qed lemma (in CRR_market_viable) CRR_market_fair_price: assumes "pyf\ borel_measurable (G matur)" shows "fair_price Mkt (\ w\ range (pseudo_proj_True matur). (prod (prob_component ((1 + r - d) / (u - d)) w) {0..m. pyf) matur) w)) pyf matur" proof - define dpf where "dpf = (discounted_value r (\m. pyf) matur)" define q where "q = (1 + r - d) / (u - d)" have "\pf. replicating_portfolio pf pyf matur" using CRR_market_complete assms unfolding complete_market_def by simp from this obtain pf where "replicating_portfolio pf pyf matur" by auto note pfprop = this define N where "N = bernoulli_stream ((1 + r - d) / (u - d))" have "fair_price Mkt (integral\<^sup>L N dpf) pyf matur" unfolding dpf_def proof (rule replicating_expectation_finite) show "risk_neutral_prob N" using assms risk_neutral_iff using CRR_viable gt_param lt_param N_def by blast have "filt_equiv nat_filtration M N" using bernoulli_stream_equiv[of N "(1+r-d)/(u-d)"] assms gt_param lt_param CRR_viable psgt pslt N_def by simp thus "filt_equiv G M N" using subfilt_filt_equiv using Filtration.filtration_def filtration geom_rand_walk_borel_adapted stoch_proc_subalg_nat_filt stock_filtration by blast show "pyf \ borel_measurable (G matur)" using assms by simp show "viable_market Mkt" using CRR_viable by simp have "infinite_cts_filtration p M nat_filtration" using bernoulli_nat_filtration[of M p] bernoulli psgt pslt by simp thus "sets (G 0) = {{}, space M}" using stock_filtration infinite_cts_filtration.stoch_proc_filt_triv_init[of p M nat_filtration geom_proc] geom_rand_walk_borel_adapted bot_nat_def unfolding init_triv_filt_def by simp show "replicating_portfolio pf pyf matur" using pfprop . show "\n. \asset\support_set pf. finite (prices Mkt asset n ` space M)" proof (intro allI ballI) fix n fix asset assume "asset \ support_set pf" hence "prices Mkt asset n \ borel_measurable (G n)" using readable pfprop unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def by auto hence "prices Mkt asset n \ borel_measurable (nat_filtration n)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration n" "G n" "prices Mkt asset n" borel] unfolding adapt_stoch_proc_def by auto thus "finite (prices Mkt asset n ` space M)" using nat_filtration_vimage_finite[of "prices Mkt asset n"] by simp qed show "\n. \asset\support_set pf. finite (pf asset n ` space M)" proof (intro allI ballI) fix n fix asset assume "asset \ support_set pf" hence "pf asset n \ borel_measurable (G n)" using pfprop predict_imp_adapt[of "pf asset"] unfolding replicating_portfolio_def trading_strategy_def adapt_stoch_proc_def by auto hence "pf asset n \ borel_measurable (nat_filtration n)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration n" "G n" "pf asset n" borel] unfolding adapt_stoch_proc_def by auto thus "finite (pf asset n ` space M)" using nat_filtration_vimage_finite[of "pf asset n"] by simp qed qed moreover have "integral\<^sup>L N dpf = (\ w\ range (pseudo_proj_True matur). (prod (prob_component q w) {0.. borel_measurable (G matur)" using assms discounted_measurable[of pyf "G matur"] unfolding dpf_def by simp thus "dpf \ borel_measurable (nat_filtration matur)" using stock_filtration stoch_proc_subalg_nat_filt geom_rand_walk_borel_adapted measurable_from_subalg[of "nat_filtration matur" "G matur" dpf] unfolding adapt_stoch_proc_def by auto qed ultimately show ?thesis unfolding dpf_def q_def by simp qed end \ No newline at end of file diff --git a/thys/DiscretePricing/Fair_Price.thy b/thys/DiscretePricing/Fair_Price.thy --- a/thys/DiscretePricing/Fair_Price.thy +++ b/thys/DiscretePricing/Fair_Price.thy @@ -1,3522 +1,3510 @@ (* Title: Fair_Price.thy Author: Mnacho Echenim, Univ. Grenoble Alpes *) section \Fair Prices\ text \This section contains the formalization of financial notions, such as markets, price processes, portfolios, arbitrages, fair prices, etc. It also defines risk-neutral probability spaces, and proves the main result about the fair price of a derivative in a risk-neutral probability space, namely that this fair price is equal to the expectation of the discounted value of the derivative's payoff.\ theory Fair_Price imports Filtration Martingale Geometric_Random_Walk begin subsection \Preliminary results\ lemma (in prob_space) finite_borel_measurable_integrable: assumes "f\ borel_measurable M" and "finite (f`(space M))" shows "integrable M f" proof - have "simple_function M f" using assms by (simp add: simple_function_borel_measurable) moreover have "emeasure M {y \ space M. f y \ 0} \ \" by simp ultimately have "Bochner_Integration.simple_bochner_integrable M f" using Bochner_Integration.simple_bochner_integrable.simps by blast hence "has_bochner_integral M f (Bochner_Integration.simple_bochner_integral M f)" using has_bochner_integral_simple_bochner_integrable by auto thus ?thesis using integrable.simps by auto qed subsubsection \On the almost everywhere filter\ lemma AE_eq_trans[trans]: assumes "AE x in M. A x = B x" and "AE x in M. B x = C x" shows "AE x in M. A x = C x" using assms by auto abbreviation AEeq where "AEeq M X Y \ AE w in M. X w = Y w" lemma AE_add: assumes "AE w in M. f w = g w" and "AE w in M. f' w = g' w" shows "AE w in M. f w + f' w = g w + g' w" using assms by auto lemma AE_sum: assumes "finite I" and "\ i\I. AE w in M. f i w = g i w" shows "AE w in M. (\i\ I. f i w) = (\i\ I. g i w)" using assms(1) subset_refl[of I] proof (induct rule: finite_subset_induct) case empty then show ?case by simp next case (insert a F) have "AEeq M (f a) (g a)" using assms(2) insert.hyps(2) by auto have "AE w in M. (\i\ insert a F. f i w) = f a w + (\i\ F. f i w)" by (simp add: insert.hyps(1) insert.hyps(3)) also have "AE w in M. f a w + (\i\ F. f i w) = g a w + (\i\ F. f i w)" using \AEeq M (f a) (g a)\ by auto also have "AE w in M. g a w + (\i\ F. f i w) = g a w + (\i\ F. g i w)" using insert.hyps(4) by auto also have "AE w in M. g a w + (\i\ F. g i w) = (\i\ insert a F. g i w)" by (simp add: insert.hyps(1) insert.hyps(3)) finally show ?case by auto qed lemma AE_eq_cst: assumes "AE w in M. (\w. c) w = (\w. d) w" and "emeasure M (space M) \ 0" shows "c = d" proof (rule ccontr) assume "c \ d" from \AE w in M. (\w. c) w = (\w. d) w\ obtain N where Nprops: "{w\ space M. \(\w. c) w = (\w. d) w} \ N" "N\ sets M" "emeasure M N = 0" by (force elim:AE_E) have "\w\ space M. (\w. c) w \ (\w. d) w" using \c\ d\ by simp hence "{w\ space M. (\w. c) w \ (\w. d) w} = space M" by auto hence "space M\ N" using Nprops by auto thus False using \emeasure M N = 0\ assms by (meson Nprops(2) \emeasure M (space M) \ 0\ \emeasure M N = 0\ \space M \ N\ emeasure_eq_0) qed subsubsection \On conditional expectations\ lemma (in prob_space) subalgebra_sigma_finite: assumes "subalgebra M N" shows "sigma_finite_subalgebra M N" unfolding sigma_finite_subalgebra_def by (simp add: assms prob_space_axioms prob_space_imp_sigma_finite prob_space_restr_to_subalg) lemma (in prob_space) trivial_subalg_cond_expect_AE: assumes "subalgebra M N" and "sets N = {{}, space M}" and "integrable M f" shows "AE x in M. real_cond_exp M N f x = (\x. expectation f) x" proof (intro sigma_finite_subalgebra.real_cond_exp_charact) show "sigma_finite_subalgebra M N" by (simp add: assms(1) subalgebra_sigma_finite) show "integrable M f" using assms by simp show "integrable M (\x. expectation f)" by auto show "(\x. expectation f) \ borel_measurable N" by simp show "\A. A \ sets N \ set_lebesgue_integral M A f = (\x\A. expectation f\M)" proof - fix A assume "A \ sets N" show "set_lebesgue_integral M A f = (\x\A. expectation f\M)" proof (cases "A = {}") case True thus ?thesis by (simp add: set_lebesgue_integral_def) next case False hence "A = space M" using assms \A\ sets N\ by auto have "set_lebesgue_integral M A f = expectation f" using \A = space M\ by (metis (mono_tags, lifting) Bochner_Integration.integral_cong indicator_simps(1) scaleR_one set_lebesgue_integral_def) also have "... = (\x\A. expectation f\M)" using \A = space M\ by (auto simp add:prob_space set_lebesgue_integral_def) finally show ?thesis . qed qed qed lemma (in prob_space) triv_subalg_borel_eq: assumes "subalgebra M F" and "sets F = {{}, space M}" and "AE x in M. f x = (c::'b::{t2_space})" and "f\ borel_measurable F" shows "\x\ space M. f x = c" proof fix x assume "x\ space M" have "space M = space F" using assms by (simp add:subalgebra_def) hence "x\ space F" using \x\ space M\ by simp have "space M \ {}" by (simp add:not_empty) hence "\d. \y\ space F. f y = d" by (metis assms(1) assms(2) assms(4) subalgebra_def triv_measurable_cst) from this obtain d where "\y \space F. f y = d" by auto hence "f x = d" using \x\ space F\ by simp also have "... = c" proof (rule ccontr) assume "d\ c" from \AE x in M. f x= c\ obtain N where Nprops: "{x\ space M. \f x = c} \ N" "N\ sets M" "emeasure M N = 0" by (force elim:AE_E) have "space M \ {x\ space M. \f x = c}" using \\y \space F. f y = d\ \space M = space F\ \d\ c\ by auto hence "space M\ N" using Nprops by auto thus False using \emeasure M N = 0\ emeasure_space_1 Nprops(2) emeasure_mono by fastforce qed finally show "f x = c" . qed lemma (in prob_space) trivial_subalg_cond_expect_eq: assumes "subalgebra M N" and "sets N = {{}, space M}" and "integrable M f" shows "\x\ space M. real_cond_exp M N f x = expectation f" proof (rule triv_subalg_borel_eq) show "subalgebra M N" "sets N = {{}, space M}" using assms by auto show "real_cond_exp M N f \ borel_measurable N" by simp show "AE x in M. real_cond_exp M N f x = expectation f" by (rule trivial_subalg_cond_expect_AE, (auto simp add:assms)) qed lemma (in sigma_finite_subalgebra) real_cond_exp_cong': assumes "\w \ space M. f w = g w" and "f\ borel_measurable M" shows "AE w in M. real_cond_exp M F f w = real_cond_exp M F g w" proof (rule real_cond_exp_cong) show "AE w in M. f w = g w" using assms by simp show "f\ borel_measurable M" using assms by simp show "g\ borel_measurable M" using assms by (metis measurable_cong) qed lemma (in sigma_finite_subalgebra) real_cond_exp_bsum : fixes f::"'b \ 'a \ real" assumes [measurable]: "\i. i\I \ integrable M (f i)" shows "AE x in M. real_cond_exp M F (\x. \i\I. f i x) x = (\i\I. real_cond_exp M F (f i) x)" proof (rule real_cond_exp_charact) fix A assume [measurable]: "A \ sets F" then have A_meas [measurable]: "A \ sets M" by (meson subsetD subalg subalgebra_def) have *: "\i. i \ I \ integrable M (\x. indicator A x * f i x)" using integrable_mult_indicator[OF \A \ sets M\ assms(1)] by auto have **: "\i. i \ I \ integrable M (\x. indicator A x * real_cond_exp M F (f i) x)" using integrable_mult_indicator[OF \A \ sets M\ real_cond_exp_int(1)[OF assms(1)]] by auto have inti: "\i. i \ I \(\x. indicator A x * f i x \M) = (\x. indicator A x * real_cond_exp M F (f i) x \M)" using real_cond_exp_intg(2)[symmetric,of "indicator A" ] using "*" \A \ sets F\ assms borel_measurable_indicator by blast have "(\x\A. (\i\I. f i x)\M) = (\x. (\i\I. indicator A x * f i x)\M)" by (simp add: sum_distrib_left set_lebesgue_integral_def) also have "... = (\i\I. (\x. indicator A x * f i x \M))" using Bochner_Integration.integral_sum[of I M "\i x. indicator A x * f i x"] * by simp also have "... = (\i\I. (\x. indicator A x * real_cond_exp M F (f i) x \M))" using inti by auto also have "... = (\x. (\i\I. indicator A x * real_cond_exp M F (f i) x)\M)" by (rule Bochner_Integration.integral_sum[symmetric], simp add: **) also have "... = (\x\A. (\i\I. real_cond_exp M F (f i) x)\M)" by (simp add: sum_distrib_left set_lebesgue_integral_def) finally show "(\x\A. (\i\I. f i x)\M) = (\x\A. (\i\I. real_cond_exp M F (f i) x)\M)" by auto qed (auto simp add: assms real_cond_exp_int(1)[OF assms(1)]) subsection \Financial formalizations\ subsubsection \Markets\ definition stk_strict_subs::"'c set \ bool" where "stk_strict_subs S \ S \ UNIV" typedef ('a,'c) discrete_market = "{(s::('c set), a::'c \ (nat \ 'a \ real)). stk_strict_subs s}" unfolding stk_strict_subs_def by fastforce definition prices where "prices Mkt = (snd (Rep_discrete_market Mkt))" definition assets where "assets Mkt = UNIV" definition stocks where "stocks Mkt = (fst (Rep_discrete_market Mkt))" definition discrete_market_of where "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" lemma prices_of: shows "prices (discrete_market_of S A) = A" proof - have "stk_strict_subs (if (stk_strict_subs S) then S else {})" proof (cases "stk_strict_subs S") case True thus ?thesis by simp next case False thus ?thesis unfolding stk_strict_subs_def by simp qed hence fct: "((if (stk_strict_subs S) then S else {}), A) \ {(s, a). stk_strict_subs s}" by simp have "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" unfolding discrete_market_of_def by simp hence "Rep_discrete_market (discrete_market_of S A) = (if (stk_strict_subs S) then S else {},A)" using Abs_discrete_market_inverse[of "(if (stk_strict_subs S) then S else {}, A)"] fct by simp thus ?thesis unfolding prices_def by simp qed lemma stocks_of: assumes "UNIV \ S" shows "stocks (discrete_market_of S A) = S" proof - have "stk_strict_subs S" using assms unfolding stk_strict_subs_def by simp hence fct: "((if (stk_strict_subs S) then S else {}), A) \ {(s, a). stk_strict_subs s}" by simp have "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" unfolding discrete_market_of_def by simp hence "Rep_discrete_market (discrete_market_of S A) = (if (stk_strict_subs S) then S else {},A)" using Abs_discrete_market_inverse[of "(if (stk_strict_subs S) then S else {}, A)"] fct by simp thus ?thesis unfolding stocks_def using \stk_strict_subs S\ by simp qed lemma mkt_stocks_assets: shows "stk_strict_subs (stocks Mkt)" unfolding stocks_def prices_def by (metis Rep_discrete_market mem_Collect_eq split_beta') subsubsection \Quantity processes and portfolios\ text \These are functions that assign quantities to assets; each quantity is a stochastic process. Basic operations are defined on these processes.\ paragraph \Basic operations\ definition qty_empty where "qty_empty = (\ (x::'a) (n::nat) w. 0::real)" definition qty_single where "qty_single asset qt_proc = (qty_empty(asset := qt_proc))" definition qty_sum::"('b \ nat \ 'a \ real) \ ('b \ nat \ 'a \ real) \ ('b \ nat \ 'a \ real)" where "qty_sum pf1 pf2 = (\x n w. pf1 x n w + pf2 x n w)" definition qty_mult_comp::"('b \ nat \ 'a \ real) \ (nat \ 'a \ real) \ ('b \ nat \ 'a \ real)" where "qty_mult_comp pf1 qty = (\x n w. (pf1 x n w) * (qty n w))" definition qty_rem_comp::"('b \ nat \ 'a \ real) \ 'b \ ('b \ nat \ 'a \ real)" where "qty_rem_comp pf1 x = pf1(x:=(\n w. 0))" definition qty_replace_comp where "qty_replace_comp pf1 x pf2 = qty_sum (qty_rem_comp pf1 x) (qty_mult_comp pf2 (pf1 x))" paragraph \Support sets\ text \If p x n w is different from 0, this means that this quantity is held on interval ]n-1, n].\ definition support_set::"('b \ nat \ 'a \ real) \ 'b set" where "support_set p = {x. \ n w. p x n w \ 0}" lemma qty_empty_support_set: shows "support_set qty_empty = {}" unfolding support_set_def qty_empty_def by simp lemma sum_support_set: shows "support_set (qty_sum pf1 pf2) \ (support_set pf1) \ (support_set pf2)" proof (intro subsetI, rule ccontr) fix x assume "x\ support_set (qty_sum pf1 pf2)" and "x \ support_set pf1 \ support_set pf2" note xprops = this hence "\ n w. (qty_sum pf1 pf2) x n w \ 0" by (simp add: support_set_def) from this obtain n w where "(qty_sum pf1 pf2) x n w \ 0" by auto note nwprops = this have "pf1 x n w = 0" "pf2 x n w = 0" using xprops by (auto simp add:support_set_def) hence "(qty_sum pf1 pf2) x n w = 0" unfolding qty_sum_def by simp thus False using nwprops by simp qed lemma mult_comp_support_set: shows "support_set (qty_mult_comp pf1 qty) \ (support_set pf1)" proof (intro subsetI, rule ccontr) fix x assume "x\ support_set (qty_mult_comp pf1 qty)" and "x \ support_set pf1" note xprops = this hence "\ n w. (qty_mult_comp pf1 qty) x n w \ 0" by (simp add: support_set_def) from this obtain n w where "qty_mult_comp pf1 qty x n w \ 0" by auto note nwprops = this have "pf1 x n w = 0" using xprops by (simp add:support_set_def) hence "(qty_mult_comp pf1 qty) x n w = 0" unfolding qty_mult_comp_def by simp thus False using nwprops by simp qed lemma remove_comp_support_set: shows "support_set (qty_rem_comp pf1 x) \ ((support_set pf1) - {x})" proof (intro subsetI, rule ccontr) fix y assume "y\ support_set (qty_rem_comp pf1 x)" and "y \ support_set pf1 - {x}" note xprops = this hence "y\ support_set pf1 \ y = x" by simp have "\ n w. (qty_rem_comp pf1 x) y n w \ 0" using xprops by (simp add: support_set_def) from this obtain n w where "(qty_rem_comp pf1 x) y n w \ 0" by auto note nwprops = this show False proof (cases "y\ support_set pf1") case True hence "pf1 y n w = 0" using xprops by (simp add:support_set_def) hence "(qty_rem_comp pf1 x) x n w = 0" unfolding qty_rem_comp_def by simp thus ?thesis using nwprops by (metis \pf1 y n w = 0\ fun_upd_apply qty_rem_comp_def) next case False hence "y = x" using \y\ support_set pf1 \ y = x\ by simp hence "(qty_rem_comp pf1 x) x n w = 0" unfolding qty_rem_comp_def by simp thus ?thesis using nwprops by (simp add: \y = x\) qed qed lemma replace_comp_support_set: shows "support_set (qty_replace_comp pf1 x pf2) \ (support_set pf1 - {x}) \ support_set pf2" proof - have "support_set (qty_replace_comp pf1 x pf2) \ support_set (qty_rem_comp pf1 x) \ support_set (qty_mult_comp pf2 (pf1 x))" unfolding qty_replace_comp_def by (simp add:sum_support_set) also have "... \ (support_set pf1 - {x}) \ support_set pf2" using remove_comp_support_set mult_comp_support_set by (metis sup.mono) finally show ?thesis . qed lemma single_comp_support: shows "support_set (qty_single asset qty) \ {asset}" proof fix x assume "x\ support_set (qty_single asset qty)" show "x\ {asset}" proof (rule ccontr) assume "x\ {asset}" hence "x\ asset" by auto have "\ n w. qty_single asset qty x n w \ 0" using \x\ support_set (qty_single asset qty)\ by (simp add:support_set_def) from this obtain n w where "qty_single asset qty x n w \ 0" by auto thus False using \x\asset\ by (simp add: qty_single_def qty_empty_def) qed qed lemma single_comp_nz_support: assumes "\ n w. qty n w\ 0" shows "support_set (qty_single asset qty) = {asset}" proof show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) have "asset\ support_set (qty_single asset qty)" using assms unfolding support_set_def qty_single_def by simp thus "{asset} \ support_set (qty_single asset qty)" by auto qed paragraph \Portfolios\ definition portfolio where "portfolio p \ finite (support_set p)" definition stock_portfolio :: "('a, 'b) discrete_market \ ('b \ nat \ 'a \ real) \ bool" where "stock_portfolio Mkt p \ portfolio p \ support_set p \ stocks Mkt" lemma sum_portfolio: assumes "portfolio pf1" and "portfolio pf2" shows "portfolio (qty_sum pf1 pf2)" unfolding portfolio_def proof - have "support_set (qty_sum pf1 pf2) \ (support_set pf1) \ (support_set pf2)" by (simp add: sum_support_set) thus "finite (support_set (qty_sum pf1 pf2))" using assms unfolding portfolio_def using infinite_super by auto qed lemma sum_basic_support_set: assumes "stock_portfolio Mkt pf1" and "stock_portfolio Mkt pf2" shows "stock_portfolio Mkt (qty_sum pf1 pf2)" using assms sum_support_set[of pf1 pf2] unfolding stock_portfolio_def by (metis Diff_subset_conv gfp.leq_trans subset_Un_eq sum_portfolio) lemma mult_comp_portfolio: assumes "portfolio pf1" shows "portfolio (qty_mult_comp pf1 qty)" unfolding portfolio_def proof - have "support_set (qty_mult_comp pf1 qty) \ (support_set pf1)" by (simp add: mult_comp_support_set) thus "finite (support_set (qty_mult_comp pf1 qty))" using assms unfolding portfolio_def using infinite_super by auto qed lemma mult_comp_basic_support_set: assumes "stock_portfolio Mkt pf1" shows "stock_portfolio Mkt (qty_mult_comp pf1 qty)" using assms mult_comp_support_set[of pf1] unfolding stock_portfolio_def using mult_comp_portfolio by blast lemma remove_comp_portfolio: assumes "portfolio pf1" shows "portfolio (qty_rem_comp pf1 x)" unfolding portfolio_def proof - have "support_set (qty_rem_comp pf1 x) \ (support_set pf1)" using remove_comp_support_set[of pf1 x] by blast thus "finite (support_set (qty_rem_comp pf1 x))" using assms unfolding portfolio_def using infinite_super by auto qed lemma remove_comp_basic_support_set: assumes "stock_portfolio Mkt pf1" shows "stock_portfolio Mkt (qty_mult_comp pf1 qty)" using assms mult_comp_support_set[of pf1] unfolding stock_portfolio_def using mult_comp_portfolio by blast lemma replace_comp_portfolio: assumes "portfolio pf1" and "portfolio pf2" shows "portfolio (qty_replace_comp pf1 x pf2)" unfolding portfolio_def proof - have "support_set (qty_replace_comp pf1 x pf2) \ (support_set pf1) \ (support_set pf2)" using replace_comp_support_set[of pf1 x pf2] by blast thus "finite (support_set (qty_replace_comp pf1 x pf2))" using assms unfolding portfolio_def using infinite_super by auto qed lemma replace_comp_stocks: assumes "support_set pf1 \ stocks Mkt \ {x}" and "support_set pf2 \ stocks Mkt" shows "support_set (qty_replace_comp pf1 x pf2) \ stocks Mkt" proof - have "support_set (qty_rem_comp pf1 x) \ stocks Mkt" using assms(1) remove_comp_support_set by fastforce moreover have "support_set (qty_mult_comp pf2 (pf1 x)) \ stocks Mkt" using assms mult_comp_support_set by fastforce ultimately show ?thesis unfolding qty_replace_comp_def using sum_support_set by fastforce qed lemma single_comp_portfolio: shows "portfolio (qty_single asset qty)" by (meson finite.emptyI finite.insertI finite_subset portfolio_def single_comp_support) paragraph \Value processes\ definition val_process where "val_process Mkt p = (if (\ (portfolio p)) then (\ n w. 0) else (\ n w . (sum (\x. ((prices Mkt) x n w) * (p x (Suc n) w)) (support_set p))))" lemma subset_val_process': assumes "finite A" and "support_set p \ A" shows "val_process Mkt p n w = (sum (\x. ((prices Mkt) x n w) * (p x (Suc n) w)) A)" proof - have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x n w) * (p x (Suc n) w)) = 0" by simp hence "val_process Mkt p n w = (\x\ (support_set p). ((prices Mkt) x n w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x n w) * (p x (Suc n) w))" unfolding val_process_def using \portfolio p\ by simp also have "... = (\ x\ A. ((prices Mkt) x n w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "val_process Mkt p n w = (\ x\ A. ((prices Mkt) x n w) * (p x (Suc n) w))" . qed lemma sum_val_process: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. val_process Mkt (qty_sum pf1 pf2) n w = (val_process Mkt pf1) n w + (val_process Mkt pf2) n w" proof (intro allI) fix n w have vp1: "val_process Mkt pf1 n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf1 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf1 \ support_set pf1 \ support_set pf2" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_val_process') qed have vp2: "val_process Mkt pf2 n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf2 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf2 \ support_set pf2 \ support_set pf1" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_val_process') qed have pf:"portfolio (qty_sum pf1 pf2)" using assms by (simp add:sum_portfolio) have fin:"finite (support_set pf1 \ support_set pf2)" using assms finite_Un unfolding portfolio_def by auto have "(val_process Mkt pf1) n w + (val_process Mkt pf2) n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf1 x (Suc n) w)) + (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf2 x (Suc n) w))" using vp1 vp2 by simp also have "... = (\ x\ (support_set pf1)\ (support_set pf2). (((prices Mkt) x n w) * (pf1 x (Suc n) w)) + ((prices Mkt) x n w) * (pf2 x (Suc n) w))" by (simp add: sum.distrib) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * ((pf1 x (Suc n) w) + (pf2 x (Suc n) w)))" by (simp add: distrib_left) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * ((qty_sum pf1 pf2) x (Suc n) w))" by (simp add: qty_sum_def) also have "... = (\ x\ (support_set (qty_sum pf1 pf2)). ((prices Mkt) x n w) * ((qty_sum pf1 pf2) x (Suc n) w))" using sum_support_set[of pf1 pf2] subset_val_process'[of "support_set pf1\ support_set pf2" "qty_sum pf1 pf2"] pf fin unfolding val_process_def by simp also have "... = val_process Mkt (qty_sum pf1 pf2) n w" by (metis (no_types, lifting) pf sum.cong val_process_def) finally have "(val_process Mkt pf1) n w + (val_process Mkt pf2) n w = val_process Mkt (qty_sum pf1 pf2) n w" . thus "val_process Mkt (qty_sum pf1 pf2) n w = (val_process Mkt pf1) n w + (val_process Mkt pf2) n w" .. qed lemma mult_comp_val_process: assumes "portfolio pf1" shows "\n w. val_process Mkt (qty_mult_comp pf1 qty) n w = ((val_process Mkt pf1) n w) * (qty (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_mult_comp pf1 qty)" using assms by (simp add:mult_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto have "((val_process Mkt pf1) n w) * (qty (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x n w) * (pf1 x (Suc n) w))*(qty (Suc n) w)" unfolding val_process_def using assms by simp also have "... = (\ x\ (support_set pf1). (((prices Mkt) x n w) * (pf1 x (Suc n) w) * (qty (Suc n) w)))" using sum_distrib_right by auto also have "... = (\ x\ (support_set pf1). ((prices Mkt) x n w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" unfolding qty_mult_comp_def by (simp add: mult.commute mult.left_commute) also have "... = (\ x\ (support_set (qty_mult_comp pf1 qty)). ((prices Mkt) x n w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" using mult_comp_support_set[of pf1] subset_val_process'[of "support_set pf1" "qty_mult_comp pf1 qty"] pf fin unfolding val_process_def by simp also have "... = val_process Mkt (qty_mult_comp pf1 qty) n w" by (metis (no_types, lifting) pf sum.cong val_process_def) finally have "(val_process Mkt pf1) n w * (qty (Suc n) w) = val_process Mkt (qty_mult_comp pf1 qty) n w" . thus "val_process Mkt (qty_mult_comp pf1 qty) n w = (val_process Mkt pf1) n w * (qty (Suc n) w)" .. qed lemma remove_comp_values: assumes "x \ y" shows "\n w. pf1 x n w = (qty_rem_comp pf1 y) x n w" proof (intro allI) fix n w show "pf1 x n w = (qty_rem_comp pf1 y) x n w" by (simp add: assms qty_rem_comp_def) qed lemma remove_comp_val_process: assumes "portfolio pf1" shows "\n w. val_process Mkt (qty_rem_comp pf1 y) n w = ((val_process Mkt pf1) n w) - (prices Mkt y n w)* (pf1 y (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_rem_comp pf1 y)" using assms by (simp add:remove_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto hence fin2: "finite (support_set pf1 - {y})" by simp have "((val_process Mkt pf1) n w) = (\ x\ (support_set pf1). ((prices Mkt) x n w) * (pf1 x (Suc n) w))" unfolding val_process_def using assms by simp also have "... = (\ x\ (support_set pf1 - {y}). (((prices Mkt) x n w) * (pf1 x (Suc n) w))) + (prices Mkt y n w)* (pf1 y (Suc n) w)" proof (cases "y\ support_set pf1") case True thus ?thesis by (simp add: fin sum_diff1) next case False hence "pf1 y (Suc n) w = 0" unfolding support_set_def by simp thus ?thesis by (simp add: fin sum_diff1) qed also have "... = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x n w) * ((qty_rem_comp pf1 y) x (Suc n) w)) + (prices Mkt y n w)* (pf1 y (Suc n) w)" proof - have "(\ x\ (support_set pf1 - {y}). (((prices Mkt) x n w) * (pf1 x (Suc n) w))) = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x n w) * ((qty_rem_comp pf1 y) x (Suc n) w))" proof (rule sum.cong,simp) fix x assume "x\ support_set pf1 - {y}" show "prices Mkt x n w * pf1 x (Suc n) w = prices Mkt x n w * qty_rem_comp pf1 y x (Suc n) w" using remove_comp_values by (metis DiffD2 \x \ support_set pf1 - {y}\ singletonI) qed thus ?thesis by simp qed also have "... = (val_process Mkt (qty_rem_comp pf1 y) n w) + (prices Mkt y n w)* (pf1 y (Suc n) w)" using subset_val_process'[of "support_set pf1 - {y}" "qty_rem_comp pf1 y"] fin2 by (simp add: remove_comp_support_set) finally have "(val_process Mkt pf1) n w = (val_process Mkt (qty_rem_comp pf1 y) n w) + (prices Mkt y n w)* (pf1 y (Suc n) w)" . thus "val_process Mkt (qty_rem_comp pf1 y) n w = ((val_process Mkt pf1) n w) - (prices Mkt y n w)* (pf1 y (Suc n) w)" by simp qed lemma replace_comp_val_process: assumes "\n w. prices Mkt x n w = val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" shows "\n w. val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt pf1 n w" proof (intro allI) fix n w have "val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt (qty_rem_comp pf1 x) n w + val_process Mkt (qty_mult_comp pf2 (pf1 x)) n w" unfolding qty_replace_comp_def using assms sum_val_process[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = val_process Mkt pf1 n w - (prices Mkt x n w * pf1 x (Suc n) w) + val_process Mkt pf2 n w * pf1 x (Suc n) w" by (simp add: assms(2) assms(3) mult_comp_val_process remove_comp_val_process) also have "... = val_process Mkt pf1 n w" using assms by simp finally show "val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt pf1 n w" . qed lemma qty_single_val_process: shows "val_process Mkt (qty_single asset qty) n w = prices Mkt asset n w * qty (Suc n) w" proof - have "val_process Mkt (qty_single asset qty) n w = (sum (\x. ((prices Mkt) x n w) * ((qty_single asset qty) x (Suc n) w)) {asset})" proof (rule subset_val_process') show "finite {asset}" by simp show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) qed also have "... = prices Mkt asset n w * qty (Suc n) w" unfolding qty_single_def by simp finally show ?thesis . qed subsubsection \Trading strategies\ locale disc_equity_market = triv_init_disc_filtr_prob_space + fixes Mkt::"('a,'b) discrete_market" paragraph \Discrete predictable processes\ paragraph \Trading strategy\ definition (in disc_filtr_prob_space) trading_strategy where "trading_strategy p \ portfolio p \ (\asset \ support_set p. borel_predict_stoch_proc F (p asset))" definition (in disc_filtr_prob_space) support_adapt:: "('a, 'b) discrete_market \ ('b \ nat \ 'a \ real) \ bool" where "support_adapt Mkt pf \ (\ asset \ support_set pf. borel_adapt_stoch_proc F (prices Mkt asset))" lemma (in disc_filtr_prob_space) quantity_adapted: assumes "\ asset \ support_set p. p asset (Suc n) \ borel_measurable (F n)" "\asset \ support_set p. prices Mkt asset n \ borel_measurable (F n)" shows "val_process Mkt p n \ borel_measurable (F n)" proof (cases "portfolio p") case False have "val_process Mkt p = (\ n w. 0)" unfolding val_process_def using False by simp thus "?thesis" by simp next case True hence "val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w)" unfolding val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc n) \ borel_measurable (F n)" using assms unfolding trading_strategy_def adapt_stoch_proc_def by simp moreover have "prices Mkt asset n \ borel_measurable (F n)" using \asset \ support_set p\ assms(2) unfolding support_adapt_def by (simp add: adapt_stoch_proc_def) ultimately show "(\x. prices Mkt asset n x * p asset (Suc n) x) \ borel_measurable (F n)" by simp qed ultimately show "val_process Mkt p n \ borel_measurable (F n)" by simp qed lemma (in disc_filtr_prob_space) trading_strategy_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (val_process Mkt p)" unfolding support_adapt_def proof (cases "portfolio p") case False have "val_process Mkt p = (\ n w. 0)" unfolding val_process_def using False by simp thus "borel_adapt_stoch_proc F (val_process Mkt p)" by (simp add: constant_process_borel_adapted) next case True show ?thesis unfolding adapt_stoch_proc_def proof fix n have "val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w)" unfolding val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc n) \ borel_measurable (F n)" using assms unfolding trading_strategy_def predict_stoch_proc_def by simp moreover have "prices Mkt asset n \ borel_measurable (F n)" using \asset \ support_set p\ assms(2) unfolding support_adapt_def by (simp add:adapt_stoch_proc_def) ultimately show "(\x. prices Mkt asset n x * p asset (Suc n) x) \ borel_measurable (F n)" by simp qed ultimately show "val_process Mkt p n \ borel_measurable (F n)" by simp qed qed lemma (in disc_equity_market) ats_val_process_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (val_process Mkt p)" unfolding support_adapt_def by (meson assms(1) assms(2) subsetCE trading_strategy_adapted) lemma (in disc_equity_market) trading_strategy_init: assumes "trading_strategy p" and "support_adapt Mkt p" shows "\c. \w \ space M. val_process Mkt p 0 w = c" using assms adapted_init ats_val_process_adapted by simp definition (in disc_equity_market) initial_value where "initial_value pf = constant_image (val_process Mkt pf 0)" lemma (in disc_equity_market) initial_valueI: assumes "trading_strategy pf" and "support_adapt Mkt pf" shows "\w\ space M. val_process Mkt pf 0 w = initial_value pf" unfolding initial_value_def proof (rule constant_imageI) show "\c. \w\space M. val_process Mkt pf 0 w = c" using trading_strategy_init assms by simp qed lemma (in disc_equity_market) inc_predict_support_trading_strat: assumes "trading_strategy pf1" shows "\ asset \ support_set pf1 \ B. borel_predict_stoch_proc F (pf1 asset)" proof fix asset assume "asset \ support_set pf1 \ B" show "borel_predict_stoch_proc F (pf1 asset)" proof (cases "asset \ support_set pf1") case True thus ?thesis using assms unfolding trading_strategy_def by simp next case False hence "\n w. pf1 asset n w = 0" unfolding support_set_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "pf1 asset 0 \ measurable (F 0) borel" using \\n w. pf1 asset n w = 0\ by (simp add: borel_measurable_const measurable_cong) next show "\n. pf1 asset (Suc n) \ borel_measurable (F n)" proof fix n have "\w. pf1 asset (Suc n) w = 0" using \\n w. pf1 asset n w = 0\ by simp have "0\ space borel" by simp thus "pf1 asset (Suc n) \ measurable (F n) borel" using measurable_const[of 0 borel "F n"] by (metis \0 \ space borel \ (\x. 0) \ borel_measurable (F n)\ \0 \ space borel\ \\n w. pf1 asset n w = 0\ measurable_cong) qed qed qed qed lemma (in disc_equity_market) inc_predict_support_trading_strat': assumes "trading_strategy pf1" and "asset \ support_set pf1\ B" shows "borel_predict_stoch_proc F (pf1 asset)" proof (cases "asset \ support_set pf1") case True thus ?thesis using assms unfolding trading_strategy_def by simp next case False hence "\n w. pf1 asset n w = 0" unfolding support_set_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "pf1 asset 0 \ measurable (F 0) borel" using \\n w. pf1 asset n w = 0\ by (simp add: borel_measurable_const measurable_cong) next show "\n. pf1 asset (Suc n) \ borel_measurable (F n)" proof fix n have "\w. pf1 asset (Suc n) w = 0" using \\n w. pf1 asset n w = 0\ by simp have "0\ space borel" by simp thus "pf1 asset (Suc n) \ measurable (F n) borel" using measurable_const[of 0 borel "F n"] by (metis \0 \ space borel \ (\x. 0) \ borel_measurable (F n)\ \0 \ space borel\ \\n w. pf1 asset n w = 0\ measurable_cong) qed qed qed lemma (in disc_equity_market) inc_support_trading_strat: assumes "trading_strategy pf1" shows "\ asset \ support_set pf1 \ B. borel_adapt_stoch_proc F (pf1 asset)" using assms by (simp add: inc_predict_support_trading_strat predict_imp_adapt) lemma (in disc_equity_market) qty_empty_trading_strat: shows "trading_strategy qty_empty" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio qty_empty" by (metis fun_upd_triv qty_single_def single_comp_portfolio) show "\asset. asset \ support_set qty_empty \ borel_predict_stoch_proc F (qty_empty asset)" using qty_empty_support_set by auto qed lemma (in disc_equity_market) sum_trading_strat: assumes "trading_strategy pf1" and "trading_strategy pf2" shows "trading_strategy (qty_sum pf1 pf2)" proof - have "\ asset \ support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)" using assms by (simp add: inc_predict_support_trading_strat) have "\ asset \ support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)" using assms by (simp add: inc_predict_support_trading_strat) have "\ asset \ support_set pf1 \ support_set pf2. borel_predict_stoch_proc F ((qty_sum pf1 pf2) asset)" proof fix asset assume "asset \ support_set pf1 \ support_set pf2" show "borel_predict_stoch_proc F (qty_sum pf1 pf2 asset)" unfolding predict_stoch_proc_def qty_sum_def proof show "(\w. pf1 asset 0 w + pf2 asset 0 w) \ borel_measurable (F 0)" proof - have "(\w. pf1 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast moreover have "(\w. pf2 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast ultimately show ?thesis by simp qed next show "\n. (\w. pf1 asset (Suc n) w + pf2 asset (Suc n) w) \ borel_measurable (F n)" proof fix n have "(\w. pf1 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast moreover have "(\w. pf2 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast ultimately show "(\w. pf1 asset (Suc n) w + pf2 asset (Suc n) w) \ borel_measurable (F n)" by simp qed qed qed thus ?thesis unfolding trading_strategy_def using sum_support_set[of pf1 pf2] by (meson assms(1) assms(2) subsetCE sum_portfolio trading_strategy_def) qed lemma (in disc_equity_market) mult_comp_trading_strat: assumes "trading_strategy pf1" and "borel_predict_stoch_proc F qty" shows "trading_strategy (qty_mult_comp pf1 qty)" proof - have "\ asset \ support_set pf1. borel_predict_stoch_proc F (pf1 asset)" using assms unfolding trading_strategy_def by simp have "\ asset \ support_set pf1. borel_predict_stoch_proc F (qty_mult_comp pf1 qty asset)" unfolding predict_stoch_proc_def qty_mult_comp_def proof (intro ballI conjI) fix asset assume "asset \ support_set pf1" show "(\w. pf1 asset 0 w * qty 0 w) \ borel_measurable (F 0)" proof - have "(\w. pf1 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ predict_stoch_proc_def by auto moreover have "(\w. qty 0 w) \ borel_measurable (F 0)" using assms predict_stoch_proc_def by auto ultimately show "(\w. pf1 asset 0 w * qty 0 w) \ borel_measurable (F 0)" by simp qed show "\n. (\w. pf1 asset (Suc n) w * qty (Suc n) w) \ borel_measurable (F n)" proof fix n have "(\w. pf1 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ predict_stoch_proc_def by blast moreover have "(\w. qty (Suc n) w) \ borel_measurable (F n)" using assms predict_stoch_proc_def by blast ultimately show "(\w. pf1 asset (Suc n) w * qty (Suc n) w) \ borel_measurable (F n)" by simp qed qed thus ?thesis unfolding trading_strategy_def using mult_comp_support_set[of pf1 qty] by (meson assms(1) mult_comp_portfolio subsetCE trading_strategy_def) qed lemma (in disc_equity_market) remove_comp_trading_strat: assumes "trading_strategy pf1" shows "trading_strategy (qty_rem_comp pf1 x)" proof - have "\ asset \ support_set pf1. borel_predict_stoch_proc F (pf1 asset)" using assms unfolding trading_strategy_def by simp have "\ asset \ support_set pf1. borel_predict_stoch_proc F (qty_rem_comp pf1 x asset)" unfolding predict_stoch_proc_def qty_rem_comp_def proof (intro ballI conjI) fix asset assume "asset \ support_set pf1" show "(pf1(x := \n w. 0)) asset 0 \ borel_measurable (F 0)" proof - show "(\w. (pf1(x := \n w. 0)) asset 0 w) \ borel_measurable (F 0)" proof (cases "asset = x") case True thus ?thesis by simp next case False thus "(\w. (pf1(x := \n w. 0)) asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ by (simp add: predict_stoch_proc_def) qed qed show "\n. (\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" proof fix n show "(\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" proof (cases "asset = x") case True thus ?thesis by simp next case False thus "(\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ by (simp add: predict_stoch_proc_def) qed qed qed thus ?thesis unfolding trading_strategy_def using remove_comp_support_set[of pf1 x] by (metis Diff_empty assms remove_comp_portfolio subsetCE subset_Diff_insert trading_strategy_def) qed lemma (in disc_equity_market) replace_comp_trading_strat: assumes "trading_strategy pf1" and "trading_strategy pf2" shows "trading_strategy (qty_replace_comp pf1 x pf2)" unfolding qty_replace_comp_def proof (rule sum_trading_strat) show "trading_strategy (qty_rem_comp pf1 x)" using assms by (simp add: remove_comp_trading_strat) show "trading_strategy (qty_mult_comp pf2 (pf1 x))" proof (cases "x\ support_set pf1") case True hence "borel_predict_stoch_proc F (pf1 x)" using assms unfolding trading_strategy_def by simp thus ?thesis using assms by (simp add: mult_comp_trading_strat) next case False thus ?thesis - proof - - obtain nn :: "'c \ ('c \ nat \ 'a \ real) \ nat" and aa :: "'c \ ('c \ nat \ 'a \ real) \ 'a" where - "\x0 x1. (\v2 v3. x1 x0 v2 v3 \ 0) = (x1 x0 (nn x0 x1) (aa x0 x1) \ 0)" - by moura - then have "\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))" - by auto - then show ?thesis - proof - - have "\f c n a. qty_mult_comp f (pf1 x) (c::'c) n a = 0" - by (metis False \\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))\ mult.commute mult_zero_left qty_mult_comp_def support_set_def) - then show ?thesis - by (metis (no_types) \\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))\ assms(2) mult_comp_portfolio support_set_def trading_strategy_def) - qed - qed + by (meson UnCI assms(1) assms(2) disc_equity_market.inc_predict_support_trading_strat + disc_equity_market_axioms insertI1 mult_comp_trading_strat) qed qed subsubsection \Self-financing portfolios\ paragraph \Closing value process\ fun up_cl_proc where "up_cl_proc Mkt p 0 = val_process Mkt p 0" | "up_cl_proc Mkt p (Suc n) = (\w. \x\support_set p. prices Mkt x (Suc n) w * p x (Suc n) w)" definition cls_val_process where "cls_val_process Mkt p = (if (\ (portfolio p)) then (\ n w. 0) else (\ n w . up_cl_proc Mkt p n w))" lemma (in disc_filtr_prob_space) quantity_updated_borel: assumes "\n. \ asset \ support_set p. p asset (Suc n) \ borel_measurable (F n)" and "\n. \asset \ support_set p. prices Mkt asset n \ borel_measurable (F n)" shows "\n. cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "portfolio p") case False have "cls_val_process Mkt p = (\ n w. 0)" unfolding cls_val_process_def using False by simp thus "?thesis" by simp next case True show "\n. cls_val_process Mkt p n \ borel_measurable (F n)" proof fix n show "cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "n = 0") case False hence "\m. n = Suc m" using old.nat.exhaust by auto from this obtain m where "n = Suc m" by auto have "cls_val_process Mkt p (Suc m) = (\w. \x\support_set p. prices Mkt x (Suc m) w * p x (Suc m) w)" unfolding cls_val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x (Suc m) w * p x (Suc m) w) \ borel_measurable (F (Suc m))" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc m) \ borel_measurable (F m)" using assms unfolding trading_strategy_def adapt_stoch_proc_def by simp hence "p asset (Suc m) \ borel_measurable (F (Suc m))" using Suc_n_not_le_n increasing_measurable_info nat_le_linear by blast moreover have "prices Mkt asset (Suc m) \ borel_measurable (F (Suc m))" using \asset \ support_set p\ assms(2) unfolding support_adapt_def adapt_stoch_proc_def by blast ultimately show "(\x. prices Mkt asset (Suc m) x * p asset (Suc m) x) \ borel_measurable (F (Suc m))" by simp qed ultimately have "cls_val_process Mkt p (Suc m) \ borel_measurable (F (Suc m))" by simp thus ?thesis using \n = Suc m\ by simp next case True thus "cls_val_process Mkt p n \ borel_measurable (F n)" by (metis (no_types, lifting) assms(1) assms(2) quantity_adapted up_cl_proc.simps(1) cls_val_process_def val_process_def) qed qed qed lemma (in disc_equity_market) cls_val_process_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (cls_val_process Mkt p)" proof (cases "portfolio p") case False have "cls_val_process Mkt p = (\ n w. 0)" unfolding cls_val_process_def using False by simp thus "borel_adapt_stoch_proc F (cls_val_process Mkt p)" by (simp add: constant_process_borel_adapted) next case True show ?thesis unfolding adapt_stoch_proc_def proof fix n show "cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "n = 0") case True thus "cls_val_process Mkt p n \ borel_measurable (F n)" using up_cl_proc.simps(1) assms by (metis (no_types, lifting) adapt_stoch_proc_def ats_val_process_adapted cls_val_process_def val_process_def) next case False hence "\m. Suc m = n" using not0_implies_Suc by blast from this obtain m where "Suc m = n" by auto hence "cls_val_process Mkt p n = up_cl_proc Mkt p n" unfolding cls_val_process_def using True by simp also have "... = (\w. \x\support_set p. prices Mkt x n w * p x n w)" using up_cl_proc.simps(2) \Suc m = n\ by auto finally have "cls_val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x n w)" . moreover have "(\w. \x\support_set p. prices Mkt x n w * p x n w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset n \ borel_measurable (F n)" using assms unfolding trading_strategy_def predict_stoch_proc_def using Suc_n_not_le_n \Suc m = n\ increasing_measurable_info nat_le_linear by blast moreover have "prices Mkt asset n \ borel_measurable (F n)" using assms \asset \ support_set p\ unfolding support_adapt_def adapt_stoch_proc_def using stock_portfolio_def by blast ultimately show "(\x. prices Mkt asset n x * p asset n x) \ borel_measurable (F n)" by simp qed ultimately show "cls_val_process Mkt p n \ borel_measurable (F n)" by simp qed qed qed lemma subset_cls_val_process: assumes "finite A" and "support_set p \ A" shows "\n w. cls_val_process Mkt p (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) A)" proof (intro allI) fix n::nat fix w::'b have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) = 0" by simp hence "cls_val_process Mkt p (Suc n) w = (\x\ (support_set p). ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" unfolding cls_val_process_def using \portfolio p\ up_cl_proc.simps(2)[of Mkt p n] by simp also have "... = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "cls_val_process Mkt p (Suc n) w = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" . qed lemma subset_cls_val_process': assumes "finite A" and "support_set p \ A" shows "cls_val_process Mkt p (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) A)" proof - have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) = 0" by simp hence "cls_val_process Mkt p (Suc n) w = (\x\ (support_set p). ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" unfolding cls_val_process_def using \portfolio p\ up_cl_proc.simps(2)[of Mkt p n] by simp also have "... = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "cls_val_process Mkt p (Suc n) w = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" . qed lemma sum_cls_val_process_Suc: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w" proof (intro allI) fix n w have vp1: "cls_val_process Mkt pf1 (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf1 \ support_set pf1 \ support_set pf2" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_cls_val_process) qed have vp2: "cls_val_process Mkt pf2 (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf2 \ support_set pf2 \ support_set pf1" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (auto simp add: subset_cls_val_process) qed have pf:"portfolio (qty_sum pf1 pf2)" using assms by (simp add:sum_portfolio) have fin:"finite (support_set pf1 \ support_set pf2)" using assms finite_Un unfolding portfolio_def by auto have "(cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w)) + (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" using vp1 vp2 by simp also have "... = (\ x\ (support_set pf1)\ (support_set pf2). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w)) + ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" by (simp add: sum.distrib) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * ((pf1 x (Suc n) w) + (pf2 x (Suc n) w)))" by (simp add: distrib_left) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * ((qty_sum pf1 pf2) x (Suc n) w))" by (simp add: qty_sum_def) also have "... = (\ x\ (support_set (qty_sum pf1 pf2)). ((prices Mkt) x (Suc n) w) * ((qty_sum pf1 pf2) x (Suc n) w))" using sum_support_set[of pf1 pf2] subset_cls_val_process[of "support_set pf1\ support_set pf2" "qty_sum pf1 pf2"] pf fin unfolding cls_val_process_def by simp also have "... = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" by (metis (no_types, lifting) pf sum.cong up_cl_proc.simps(2) cls_val_process_def) finally have "(cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" . thus "cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w" .. qed lemma sum_cls_val_process0: assumes "portfolio pf1" and "portfolio pf2" shows "\w. cls_val_process Mkt (qty_sum pf1 pf2) 0 w = (cls_val_process Mkt pf1) 0 w + (cls_val_process Mkt pf2) 0 w" unfolding cls_val_process_def by (simp add: sum_val_process assms(1) assms(2) sum_portfolio) lemma sum_cls_val_process: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_sum pf1 pf2) n w = (cls_val_process Mkt pf1) n w + (cls_val_process Mkt pf2) n w" by (metis (no_types, lifting) assms(1) assms(2) sum_cls_val_process0 sum_cls_val_process_Suc up_cl_proc.elims) lemma mult_comp_cls_val_process0: assumes "portfolio pf1" shows "\w. cls_val_process Mkt (qty_mult_comp pf1 qty) 0 w = ((cls_val_process Mkt pf1) 0 w) * (qty (Suc 0) w)" unfolding cls_val_process_def by (simp add: assms mult_comp_portfolio mult_comp_val_process) lemma mult_comp_cls_val_process_Suc: assumes "portfolio pf1" shows "\n w. cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) * (qty (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_mult_comp pf1 qty)" using assms by (simp add:mult_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto have "((cls_val_process Mkt pf1) (Suc n) w) * (qty (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))*(qty (Suc n) w)" unfolding cls_val_process_def using assms by simp also have "... = (\ x\ (support_set pf1). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w) * (qty (Suc n) w)))" using sum_distrib_right by auto also have "... = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" unfolding qty_mult_comp_def by (simp add: mult.commute mult.left_commute) also have "... = (\ x\ (support_set (qty_mult_comp pf1 qty)). ((prices Mkt) x (Suc n) w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" using mult_comp_support_set[of pf1 qty] subset_cls_val_process[of "support_set pf1" "qty_mult_comp pf1 qty"] pf fin up_cl_proc.simps(2) unfolding cls_val_process_def by simp also have "... = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" by (metis (no_types, lifting) pf sum.cong cls_val_process_def up_cl_proc.simps(2)) finally have "(cls_val_process Mkt pf1) (Suc n) w * (qty (Suc n) w) = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" . thus "cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w * (qty (Suc n) w)" .. qed lemma remove_comp_cls_val_process0: assumes "portfolio pf1" shows "\w. cls_val_process Mkt (qty_rem_comp pf1 y) 0 w = ((cls_val_process Mkt pf1) 0 w) - (prices Mkt y 0 w)* (pf1 y (Suc 0) w)" unfolding cls_val_process_def by (simp add: assms remove_comp_portfolio remove_comp_val_process) lemma remove_comp_cls_val_process_Suc: assumes "portfolio pf1" shows "\n w. cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) - (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_rem_comp pf1 y)" using assms by (simp add:remove_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto hence fin2: "finite (support_set pf1 - {y})" by simp have "((cls_val_process Mkt pf1) (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))" unfolding cls_val_process_def using assms by simp also have "... = (\ x\ (support_set pf1 - {y}). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof (cases "y\ support_set pf1") case True thus ?thesis by (simp add: fin sum_diff1) next case False hence "pf1 y (Suc n) w = 0" unfolding support_set_def by simp thus ?thesis by (simp add: fin sum_diff1) qed also have "... = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x (Suc n) w) * ((qty_rem_comp pf1 y) x (Suc n) w)) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof - have "(\ x\ (support_set pf1 - {y}). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))) = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x (Suc n) w) * ((qty_rem_comp pf1 y) x (Suc n) w))" proof (rule sum.cong,simp) fix x assume "x\ support_set pf1 - {y}" show "prices Mkt x (Suc n) w * pf1 x (Suc n) w = prices Mkt x (Suc n) w * qty_rem_comp pf1 y x (Suc n) w" using remove_comp_values by (metis DiffD2 \x \ support_set pf1 - {y}\ singletonI) qed thus ?thesis by simp qed also have "... = (cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" using subset_cls_val_process[of "support_set pf1 - {y}" "qty_rem_comp pf1 y"] fin2 by (simp add: remove_comp_support_set) finally have "(cls_val_process Mkt pf1) (Suc n) w = (cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" . thus "cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) - (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" by simp qed lemma replace_comp_cls_val_process0: assumes "\w. prices Mkt x 0 w = cls_val_process Mkt pf2 0 w" and "portfolio pf1" and "portfolio pf2" shows "\w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt pf1 0 w" proof fix w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt (qty_rem_comp pf1 x) 0 w + cls_val_process Mkt (qty_mult_comp pf2 (pf1 x)) 0 w" unfolding qty_replace_comp_def using assms sum_cls_val_process0[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = cls_val_process Mkt pf1 0 w - (prices Mkt x 0 w * pf1 x (Suc 0) w) + cls_val_process Mkt pf2 0 w * pf1 x (Suc 0) w" by (simp add: assms(2) assms(3) mult_comp_cls_val_process0 remove_comp_cls_val_process0) also have "... = cls_val_process Mkt pf1 0 w" using assms by simp finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt pf1 0 w" . qed lemma replace_comp_cls_val_process_Suc: assumes "\n w. prices Mkt x (Suc n) w = cls_val_process Mkt pf2 (Suc n) w" and "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" proof (intro allI) fix n w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt (qty_rem_comp pf1 x) (Suc n) w + cls_val_process Mkt (qty_mult_comp pf2 (pf1 x)) (Suc n) w" unfolding qty_replace_comp_def using assms sum_cls_val_process_Suc[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = cls_val_process Mkt pf1 (Suc n) w - (prices Mkt x (Suc n) w * pf1 x (Suc n) w) + cls_val_process Mkt pf2 (Suc n) w * pf1 x (Suc n) w" by (simp add: assms(2) assms(3) mult_comp_cls_val_process_Suc remove_comp_cls_val_process_Suc) also have "... = cls_val_process Mkt pf1 (Suc n) w" using assms by simp finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" . qed lemma replace_comp_cls_val_process: assumes "\n w. prices Mkt x n w = cls_val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) n w = cls_val_process Mkt pf1 n w" by (metis (no_types, lifting) assms replace_comp_cls_val_process0 replace_comp_cls_val_process_Suc up_cl_proc.elims) lemma qty_single_updated: shows "cls_val_process Mkt (qty_single asset qty) (Suc n) w = prices Mkt asset (Suc n) w * qty (Suc n) w" proof - have "cls_val_process Mkt (qty_single asset qty) (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * ((qty_single asset qty) x (Suc n) w)) {asset})" proof (rule subset_cls_val_process') show "finite {asset}" by simp show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) qed also have "... = prices Mkt asset (Suc n) w * qty (Suc n) w" unfolding qty_single_def by simp finally show ?thesis . qed paragraph \Self-financing\ definition self_financing where "self_financing Mkt p \ (\n. val_process Mkt p (Suc n) = cls_val_process Mkt p (Suc n))" lemma self_financingE: assumes "self_financing Mkt p" shows "\n. val_process Mkt p n = cls_val_process Mkt p n" proof fix n show "val_process Mkt p n = cls_val_process Mkt p n" proof (cases "n = 0") case False thus ?thesis using assms unfolding self_financing_def by (metis up_cl_proc.elims) next case True thus ?thesis by (simp add: cls_val_process_def val_process_def) qed qed lemma static_portfolio_self_financing: assumes "\ x \ support_set p. (\w i. p x i w = p x (Suc i) w)" shows "self_financing Mkt p" unfolding self_financing_def proof (intro allI impI) fix n show "val_process Mkt p (Suc n) = cls_val_process Mkt p (Suc n)" proof (cases "portfolio p") case False thus ?thesis unfolding val_process_def cls_val_process_def by simp next case True have "\w. (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = cls_val_process Mkt p (Suc n) w" proof fix w show "(\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = cls_val_process Mkt p (Suc n) w" proof - have "(\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc n) w)" proof (rule sum.cong, simp) fix x assume "x\ support_set p" hence "p x (Suc n) w = p x (Suc (Suc n)) w" using assms by blast thus "prices Mkt x (Suc n) w * p x (Suc (Suc n)) w = prices Mkt x (Suc n) w * p x (Suc n) w" by simp qed also have "... = cls_val_process Mkt p (Suc n) w" using up_cl_proc.simps(2)[of Mkt p n] by (metis True cls_val_process_def) finally show ?thesis . qed qed moreover have "\w. val_process Mkt p (Suc n) w = (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w)" unfolding val_process_def using True by simp ultimately show ?thesis by auto qed qed lemma sum_self_financing: assumes "portfolio pf1" and "portfolio pf2" and "self_financing Mkt pf1" and "self_financing Mkt pf2" shows "self_financing Mkt (qty_sum pf1 pf2)" proof - have "\ n w. val_process Mkt (qty_sum pf1 pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" proof (intro allI) fix n w have "val_process Mkt (qty_sum pf1 pf2) (Suc n) w = val_process Mkt pf1 (Suc n) w + val_process Mkt pf2 (Suc n) w" using assms by (simp add:sum_val_process) also have "... = cls_val_process Mkt pf1 (Suc n) w + val_process Mkt pf2 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt pf1 (Suc n) w + cls_val_process Mkt pf2 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" using assms by (simp add: sum_cls_val_process) finally show "val_process Mkt (qty_sum pf1 pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed lemma mult_time_constant_self_financing: assumes "portfolio pf1" and "self_financing Mkt pf1" and "\n w. qty n w = qty (Suc n) w" shows "self_financing Mkt (qty_mult_comp pf1 qty)" proof - have "\ n w. val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" proof (intro allI) fix n w have "val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = val_process Mkt pf1 (Suc n) w * qty (Suc n) w" using assms by (simp add:mult_comp_val_process) also have "... = cls_val_process Mkt pf1 (Suc n) w * qty (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" using assms by (auto simp add: mult_comp_cls_val_process_Suc) finally show "val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed lemma replace_comp_self_financing: assumes "\n w. prices Mkt x n w = cls_val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" and "self_financing Mkt pf1" and "self_financing Mkt pf2" shows "self_financing Mkt (qty_replace_comp pf1 x pf2)" proof - have sfeq: "\n w. prices Mkt x n w = val_process Mkt pf2 n w" using assms by (simp add: self_financingE) have "\ n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" proof (intro allI) fix n w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" using assms by (simp add:replace_comp_cls_val_process) also have "... = val_process Mkt pf1 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" using assms sfeq by (simp add: replace_comp_val_process self_financing_def) finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed paragraph \Make a portfolio self-financing\ fun remaining_qty where init: "remaining_qty Mkt v pf asset 0 = (\w. 0)" | first: "remaining_qty Mkt v pf asset (Suc 0) = (\w. (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" | step: "remaining_qty Mkt v pf asset (Suc (Suc n)) = (\w. (remaining_qty Mkt v pf asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" lemma (in disc_equity_market) remaining_qty_predict': assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F n)" proof (induct n) case 0 have "(\w. (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))\ borel_measurable (F 0)" proof (rule borel_measurable_divide) have "val_process Mkt pf 0 \ borel_measurable (F 0)" using assms ats_val_process_adapted by (simp add:adapt_stoch_proc_def) thus "(\x. v - val_process Mkt pf 0 x) \ borel_measurable (F 0)" by simp show "prices Mkt asset 0 \ borel_measurable (F 0)" using assms unfolding adapt_stoch_proc_def by simp qed thus ?case by simp next case (Suc n) have "(\w. (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/ (prices Mkt asset (Suc n) w)) \ borel_measurable (F (Suc n))" proof (rule borel_measurable_divide) show "(\w. (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" proof (rule borel_measurable_diff) show "(\w. (cls_val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" using assms cls_val_process_adapted unfolding adapt_stoch_proc_def by auto show "(\w. (val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" using assms ats_val_process_adapted by (simp add:adapt_stoch_proc_def) qed show "prices Mkt asset (Suc n) \ borel_measurable (F (Suc n))" using assms unfolding adapt_stoch_proc_def by simp qed moreover have "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F (Suc n))" using Suc Suc_n_not_le_n increasing_measurable_info nat_le_linear by blast ultimately show ?case using Suc remaining_qty.simps(3)[of Mkt v pf asset n] by simp qed lemma (in disc_equity_market) remaining_qty_predict: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "borel_predict_stoch_proc F (remaining_qty Mkt v pf asset)" unfolding predict_stoch_proc_def proof (intro conjI allI) show "remaining_qty Mkt v pf asset 0 \ borel_measurable (F 0)" using init by simp fix n show "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F n)" using assms by (simp add: remaining_qty_predict') qed lemma (in disc_equity_market) remaining_qty_adapt: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "remaining_qty Mkt v pf asset n \ borel_measurable (F n)" using adapt_stoch_proc_def assms(1) assms(2) predict_imp_adapt remaining_qty_predict by (metis assms(3)) lemma (in disc_equity_market) remaining_qty_adapted: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "borel_adapt_stoch_proc F (remaining_qty Mkt v pf asset)" using assms unfolding adapt_stoch_proc_def using assms remaining_qty_adapt by blast definition self_finance where "self_finance Mkt v pf (asset::'a) = qty_sum pf (qty_single asset (remaining_qty Mkt v pf asset))" lemma self_finance_portfolio: assumes "portfolio pf" shows "portfolio (self_finance Mkt v pf asset)" unfolding self_finance_def by (simp add: assms single_comp_portfolio sum_portfolio) lemma self_finance_init: assumes "\w. prices Mkt asset 0 w \ 0" and "portfolio pf" shows "val_process Mkt (self_finance Mkt v pf asset) 0 w = v" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "val_process Mkt (self_finance Mkt v pf asset) 0 w = val_process Mkt pf 0 w + val_process Mkt scp 0 w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = val_process Mkt pf 0 w + (sum (\x. ((prices Mkt) x 0 w) * (scp x (Suc 0) w)) {asset})" using subset_val_process'[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * scp asset (Suc 0) w" by auto also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * (remaining_qty Mkt v pf asset) (Suc 0) w" unfolding scp_def qty_single_def by simp also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w)" by simp also have "... = val_process Mkt pf 0 w + (v - val_process Mkt pf 0 w)" using assms by simp also have "... = v" by simp finally show ?thesis . qed lemma self_finance_succ: assumes "prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + cls_val_process Mkt pf (Suc n) w" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = val_process Mkt pf (Suc n) w + val_process Mkt scp (Suc n) w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = val_process Mkt pf (Suc n) w + (sum (\x. ((prices Mkt) x (Suc n) w) * (scp x (Suc (Suc n)) w)) {asset})" using subset_val_process'[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * scp asset (Suc (Suc n)) w" by auto also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc (Suc n)) w" unfolding scp_def qty_single_def by simp also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset (Suc n) w + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" by simp also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + prices Mkt asset (Suc n) w * (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w)" by (simp add: distrib_left) also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)" using assms by simp also have "... = prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + cls_val_process Mkt pf (Suc n) w" by simp finally show ?thesis . qed lemma self_finance_updated: assumes "prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc n) w" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + cls_val_process Mkt scp (Suc n) w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_cls_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = cls_val_process Mkt pf (Suc n) w + (sum (\x. ((prices Mkt) x (Suc n) w) * (scp x (Suc n) w)) {asset})" using subset_cls_val_process[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * scp asset (Suc n) w" by auto also have "... = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc n) w" unfolding scp_def qty_single_def by simp finally show ?thesis . qed lemma self_finance_charact: assumes "\ n w. prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "self_financing Mkt (self_finance Mkt v pf asset)" proof- have "\n w. val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w" proof (intro allI) fix n w show "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w" using assms self_finance_succ[of Mkt asset] by (simp add: self_finance_updated) qed thus ?thesis unfolding self_financing_def by auto qed subsubsection \Replicating portfolios\ definition (in disc_filtr_prob_space) price_structure::"('a \ real) \ nat \ real \ (nat \ 'a \ real) \ bool" where "price_structure pyf T \ pr \ ((\ w\ space M. pr 0 w = \) \ (AE w in M. pr T w = pyf w) \ (pr T \ borel_measurable (F T)))" lemma (in disc_filtr_prob_space) price_structure_init: assumes "price_structure pyf T \ pr" shows "\ w\ space M. pr 0 w = \" using assms unfolding price_structure_def by simp lemma (in disc_filtr_prob_space) price_structure_borel_measurable: assumes "price_structure pyf T \ pr" shows "pr T \ borel_measurable (F T)" using assms unfolding price_structure_def by simp lemma (in disc_filtr_prob_space) price_structure_maturity: assumes "price_structure pyf T \ pr" shows "AE w in M. pr T w = pyf w" using assms unfolding price_structure_def by simp definition (in disc_equity_market) replicating_portfolio where "replicating_portfolio pf der matur \ (stock_portfolio Mkt pf) \ (trading_strategy pf) \ (self_financing Mkt pf) \ (AE w in M. cls_val_process Mkt pf matur w = der w)" definition (in disc_equity_market) is_attainable where "is_attainable der matur \ (\ pf. replicating_portfolio pf der matur)" lemma (in disc_equity_market) replicating_price_process: assumes "replicating_portfolio pf der matur" and "support_adapt Mkt pf" shows "price_structure der matur (initial_value pf) (cls_val_process Mkt pf)" unfolding price_structure_def proof (intro conjI) show "AE w in M. cls_val_process Mkt pf matur w = der w" using assms unfolding replicating_portfolio_def by simp show "\w\space M. cls_val_process Mkt pf 0 w = initial_value pf" proof fix w assume "w\ space M" thus "cls_val_process Mkt pf 0 w = initial_value pf" unfolding initial_value_def using constant_imageI[of "cls_val_process Mkt pf 0"] trading_strategy_init[of pf] assms replicating_portfolio_def [of pf der matur] by (simp add: stock_portfolio_def cls_val_process_def) qed show "cls_val_process Mkt pf matur \ borel_measurable (F matur)" using assms unfolding replicating_portfolio_def using ats_val_process_adapted[of pf] cls_val_process_adapted by (simp add:adapt_stoch_proc_def) qed subsubsection \Arbitrages\ definition (in disc_filtr_prob_space) arbitrage_process where "arbitrage_process Mkt p \ (\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" lemma (in disc_filtr_prob_space) arbitrage_processE: assumes "arbitrage_process Mkt p" shows "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" using assms disc_filtr_prob_space.arbitrage_process_def disc_filtr_prob_space_axioms self_financingE by fastforce lemma (in disc_filtr_prob_space) arbitrage_processI: assumes "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" shows "arbitrage_process Mkt p" unfolding arbitrage_process_def using assms by (simp add: self_financingE cls_val_process_def) definition (in disc_filtr_prob_space) viable_market where "viable_market Mkt \ (\p. stock_portfolio Mkt p \ \ arbitrage_process Mkt p)" lemma (in disc_filtr_prob_space) arbitrage_val_process: assumes "arbitrage_process Mkt pf1" and "self_financing Mkt pf2" and "trading_strategy pf2" and "\ n w. cls_val_process Mkt pf1 n w = cls_val_process Mkt pf2 n w" shows "arbitrage_process Mkt pf2" proof - have "(\ m. (self_financing Mkt pf1) \ (trading_strategy pf1) \ (\w \ space M. cls_val_process Mkt pf1 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt pf1 m w) \ 0 < \

(w in M. cls_val_process Mkt pf1 m w > 0))" using assms arbitrage_processE[of Mkt pf1] by simp from this obtain m where "(self_financing Mkt pf1)" and "(trading_strategy pf1)" and "(\w \ space M. cls_val_process Mkt pf1 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt pf1 m w)" "0 < \

(w in M. cls_val_process Mkt pf1 m w > 0)" by auto have ae_eq:"\w\ space M. (cls_val_process Mkt pf1 0 w = 0) = (cls_val_process Mkt pf2 0 w = 0)" proof fix w assume "w\ space M" show "(cls_val_process Mkt pf1 0 w = 0) = (cls_val_process Mkt pf2 0 w = 0) " using assms by simp qed have ae_geq:"almost_everywhere M (\w. cls_val_process Mkt pf1 m w \ 0) = almost_everywhere M (\w. cls_val_process Mkt pf2 m w \ 0)" proof (rule AE_cong) fix w assume "w\ space M" show "(cls_val_process Mkt pf1 m w \ 0) = (cls_val_process Mkt pf2 m w \ 0) " using assms by simp qed have "self_financing Mkt pf2" using assms by simp moreover have "trading_strategy pf2" using assms by simp moreover have "(\w \ space M. cls_val_process Mkt pf2 0 w = 0)" using \(\w \ space M. cls_val_process Mkt pf1 0 w = 0)\ ae_eq by simp moreover have "AE w in M. 0 \ cls_val_process Mkt pf2 m w" using \AE w in M. 0 \ cls_val_process Mkt pf1 m w\ ae_geq by simp moreover have "0 < prob {w \ space M. 0 < cls_val_process Mkt pf2 m w}" proof - have "{w \ space M. 0 < cls_val_process Mkt pf2 m w} = {w \ space M. 0 < cls_val_process Mkt pf1 m w}" by (simp add: assms(4)) thus ?thesis by (simp add: \0 < prob {w \ space M. 0 < cls_val_process Mkt pf1 m w}\) qed ultimately show ?thesis using arbitrage_processI by blast qed definition coincides_on where "coincides_on Mkt Mkt2 A \ (stocks Mkt = stocks Mkt2 \ (\x. x\ A \ prices Mkt x = prices Mkt2 x))" lemma coincides_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. val_process Mkt pf n w = val_process Mkt2 pf n w" proof (intro allI) fix n w show "val_process Mkt pf n w = val_process Mkt2 pf n w" proof (cases "portfolio pf") case False thus ?thesis unfolding val_process_def by simp next case True hence "val_process Mkt pf n w = (\x\ support_set pf. prices Mkt x n w * pf x (Suc n) w)" using assms unfolding val_process_def by simp also have "... = (\x\ support_set pf. prices Mkt2 x n w * pf x (Suc n) w)" proof (rule sum.cong, simp) fix y assume "y\ support_set pf" hence "y\ A" using assms unfolding stock_portfolio_def by auto hence "prices Mkt y n w = prices Mkt2 y n w" using assms unfolding coincides_on_def by auto thus "prices Mkt y n w * pf y (Suc n) w = prices Mkt2 y n w * pf y (Suc n) w" by simp qed also have "... = val_process Mkt2 pf n w" by (metis (mono_tags, lifting) calculation val_process_def) finally show "val_process Mkt pf n w = val_process Mkt2 pf n w" . qed qed lemma coincides_cls_val_process': assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" proof (intro allI) fix n w show "cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" proof (cases "portfolio pf") case False thus ?thesis unfolding cls_val_process_def by simp next case True hence "cls_val_process Mkt pf (Suc n) w = (\x\ support_set pf. prices Mkt x (Suc n) w * pf x (Suc n) w)" using assms unfolding cls_val_process_def by simp also have "... = (\x\ support_set pf. prices Mkt2 x (Suc n) w * pf x (Suc n) w)" proof (rule sum.cong, simp) fix y assume "y\ support_set pf" hence "y\ A" using assms unfolding stock_portfolio_def by auto hence "prices Mkt y (Suc n) w = prices Mkt2 y (Suc n) w" using assms unfolding coincides_on_def by auto thus "prices Mkt y (Suc n) w * pf y (Suc n) w = prices Mkt2 y (Suc n) w * pf y (Suc n) w" by simp qed also have "... = cls_val_process Mkt2 pf (Suc n) w" by (metis (mono_tags, lifting) True up_cl_proc.simps(2) cls_val_process_def) finally show "cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" . qed qed lemma coincides_cls_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (intro allI) fix n w show "cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (cases "portfolio pf") case False thus ?thesis unfolding cls_val_process_def by simp next case True show "cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (induct n) case 0 with assms show ?case by (simp add: cls_val_process_def coincides_val_process) next case Suc thus ?case using coincides_cls_val_process' assms by blast qed qed qed lemma (in disc_filtr_prob_space) coincides_on_self_financing: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "self_financing Mkt p" shows "self_financing Mkt2 p" proof - have "\ n w. val_process Mkt2 p (Suc n) w = cls_val_process Mkt2 p (Suc n) w" proof (intro allI) fix n w have "val_process Mkt2 p (Suc n) w = val_process Mkt p (Suc n) w" using assms(1) assms(2) coincides_val_process stock_portfolio_def by fastforce also have "... = cls_val_process Mkt p (Suc n) w" by (metis \self_financing Mkt p\ self_financing_def) also have "... = cls_val_process Mkt2 p (Suc n) w" using assms(1) assms(2) coincides_cls_val_process stock_portfolio_def by blast finally show "val_process Mkt2 p (Suc n) w = cls_val_process Mkt2 p (Suc n) w" . qed thus "self_financing Mkt2 p" unfolding self_financing_def by auto qed lemma (in disc_filtr_prob_space) coincides_on_arbitrage: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "arbitrage_process Mkt p" shows "arbitrage_process Mkt2 p" proof - have "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w\ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" using assms using arbitrage_processE by simp from this obtain m where "(self_financing Mkt p)" and "(trading_strategy p)" and "(\w\ space M. cls_val_process Mkt p 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt p m w)" "0 < \

(w in M. cls_val_process Mkt p m w > 0)" by auto have ae_eq:"\w\ space M. (cls_val_process Mkt2 p 0 w = 0) = (cls_val_process Mkt p 0 w = 0)" proof fix w assume "w\ space M" show "(cls_val_process Mkt2 p 0 w = 0) = (cls_val_process Mkt p 0 w = 0) " using assms coincides_cls_val_process by (metis) qed have ae_geq:"almost_everywhere M (\w. cls_val_process Mkt2 p m w \ 0) = almost_everywhere M (\w. cls_val_process Mkt p m w \ 0)" proof (rule AE_cong) fix w assume "w\ space M" show "(cls_val_process Mkt2 p m w \ 0) = (cls_val_process Mkt p m w \ 0) " using assms coincides_cls_val_process by (metis) qed have "self_financing Mkt2 p" using assms coincides_on_self_financing using \self_financing Mkt p\ by blast moreover have "trading_strategy p" using \trading_strategy p\ . moreover have "(\w\ space M. cls_val_process Mkt2 p 0 w = 0)" using \(\w\ space M. cls_val_process Mkt p 0 w = 0)\ ae_eq by simp moreover have "AE w in M. 0 \ cls_val_process Mkt2 p m w" using \AE w in M. 0 \ cls_val_process Mkt p m w\ ae_geq by simp moreover have "0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" proof - have "{w \ space M. 0 < cls_val_process Mkt2 p m w} = {w \ space M. 0 < cls_val_process Mkt p m w}" by (metis (no_types, lifting) assms(1) assms(2) coincides_cls_val_process) thus ?thesis by (simp add: \0 < prob {w \ space M. 0 < cls_val_process Mkt p m w}\) qed ultimately show ?thesis using arbitrage_processI by blast qed lemma (in disc_filtr_prob_space) coincides_on_stocks_viable: assumes "coincides_on Mkt Mkt2 (stocks Mkt)" and "viable_market Mkt" shows "viable_market Mkt2" using coincides_on_arbitrage by (metis (mono_tags, opaque_lifting) assms(1) assms(2) coincides_on_def stock_portfolio_def viable_market_def) lemma coincides_stocks_val_process: assumes "stock_portfolio Mkt pf" and "coincides_on Mkt Mkt2 (stocks Mkt)" shows "\n w. val_process Mkt pf n w = val_process Mkt2 pf n w" using assms unfolding stock_portfolio_def by (simp add: coincides_val_process) lemma coincides_stocks_cls_val_process: assumes "stock_portfolio Mkt pf" and "coincides_on Mkt Mkt2 (stocks Mkt)" shows "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using assms unfolding stock_portfolio_def by (simp add: coincides_cls_val_process) lemma (in disc_filtr_prob_space) coincides_on_adapted_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "borel_adapt_stoch_proc F (val_process Mkt p)" shows "borel_adapt_stoch_proc F (val_process Mkt2 p)" unfolding adapt_stoch_proc_def proof fix n have "val_process Mkt p n \ borel_measurable (F n)" using assms unfolding adapt_stoch_proc_def by simp moreover have "\w. val_process Mkt p n w = val_process Mkt2 p n w" using assms coincides_val_process[of Mkt Mkt2 A] by auto thus "val_process Mkt2 p n \ borel_measurable (F n)" using calculation by presburger qed lemma (in disc_filtr_prob_space) coincides_on_adapted_cls_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "borel_adapt_stoch_proc F (cls_val_process Mkt p)" shows "borel_adapt_stoch_proc F (cls_val_process Mkt2 p)" unfolding adapt_stoch_proc_def proof fix n have "cls_val_process Mkt p n \ borel_measurable (F n)" using assms unfolding adapt_stoch_proc_def by simp moreover have "\w. cls_val_process Mkt p n w = cls_val_process Mkt2 p n w" using assms coincides_cls_val_process[of Mkt Mkt2 A] by auto thus "cls_val_process Mkt2 p n \ borel_measurable (F n)" using calculation by presburger qed subsubsection \Fair prices\ definition (in disc_filtr_prob_space) fair_price where "fair_price Mkt \ pyf matur \ (\ pr. price_structure pyf matur \ pr \ (\ x Mkt2 p. (x\ stocks Mkt \ ((coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" lemma (in disc_filtr_prob_space) fair_priceI: assumes "fair_price Mkt \ pyf matur" shows "(\ pr. price_structure pyf matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using assms unfolding fair_price_def by simp paragraph \Existence when replicating portfolio\ lemma (in disc_equity_market) replicating_fair_price: assumes "viable_market Mkt" and "replicating_portfolio pf der matur" and "support_adapt Mkt pf" shows "fair_price Mkt (initial_value pf) der matur" proof (rule ccontr) define \ where "\ = (initial_value pf)" assume "\ fair_price Mkt \ der matur" hence imps: "\pr. (price_structure der matur \ pr) \ (\ x Mkt2 p. (x\ stocks Mkt \ (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p))" unfolding fair_price_def by simp have "(price_structure der matur \ (cls_val_process Mkt pf))" unfolding \_def using replicating_price_process assms by simp hence "\ x Mkt2 p. (x\ stocks Mkt \ (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = (cls_val_process Mkt pf)) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p)" using imps by simp from this obtain x Mkt2 p where "x\ stocks Mkt" and "coincides_on Mkt Mkt2 (stocks Mkt)" and "prices Mkt2 x = cls_val_process Mkt pf" and "portfolio p" and "support_set p\ stocks Mkt \ {x}" and "arbitrage_process Mkt2 p" by auto have "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using coincides_stocks_cls_val_process[of Mkt pf Mkt2] assms \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by simp have "\m. self_financing Mkt2 p \ trading_strategy p \ (AE w in M. cls_val_process Mkt2 p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt2 p m w) \ 0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" using \arbitrage_process Mkt2 p\ using arbitrage_processE[of Mkt2] by simp from this obtain m where "self_financing Mkt2 p" "trading_strategy p \ (AE w in M. cls_val_process Mkt2 p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt2 p m w) \ 0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" by auto note mprop = this define eq_stock where "eq_stock = qty_replace_comp p x pf" have "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using assms unfolding replicating_portfolio_def using coincides_cls_val_process using \\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w\ by blast hence prx: "\n w. prices Mkt2 x n w = cls_val_process Mkt2 pf n w" using \prices Mkt2 x = cls_val_process Mkt pf\ by simp have "stock_portfolio Mkt2 eq_stock" by (metis (no_types, lifting) \coincides_on Mkt Mkt2 (stocks Mkt)\ \portfolio p\ \support_set p \ stocks Mkt \ {x}\ assms(2) coincides_on_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms eq_stock_def replace_comp_portfolio replace_comp_stocks stock_portfolio_def) moreover have "arbitrage_process Mkt2 eq_stock" proof (rule arbitrage_val_process) show "arbitrage_process Mkt2 p" using \arbitrage_process Mkt2 p\ . show vp: "\n w. cls_val_process Mkt2 p n w = cls_val_process Mkt2 eq_stock n w" using replace_comp_cls_val_process \prices Mkt2 x = cls_val_process Mkt pf\ unfolding eq_stock_def by (metis (no_types, lifting) \\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w\ \portfolio p\ assms(2) replicating_portfolio_def stock_portfolio_def) show "trading_strategy eq_stock" by (metis \arbitrage_process Mkt2 p\ arbitrage_process_def assms(2) eq_stock_def replace_comp_trading_strat replicating_portfolio_def) show "self_financing Mkt2 eq_stock" unfolding eq_stock_def proof (rule replace_comp_self_financing) show "portfolio pf" using assms unfolding replicating_portfolio_def stock_portfolio_def by simp show "portfolio p" using \portfolio p\ . show "\n w. prices Mkt2 x n w = cls_val_process Mkt2 pf n w" using prx . show "self_financing Mkt2 p" using \self_financing Mkt2 p\ . show "self_financing Mkt2 pf" using coincides_on_self_financing[of Mkt Mkt2 "stocks Mkt" pf] \coincides_on Mkt Mkt2 (stocks Mkt)\ assms(2) (*disc_equity_market.replicating_portfolio_def disc_equity_market_axioms*) unfolding stock_portfolio_def replicating_portfolio_def by auto qed qed moreover have "viable_market Mkt2" using assms coincides_on_stocks_viable[of Mkt Mkt2] by (simp add: \coincides_on Mkt Mkt2 (stocks Mkt)\) ultimately show False unfolding viable_market_def by simp qed paragraph \Uniqueness when replicating portfolio\ text \The proof of uniqueness requires the existence of a stock that always takes strictly positive values.\ locale disc_market_pos_stock = disc_equity_market + fixes pos_stock assumes in_stock: "pos_stock \ stocks Mkt" and positive: "\ n w. prices Mkt pos_stock n w > 0" and readable: "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" lemma (in disc_market_pos_stock) pos_stock_borel_adapted: shows "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using assets_def readable in_stock by auto definition static_quantities where "static_quantities p \ (\asset \ support_set p. \c::real. p asset = (\ n w. c))" lemma (in disc_filtr_prob_space) static_quantities_trading_strat: assumes "static_quantities p" and "finite (support_set p)" shows "trading_strategy p" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio p" using assms unfolding portfolio_def by simp fix asset assume "asset \ support_set p" hence "\c. p asset = (\ n w. c)" using assms unfolding static_quantities_def by simp then obtain c where "p asset = (\ n w. c)" by auto show "borel_predict_stoch_proc F (p asset)" unfolding predict_stoch_proc_def proof (intro conjI) show "p asset 0 \ borel_measurable (F 0)" using \p asset = (\ n w. c)\ by simp show "\n. p asset (Suc n) \ borel_measurable (F n)" proof fix n have "p asset (Suc n) = (\ w. c)" using \p asset = (\ n w. c)\ by simp thus "p asset (Suc n) \ borel_measurable (F n)" by simp qed qed qed lemma two_component_support_set: assumes "\ n w. a n w \ 0" and "\ n w. b n w\ 0" and "x \ y" shows "support_set ((\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b)) = {x,y}" proof let ?arb_pf = "(\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b)" have "\ n w. ?arb_pf x n w \ 0" using assms by simp moreover have "\n w. ?arb_pf y n w \ 0" using assms by simp ultimately show "{x, y} \ support_set ?arb_pf" unfolding support_set_def by simp show "support_set ?arb_pf \ {x, y}" proof (rule ccontr) assume "\ support_set ?arb_pf \ {x, y}" hence "\z. z\ support_set ?arb_pf \ z\ {x, y}" by auto from this obtain z where "z\ support_set ?arb_pf" and "z\ {x, y}" by auto have "\n w. ?arb_pf z n w \ 0" using \z\ support_set ?arb_pf\ unfolding support_set_def by simp from this obtain n w where "?arb_pf z n w \ 0" by auto have "?arb_pf z n w = 0" using \z\ {x, y}\ by simp thus False using \?arb_pf z n w \ 0\ by simp qed qed lemma two_component_val_process: assumes "arb_pf = ((\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b))" and "portfolio arb_pf" and "x \ y" and "\ n w. a n w \ 0" and "\ n w. b n w\ 0" shows "val_process Mkt arb_pf n w = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" proof - have "support_set arb_pf = {x,y}" using assms by (simp add:two_component_support_set) have "val_process Mkt arb_pf n w = (\x\support_set arb_pf. prices Mkt x n w * arb_pf x (Suc n) w)" unfolding val_process_def using \portfolio arb_pf\ by simp also have "... = (\x\{x, y}. prices Mkt x n w * arb_pf x (Suc n) w)" using \support_set arb_pf = {x, y}\ by simp also have "... = (\x\{y}. prices Mkt x n w * arb_pf x (Suc n) w) + prices Mkt x n w * arb_pf x (Suc n) w" using sum.insert[of "{y}" x "\x. prices Mkt x n w * arb_pf x n w"] assms(3) by auto also have "... = prices Mkt y n w * arb_pf y (Suc n) w + prices Mkt x n w * arb_pf x (Suc n) w" by simp also have "... = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" using assms by auto finally show "val_process Mkt arb_pf n w = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" . qed lemma quantity_update_support_set: assumes "\n w. pr n w \ 0" and "x\ support_set p" shows "support_set (p(x:=pr)) = support_set p \ {x}" proof show "support_set (p(x := pr)) \ support_set p \ {x}" proof fix y assume "y\ support_set (p(x := pr))" show "y \ support_set p \ {x}" proof (rule ccontr) assume "\y \ support_set p \ {x}" hence "y \ x" by simp have "\n w. (p(x := pr)) y n w \ 0" using \y\ support_set (p(x := pr))\ unfolding support_set_def by simp then obtain n w where nwprop: "(p(x := pr)) y n w \ 0" by auto have "y\ support_set p" using \\y \ support_set p \ {x}\ by simp hence "y = x" using nwprop using support_set_def by force thus False using \y\ x\ by simp qed qed show "support_set p \ {x} \ support_set (p(x := pr))" proof fix y assume "y \ support_set p \ {x}" show "y\ support_set (p(x := pr))" proof (cases "y\ support_set p") case True thus ?thesis proof - have f1: "y \ {b. \n a. p b n a \ 0}" by (metis True support_set_def) then have "y \ x" using assms(2) support_set_def by force then show ?thesis using f1 by (simp add: support_set_def) qed next case False hence "y = x" using \y \ support_set p \ {x}\ by auto thus ?thesis using assms by (simp add: support_set_def) qed qed qed lemma fix_asset_price: shows "\x Mkt2. x \ stocks Mkt \ coincides_on Mkt Mkt2 (stocks Mkt) \ prices Mkt2 x = pr" proof - have "\x. x\ stocks Mkt" by (metis UNIV_eq_I stk_strict_subs_def mkt_stocks_assets) from this obtain x where "x\ stocks Mkt" by auto let ?res = "discrete_market_of (stocks Mkt) ((prices Mkt)(x:=pr))" have "coincides_on Mkt ?res (stocks Mkt)" proof - have "stocks Mkt = stocks (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr)))" by (metis (no_types) stk_strict_subs_def mkt_stocks_assets stocks_of) then show ?thesis by (simp add: \x \ stocks Mkt\ coincides_on_def prices_of) qed have "prices ?res x = pr" by (simp add: prices_of) show ?thesis using \coincides_on Mkt (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr))) (stocks Mkt)\ \prices (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr))) x = pr\ \x \ stocks Mkt\ by blast qed lemma (in disc_market_pos_stock) arbitrage_portfolio_properties: assumes "price_structure der matur \ pr" and "replicating_portfolio pf der matur" and "(coincides_on Mkt Mkt2 (stocks Mkt))" and "(prices Mkt2 x = pr)" and "x\ stocks Mkt" and "diff_inv = (\ - initial_value pf) / constant_image (prices Mkt pos_stock 0)" and "diff_inv \ 0" and "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" and "contr_pf = qty_sum arb_pf pf" shows "self_financing Mkt2 contr_pf" and "trading_strategy contr_pf" and "\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0" and "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" and "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" and "support_set arb_pf = {x, pos_stock}" and "portfolio contr_pf" proof - have "0 < constant_image (prices Mkt pos_stock 0)" using trading_strategy_init proof - have "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using pos_stock_borel_adapted by simp hence "\c. \w\space M. prices Mkt pos_stock 0 w = c" using adapted_init[of "prices Mkt pos_stock"] by simp moreover have "\w\ space M. 0 < prices Mkt pos_stock 0 w" using positive by simp ultimately show ?thesis using constant_image_pos by simp qed show "support_set arb_pf = {x, pos_stock}" proof - have "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" using \arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))\ . moreover have "\n w. diff_inv \ 0" using assms by simp moreover have "x\ pos_stock" using \x \ stocks Mkt\ in_stock by auto ultimately show ?thesis by (simp add:two_component_support_set) qed hence "portfolio arb_pf" unfolding portfolio_def by simp have arb_vp:"\n w. val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" proof (intro allI) fix n w have "val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * (\ n w. diff_inv) n w + prices Mkt2 x n w * (\ n w. -1) n w" proof (rule two_component_val_process) show "x\ pos_stock" using \x \ stocks Mkt\ in_stock by auto show "arb_pf = (\x n w. 0)(x := \a b. - 1, pos_stock := \a b. diff_inv)" using assms by simp show "portfolio arb_pf" using \portfolio arb_pf\ by simp show "\n w. - (1::real) \ 0" by simp show "\n w. diff_inv \ 0" using assms by auto qed also have "... = prices Mkt2 pos_stock n w * diff_inv - pr n w" using \prices Mkt2 x = pr\ by simp finally show "val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" . qed have "static_quantities arb_pf" unfolding static_quantities_def proof fix asset assume "asset \ support_set arb_pf" thus "\c. arb_pf asset = (\n w. c)" proof (cases "asset = x") case True thus ?thesis using assms by auto next case False hence "asset = pos_stock" using \support_set arb_pf = {x, pos_stock}\ using \asset \ support_set arb_pf\ by blast thus ?thesis using assms by auto qed qed hence "trading_strategy arb_pf" using \portfolio arb_pf\ portfolio_def static_quantities_trading_strat by blast have "self_financing Mkt2 arb_pf" by (simp add: static_portfolio_self_financing \arb_pf = (\x n w. 0) (x := \n w. -1, pos_stock := \n w. diff_inv)\) hence arb_uvp: "\n w. cls_val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" using assms arb_vp by (simp add:self_financingE) show "portfolio contr_pf" using assms by (metis \support_set arb_pf = {x, pos_stock}\ replicating_portfolio_def finite.emptyI finite.insertI portfolio_def stock_portfolio_def sum_portfolio) have "support_set contr_pf \ stocks Mkt \ {x}" proof - have "support_set contr_pf \ support_set arb_pf \ support_set pf" using assms by (simp add:sum_support_set) moreover have "support_set arb_pf \ stocks Mkt \ {x}" using \support_set arb_pf = {x, pos_stock}\ in_stock by simp moreover have "support_set pf \ stocks Mkt \ {x}" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto ultimately show ?thesis by auto qed show "self_financing Mkt2 contr_pf" proof - have "self_financing Mkt2 (qty_sum arb_pf pf)" proof (rule sum_self_financing) show "portfolio arb_pf" using \support_set arb_pf = {x, pos_stock}\ unfolding portfolio_def by auto show "portfolio pf" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto show "self_financing Mkt2 pf" using coincides_on_self_financing \(coincides_on Mkt Mkt2 (stocks Mkt))\ \(prices Mkt2 x = pr)\ assms(2) unfolding replicating_portfolio_def stock_portfolio_def by blast show "self_financing Mkt2 arb_pf" by (simp add: static_portfolio_self_financing \arb_pf = (\x n w. 0) (x := \n w. -1, pos_stock := \n w. diff_inv)\) qed thus ?thesis using assms by simp qed show "trading_strategy contr_pf" proof - have "trading_strategy (qty_sum arb_pf pf)" proof (rule sum_trading_strat) show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "trading_strategy arb_pf" using \trading_strategy arb_pf\ . qed thus ?thesis using assms by simp qed show "\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt2 contr_pf 0 w = cls_val_process Mkt2 arb_pf 0 w + cls_val_process Mkt2 pf 0 w" using sum_cls_val_process0[of arb_pf pf Mkt2] using \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock 0 w * diff_inv - pr 0 w + cls_val_process Mkt2 pf 0 w" using arb_uvp by simp also have "... = constant_image (prices Mkt pos_stock 0) * diff_inv - pr 0 w + cls_val_process Mkt2 pf 0 w" proof - have f1: "prices Mkt pos_stock = prices Mkt2 pos_stock" using \coincides_on Mkt Mkt2 (stocks Mkt)\ in_stock unfolding coincides_on_def by blast have "prices Mkt pos_stock 0 w = constant_image (prices Mkt pos_stock 0)" using \w \ space M\ adapted_init constant_imageI pos_stock_borel_adapted by blast then show ?thesis using f1 by simp qed also have "... = (\ - initial_value pf) - pr 0 w + cls_val_process Mkt2 pf 0 w" using \0 < constant_image (prices Mkt pos_stock 0)\ assms by simp also have "... = (\ - initial_value pf) - \ + cls_val_process Mkt2 pf 0 w" using \price_structure der matur \ pr\ price_structure_init[of der matur \ pr] by (simp add: \w \ space M\) also have "... = (\ - initial_value pf) - \ + (initial_value pf)" using initial_valueI assms unfolding replicating_portfolio_def using \w \ space M\ coincides_stocks_cls_val_process self_financingE readable by (metis (no_types, opaque_lifting) support_adapt_def stock_portfolio_def subsetCE) also have "... = 0" by simp finally show "cls_val_process Mkt2 contr_pf 0 w = 0" . qed show "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" proof assume "0 < diff_inv" show "AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w" proof (rule AE_mp) have "AE w in M. prices Mkt2 x matur w = der w" using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ unfolding price_structure_def by auto moreover have "AE w in M. cls_val_process Mkt2 pf matur w = der w" using assms coincides_stocks_cls_val_process[of Mkt pf Mkt2] \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by auto ultimately show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" by auto show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w \ 0 < cls_val_process Mkt2 contr_pf matur w" proof (rule AE_I2, rule impI) fix w assume "w\ space M" and "prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" have "cls_val_process Mkt2 contr_pf matur w = cls_val_process Mkt2 arb_pf matur w + cls_val_process Mkt2 pf matur w" using sum_cls_val_process[of arb_pf pf Mkt2] \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock matur w * diff_inv - pr matur w + cls_val_process Mkt2 pf matur w" using arb_uvp by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv - prices Mkt2 x matur w + cls_val_process Mkt2 pf matur w" using \prices Mkt2 x = pr\ by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv" using \prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w\ by simp also have "... > 0" using positive \0 < diff_inv\ by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ coincides_on_def in_stock mult_pos_pos) finally have "cls_val_process Mkt2 contr_pf matur w > 0". thus "0 < cls_val_process Mkt2 contr_pf matur w" by simp qed qed qed show "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" proof assume "diff_inv < 0" show "AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w" proof (rule AE_mp) have "AE w in M. prices Mkt2 x matur w = der w" using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ unfolding price_structure_def by auto moreover have "AE w in M. cls_val_process Mkt2 pf matur w = der w" using assms coincides_stocks_cls_val_process[of Mkt pf Mkt2] \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by auto ultimately show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" by auto show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w \ 0 > cls_val_process Mkt2 contr_pf matur w" proof (rule AE_I2, rule impI) fix w assume "w\ space M" and "prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" have "cls_val_process Mkt2 contr_pf matur w = cls_val_process Mkt2 arb_pf matur w + cls_val_process Mkt2 pf matur w" using sum_cls_val_process[of arb_pf pf Mkt2] \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock matur w * diff_inv - pr matur w + cls_val_process Mkt2 pf matur w" using arb_uvp by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv - prices Mkt2 x matur w + cls_val_process Mkt2 pf matur w" using \prices Mkt2 x = pr\ by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv" using \prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w\ by simp also have "... < 0" using positive \diff_inv < 0\ by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ coincides_on_def in_stock mult_pos_neg) finally have "cls_val_process Mkt2 contr_pf matur w < 0". thus "0 > cls_val_process Mkt2 contr_pf matur w" by simp qed qed qed qed lemma (in disc_equity_market) mult_comp_cls_val_process_measurable': assumes "cls_val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "qty n \ borel_measurable (F n)" and "0 \ n" shows "cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof - have "\m. n = Suc m" using assms by presburger from this obtain m where "n = Suc m" by auto hence "cls_val_process Mkt2 (qty_mult_comp pf qty) (Suc m) \ borel_measurable (F (Suc m))" using mult_comp_cls_val_process_Suc[of pf Mkt2 qty] borel_measurable_times[of "cls_val_process Mkt2 pf (Suc m)" "F (Suc m)" "qty (Suc m)"] assms \n= Suc m\ by presburger thus ?thesis using \n = Suc m\ by simp qed lemma (in disc_equity_market) mult_comp_cls_val_process_measurable: assumes "\n. cls_val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "\n. qty (Suc n) \ borel_measurable (F n)" shows "\n. cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof fix n show "cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof (cases "n=0") case False hence "\m. n = Suc m" by presburger from this obtain m where "n = Suc m" by auto have "qty n \ borel_measurable (F n)" using Suc_n_not_le_n \n = Suc m\ assms(3) increasing_measurable_info nat_le_linear by blast hence "qty (Suc m) \ borel_measurable (F (Suc m))" using \n = Suc m\ by simp hence "cls_val_process Mkt2 (qty_mult_comp pf qty) (Suc m) \ borel_measurable (F (Suc m))" using mult_comp_cls_val_process_Suc[of pf Mkt2 qty] borel_measurable_times[of "cls_val_process Mkt2 pf (Suc m)" "F (Suc m)" "qty (Suc m)"] assms \n= Suc m\ by presburger thus ?thesis using \n = Suc m\ by simp next case True have "qty (Suc 0) \ borel_measurable (F 0)" using assms by simp moreover have "cls_val_process Mkt2 pf 0 \ borel_measurable (F 0)" using assms by simp ultimately have "(\w. cls_val_process Mkt2 pf 0 w * qty (Suc 0) w) \ borel_measurable (F 0)" by simp thus ?thesis using assms(2) True mult_comp_cls_val_process0 by (simp add: \(\w. cls_val_process Mkt2 pf 0 w * qty (Suc 0) w) \ borel_measurable (F 0)\ mult_comp_cls_val_process0 measurable_cong) qed qed lemma (in disc_equity_market) mult_comp_val_process_measurable: assumes "val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "qty (Suc n) \ borel_measurable (F n)" shows "val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" using mult_comp_val_process[of pf Mkt2 qty] borel_measurable_times[of "val_process Mkt2 pf n" "F n" "qty (Suc n)"] assms by presburger lemma (in disc_market_pos_stock) repl_fair_price_unique: assumes "replicating_portfolio pf der matur" and "fair_price Mkt \ der matur" shows "\ = initial_value pf" proof - have expr: "(\ pr. price_structure der matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using assms fair_priceI by simp then obtain pr where "price_structure der matur \ pr" and xasset: "(\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p)))" by auto define diff_inv where "diff_inv = (\ - initial_value pf) / constant_image (prices Mkt pos_stock 0)" { fix x assume "x\ stocks Mkt" hence mkprop: "(\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p)" using xasset by simp fix Mkt2 assume "(coincides_on Mkt Mkt2 (stocks Mkt))" and "(prices Mkt2 x = pr)" have "0 < constant_image (prices Mkt pos_stock 0)" using trading_strategy_init proof - have "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using pos_stock_borel_adapted by simp hence "\c. \w\space M. prices Mkt pos_stock 0 w = c" using adapted_init[of "prices Mkt pos_stock"] by simp moreover have "\w\ space M. 0 < prices Mkt pos_stock 0 w" using positive by simp ultimately show ?thesis using constant_image_pos by simp qed define arb_pf where "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" define contr_pf where "contr_pf = qty_sum arb_pf pf" have 1:"0 \ diff_inv \ self_financing Mkt2 contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 2:"0 \ diff_inv \ trading_strategy contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 3:"0 \ diff_inv \ (\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 4: "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 5: "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 6: "0 \ diff_inv \ support_set arb_pf = {x, pos_stock}" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 7: "0 \ diff_inv \support_set contr_pf \ stocks Mkt \ {x}" proof - have "0 \ diff_inv \ support_set contr_pf \ support_set arb_pf \ support_set pf" unfolding contr_pf_def by (simp add:sum_support_set) moreover have "0 \ diff_inv \support_set arb_pf \ stocks Mkt \ {x}" using \0 \ diff_inv \ support_set arb_pf = {x, pos_stock}\ in_stock by simp moreover have "0 \ diff_inv \support_set pf \ stocks Mkt \ {x}" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto ultimately show ?thesis by auto qed have 8:"0 \ diff_inv \portfolio contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 9: "0 \ diff_inv \ cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" proof assume "0 \ diff_inv" have 10:"\ asset \ support_set arb_pf \ support_set pf. prices Mkt2 asset matur \ borel_measurable (F matur)" proof fix asset assume "asset \ support_set arb_pf \ support_set pf" show "prices Mkt2 asset matur \ borel_measurable (F matur)" proof (cases "asset \ support_set pf") case True thus ?thesis using assms readable by (metis (no_types, lifting) \coincides_on Mkt Mkt2 (stocks Mkt)\ adapt_stoch_proc_def coincides_on_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms stock_portfolio_def subsetCE) next case False hence "asset\ support_set arb_pf" using \asset \ support_set arb_pf \ support_set pf\ by auto show ?thesis proof (cases "asset = x") case True thus ?thesis using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ price_structure_borel_measurable by blast next case False hence "asset = pos_stock" using \asset\ support_set arb_pf\ \0 \ diff_inv \ support_set arb_pf = {x, pos_stock}\ \0 \ diff_inv\ by auto thus ?thesis by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ adapt_stoch_proc_def coincides_on_def in_stock pos_stock_borel_adapted) qed qed qed moreover have "\asset\support_set contr_pf. contr_pf asset matur \ borel_measurable (F matur)" using \0 \ diff_inv \trading_strategy contr_pf\ \0 \ diff_inv\ by (metis adapt_stoch_proc_def disc_filtr_prob_space.predict_imp_adapt disc_filtr_prob_space_axioms trading_strategy_def) ultimately show "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" proof- have "\asset\support_set contr_pf. contr_pf asset (Suc matur) \ borel_measurable (F matur)" using \0 \ diff_inv \trading_strategy contr_pf\ \0 \ diff_inv\ by (simp add: predict_stoch_proc_def trading_strategy_def) moreover have "\asset\support_set contr_pf. prices Mkt2 asset matur \ borel_measurable (F matur)" using 10 unfolding contr_pf_def using sum_support_set[of arb_pf pf] by auto ultimately show ?thesis by (metis (no_types, lifting) "1" \0 \ diff_inv\ quantity_adapted self_financingE) qed qed { assume "0 > diff_inv" define opp_pf where "opp_pf = qty_mult_comp contr_pf (\ n w. -1)" have "arbitrage_process Mkt2 opp_pf" proof (rule arbitrage_processI, rule exI, intro conjI) show "self_financing Mkt2 opp_pf" using 1 \0 > diff_inv\ mult_time_constant_self_financing[of contr_pf] 8 unfolding opp_pf_def by auto show "trading_strategy opp_pf" unfolding opp_pf_def proof (rule mult_comp_trading_strat) show "trading_strategy contr_pf" using 2 \0 > diff_inv\ by auto show "borel_predict_stoch_proc F (\n w. - 1)" by (simp add: constant_process_borel_predictable) qed show "\w\space M. cls_val_process Mkt2 opp_pf 0 w = 0" proof fix w assume "w\ space M" show "cls_val_process Mkt2 opp_pf 0 w = 0" using 3 8 \0 > diff_inv\ using \w \ space M\ mult_comp_cls_val_process0 opp_pf_def by fastforce qed have "AE w in M. 0 < cls_val_process Mkt2 opp_pf matur w" proof (rule AE_mp) show "AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w" using 5 \0 > diff_inv\ by auto show "AE w in M. cls_val_process Mkt2 contr_pf matur w < 0 \ 0 < cls_val_process Mkt2 opp_pf matur w" proof fix w assume "w\ space M" show "cls_val_process Mkt2 contr_pf matur w < 0 \ 0 < cls_val_process Mkt2 opp_pf matur w" proof assume "cls_val_process Mkt2 contr_pf matur w < 0" show "0 < cls_val_process Mkt2 opp_pf matur w" proof (cases "matur = 0") case False hence "\m. Suc m = matur" by presburger from this obtain m where "Suc m = matur" by auto hence "0 < cls_val_process Mkt2 opp_pf (Suc m) w" using 3 8 \0 > diff_inv\ \w \ space M\ mult_comp_cls_val_process_Suc opp_pf_def using \cls_val_process Mkt2 contr_pf matur w < 0\ by fastforce thus ?thesis using \Suc m = matur\ by simp next case True thus ?thesis using 3 8 \0 > diff_inv\ \w \ space M\ mult_comp_cls_val_process0 opp_pf_def using \cls_val_process Mkt2 contr_pf matur w < 0\ by auto qed qed qed qed thus "AE w in M. 0 \ cls_val_process Mkt2 opp_pf matur w" by auto show "0 < prob {w \ space M. 0 < cls_val_process Mkt2 opp_pf matur w}" proof - let ?P = "{w\ space M. 0 < cls_val_process Mkt2 opp_pf matur w}" have "cls_val_process Mkt2 opp_pf matur \ borel_measurable (F matur)" (*unfolding opp_pf_def *) proof - have "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" using 9 \0 > diff_inv\ by simp moreover have "portfolio contr_pf" using 8 \0 > diff_inv\ by simp moreover have "(\w. - 1) \ borel_measurable (F matur)" by (simp add:constant_process_borel_adapted) ultimately show ?thesis using mult_comp_cls_val_process_measurable proof - have "diff_inv \ 0" using \diff_inv < 0\ by blast then have "self_financing Mkt2 contr_pf" by (metis "1") then show ?thesis by (metis (no_types) \(\w. - 1) \ borel_measurable (F matur)\ \portfolio contr_pf\ \self_financing Mkt2 opp_pf\ \cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)\ mult_comp_val_process_measurable opp_pf_def self_financingE) qed qed moreover have "space M = space (F matur)" using filtration by (simp add: filtration_def subalgebra_def) ultimately have "?P \ sets (F matur)" using borel_measurable_iff_greater[of "val_process Mkt2 contr_pf matur" "F matur"] by auto hence "?P \ sets M" by (meson filtration filtration_def subalgebra_def subsetCE) hence "measure M ?P = 1" using prob_Collect_eq_1[of "\x. 0 < cls_val_process Mkt2 opp_pf matur x"] \AE w in M. 0 < cls_val_process Mkt2 opp_pf matur w\ \0 > diff_inv\ by blast thus ?thesis by simp qed qed have "\ p. portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p" proof(intro exI conjI) show "arbitrage_process Mkt2 opp_pf" using \arbitrage_process Mkt2 opp_pf\ . show "portfolio opp_pf" unfolding opp_pf_def using 8 \0 > diff_inv\ by (auto simp add: mult_comp_portfolio) show "support_set opp_pf \ stocks Mkt \ {x}" unfolding opp_pf_def using 7 \0 > diff_inv\ using mult_comp_support_set by fastforce qed } note negp = this { assume "0 < diff_inv" have "arbitrage_process Mkt2 contr_pf" proof (rule arbitrage_processI, rule exI, intro conjI) show "self_financing Mkt2 contr_pf" using 1 \0 < diff_inv\ by auto show "trading_strategy contr_pf" using 2 \0 < diff_inv\ by auto show "\w\space M. cls_val_process Mkt2 contr_pf 0 w = 0" using 3 \0 < diff_inv\ by auto show "AE w in M. 0 \ cls_val_process Mkt2 contr_pf matur w" using 4 \0 < diff_inv\ by auto show "0 < prob {w \ space M. 0 < cls_val_process Mkt2 contr_pf matur w}" proof - let ?P = "{w\ space M. 0 < cls_val_process Mkt2 contr_pf matur w}" have "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" using 9 \0 < diff_inv\ by auto moreover have "space M = space (F matur)" using filtration by (simp add: filtration_def subalgebra_def) ultimately have "?P \ sets (F matur)" using borel_measurable_iff_greater[of "val_process Mkt2 contr_pf matur" "F matur"] by auto hence "?P \ sets M" by (meson filtration filtration_def subalgebra_def subsetCE) hence "measure M ?P = 1" using prob_Collect_eq_1[of "\x. 0 < cls_val_process Mkt2 contr_pf matur x"] 4 \0 < diff_inv\ by blast thus ?thesis by simp qed qed have "\ p. portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p" proof(intro exI conjI) show "arbitrage_process Mkt2 contr_pf" using \arbitrage_process Mkt2 contr_pf\ . show "portfolio contr_pf" using 8 \0 < diff_inv\ by auto show "support_set contr_pf \ stocks Mkt \ {x}" using 7 \0 < diff_inv\ by auto qed } note posp = this have "diff_inv \ 0 \ \(\ pr. price_structure der matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using \coincides_on Mkt Mkt2 (stocks Mkt)\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ xasset posp negp by force } hence "diff_inv = 0" using fix_asset_price expr by metis moreover have "constant_image (prices Mkt pos_stock 0) > 0" by (simp add: adapted_init constant_image_pos pos_stock_borel_adapted positive) ultimately show ?thesis unfolding diff_inv_def by auto qed subsection \Risk-neutral probability space\ subsubsection \risk-free rate and discount factor processes\ fun disc_rfr_proc:: "real \ nat \ 'a \ real" where rfr_base: "(disc_rfr_proc r) 0 w = 1"| rfr_step: "(disc_rfr_proc r) (Suc n) w = (1+r) * (disc_rfr_proc r) n w" lemma disc_rfr_proc_borel_measurable: shows "(disc_rfr_proc r) n \ borel_measurable M" proof (induct n) case (Suc n) thus ?case by (simp add:borel_measurable_times) qed auto lemma disc_rfr_proc_nonrandom: fixes r::real shows "\n. disc_rfr_proc r n \ borel_measurable (F 0)" using disc_rfr_proc_borel_measurable by auto lemma (in disc_equity_market) disc_rfr_constant_time: shows "\c. \w \ space (F 0). (disc_rfr_proc r n) w = c" proof (rule triv_measurable_cst) show "space (F 0) = space M" using filtration by (simp add:filtration_def subalgebra_def) show "sets (F 0) = {{}, space M}" using info_disc_filtr by (simp add: bot_nat_def init_triv_filt_def) show "(disc_rfr_proc r n) \ borel_measurable (F 0)" using disc_rfr_proc_nonrandom by blast show "space M \ {}" by (simp add:not_empty) qed lemma (in disc_filtr_prob_space) disc_rfr_proc_borel_adapted: shows "borel_adapt_stoch_proc F (disc_rfr_proc r)" unfolding adapt_stoch_proc_def using disc_rfr_proc_nonrandom filtration unfolding filtration_def by (meson increasing_measurable_info le0) lemma disc_rfr_proc_positive: assumes "-1 < r" shows "\n w . 0 < disc_rfr_proc r n w" proof - fix n fix w::'a show "0 < disc_rfr_proc r n w" proof (induct n) case 0 thus ?case using assms "disc_rfr_proc.simps" by simp next case (Suc n) thus ?case using assms "disc_rfr_proc.simps" by simp qed qed lemma (in prob_space) disc_rfr_constant_time_pos: assumes "-1 < r" shows "\c > 0. \w \ space M. (disc_rfr_proc r n) w = c" proof - let ?F = "sigma (space M) {{}, space M}" have ex: "\c. \w \ space ?F. (disc_rfr_proc r n) w = c" proof (rule triv_measurable_cst) show "space ?F = space M" by simp show "sets ?F = {{}, space M}" by (meson sigma_algebra.sets_measure_of_eq sigma_algebra_trivial) show "(disc_rfr_proc r n) \ borel_measurable ?F" using disc_rfr_proc_borel_measurable by blast show "space M \ {}" by (simp add:not_empty) qed from this obtain c where "\w \ space ?F. (disc_rfr_proc r n) w = c" by auto note cprops = this have "c>0" proof - have "\ w. w\ space M" using subprob_not_empty by blast from this obtain w where "w\ space M" by auto hence "c = disc_rfr_proc r n w" using cprops by simp also have "... > 0" using disc_rfr_proc_positive[of r n] assms by simp finally show ?thesis . qed moreover have "space M = space ?F" by simp ultimately show ?thesis using ex using cprops by blast qed lemma disc_rfr_proc_Suc_div: assumes "-1 < r" shows "\w. disc_rfr_proc r (Suc n) w/disc_rfr_proc r n w = 1+r" proof - fix w::'a show "disc_rfr_proc r (Suc n) w/disc_rfr_proc r n w = 1+r" using disc_rfr_proc_positive assms by (metis rfr_step less_irrefl nonzero_eq_divide_eq) qed definition discount_factor where "discount_factor r n = (\w. inverse (disc_rfr_proc r n w))" lemma discount_factor_times_rfr: assumes "-1 < r" shows "(1+r) * discount_factor r (Suc n) w = discount_factor r n w" unfolding discount_factor_def using assms by simp lemma discount_factor_borel_measurable: shows "discount_factor r n \ borel_measurable M" unfolding discount_factor_def proof (rule borel_measurable_inverse) show "disc_rfr_proc r n \ borel_measurable M" by (simp add:disc_rfr_proc_borel_measurable) qed lemma discount_factor_init: shows "discount_factor r 0 = (\w. 1)" unfolding discount_factor_def by simp lemma discount_factor_nonrandom: shows "discount_factor r n \ borel_measurable M" unfolding discount_factor_def proof (rule borel_measurable_inverse) show "disc_rfr_proc r n \ borel_measurable M" by (simp add:disc_rfr_proc_borel_measurable) qed lemma discount_factor_positive: assumes "-1 < r" shows "\n w . 0 < discount_factor r n w" using assms disc_rfr_proc_positive unfolding discount_factor_def by auto lemma (in prob_space) discount_factor_constant_time_pos: assumes "-1 < r" shows "\c > 0. \w \ space M. (discount_factor r n) w = c" using disc_rfr_constant_time_pos unfolding discount_factor_def by (metis assms inverse_positive_iff_positive) locale rsk_free_asset = fixes Mkt r risk_free_asset assumes acceptable_rate: "-1 < r" and rf_price: "prices Mkt risk_free_asset = disc_rfr_proc r" and rf_stock: "risk_free_asset \ stocks Mkt" locale rfr_disc_equity_market = disc_equity_market + rsk_free_asset + assumes rd: "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" sublocale rfr_disc_equity_market \ disc_market_pos_stock _ _ _ "risk_free_asset" by (unfold_locales, (auto simp add: rf_stock rd disc_rfr_proc_positive rf_price acceptable_rate)) subsubsection \Discounted value of a stochastic process\ definition discounted_value where "discounted_value r X = (\ n w. discount_factor r n w * X n w)" lemma (in rfr_disc_equity_market) discounted_rfr: shows "discounted_value r (prices Mkt risk_free_asset) n w = 1" unfolding discounted_value_def discount_factor_def using rf_price by (metis less_irrefl mult.commute positive right_inverse) lemma discounted_init: shows "\w. discounted_value r X 0 w = X 0 w" unfolding discounted_value_def by (simp add: discount_factor_init) lemma discounted_mult: shows "\n w. discounted_value r (\m x. X m x * Y m x) n w = X n w * (discounted_value r Y) n w" by (simp add: discounted_value_def) lemma discounted_mult': shows "discounted_value r (\m x. X m x * Y m x) n w = X n w * (discounted_value r Y) n w" by (simp add: discounted_value_def) lemma discounted_mult_times_rfr: assumes "-1 < r" shows "discounted_value r (\m w. (1+r) * X w) (Suc n) w = discounted_value r (\m w. X w) n w" unfolding discounted_value_def using assms discount_factor_times_rfr discounted_mult by (simp add: discount_factor_times_rfr mult.commute) lemma discounted_cong: assumes "\n w. X n w = Y n w" shows "\ n w. discounted_value r X n w = discounted_value r Y n w" by (simp add: assms discounted_value_def) lemma discounted_cong': assumes "X n w = Y n w" shows "discounted_value r X n w = discounted_value r Y n w" by (simp add: assms discounted_value_def) lemma discounted_AE_cong: assumes "AE w in N. X n w = Y n w" shows "AE w in N. discounted_value r X n w = discounted_value r Y n w" proof (rule AE_mp) show "AE w in N. X n w = Y n w" using assms by simp show "AE w in N. X n w = Y n w \ discounted_value r X n w = discounted_value r Y n w" proof fix w assume "w\ space N" thus "X n w = Y n w \ discounted_value r X n w = discounted_value r Y n w " by (simp add:discounted_value_def) qed qed lemma discounted_sum: assumes "finite I" shows "\n w. (\ i\ I. (discounted_value r (\m x. f i m x)) n w) = (discounted_value r (\m x. (\i\ I. f i m x)) n w)" using assms(1) subset_refl[of I] proof (induct rule: finite_subset_induct) case empty then show ?case by (simp add: discounted_value_def) next case (insert a F) show ?case proof (intro allI) fix n w have "(\i\insert a F. discounted_value r (f i) n w) = discounted_value r (f a) n w + (\i\F. discounted_value r (f i) n w)" by (simp add: insert.hyps(1) insert.hyps(3)) also have "... = discounted_value r (f a) n w + discounted_value r (\m x. \i\F. f i m x) n w" using insert.hyps(4) by simp also have "... = discounted_value r (\m x. \i\insert a F. f i m x) n w" by (simp add: discounted_value_def insert.hyps(1) insert.hyps(3) ring_class.ring_distribs(1)) finally show "(\i\insert a F. discounted_value r (f i) n w) = discounted_value r (\m x. \i\insert a F. f i m x) n w" . qed qed lemma discounted_adapted: assumes "borel_adapt_stoch_proc F X" shows "borel_adapt_stoch_proc F (discounted_value r X)" unfolding adapt_stoch_proc_def proof fix t show "discounted_value r X t \ borel_measurable (F t)" unfolding discounted_value_def proof (rule borel_measurable_times) show "X t \ borel_measurable (F t)" using assms unfolding adapt_stoch_proc_def by simp show "discount_factor r t \ borel_measurable (F t)" using discount_factor_borel_measurable by auto qed qed lemma discounted_measurable: assumes "X\ borel_measurable N" shows "discounted_value r (\m. X) m \ borel_measurable N" unfolding discounted_value_def proof (rule borel_measurable_times) show "X\ borel_measurable N" using assms by simp show "discount_factor r m \ borel_measurable N" using discount_factor_borel_measurable by auto qed lemma (in prob_space) discounted_integrable: assumes "integrable N (X n)" and "-1 < r" and "space N = space M" shows "integrable N (discounted_value r X n)" unfolding discounted_value_def proof - have "\c> 0. \w \ space M. (discount_factor r n) w = c" using discount_factor_constant_time_pos assms by simp from this obtain c where "c > 0" and "\w \ space M. (discount_factor r n) w = c" by auto note cprops = this hence "\w \ space M. discount_factor r n w = c" using cprops by simp hence "\w \ space N. discount_factor r n w = c" using assms by simp thus "integrable N (\w. discount_factor r n w * X n w)" using \\w \ space N. discount_factor r n w = c\ assms Bochner_Integration.integrable_cong[of N N "(\w. discount_factor r n w * X n w)" "(\w. c * X n w)"] by simp qed subsubsection \Results on risk-neutral probability spaces\ definition (in rfr_disc_equity_market) risk_neutral_prob where "risk_neutral_prob N \ (prob_space N) \ (\ asset \ stocks Mkt. martingale N F (discounted_value r (prices Mkt asset)))" lemma integrable_val_process: assumes "\ asset \ support_set pf. integrable M (\w. prices Mkt asset n w * pf asset (Suc n) w)" shows "integrable M (val_process Mkt pf n)" proof (cases "portfolio pf") case False thus ?thesis unfolding val_process_def by simp next case True hence "val_process Mkt pf n = (\w. \x\support_set pf. prices Mkt x n w * pf x (Suc n) w)" unfolding val_process_def by simp moreover have "integrable M (\w. \x\support_set pf. prices Mkt x n w * pf x (Suc n) w)" using assms by simp ultimately show ?thesis by simp qed lemma integrable_self_fin_uvp: assumes "\ asset \ support_set pf. integrable M (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "self_financing Mkt pf" shows "integrable M (cls_val_process Mkt pf n)" proof - have "val_process Mkt pf n = cls_val_process Mkt pf n" using assms by (simp add:self_financingE) moreover have "integrable M (val_process Mkt pf n)" using assms by (simp add:integrable_val_process) ultimately show ?thesis by simp qed lemma (in rfr_disc_equity_market) stocks_portfolio_risk_neutral: assumes "risk_neutral_prob N" and "trading_strategy pf" and "subalgebra N M" and "support_set pf \ stocks Mkt" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" shows "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" proof have nsigfin: "\n. sigma_finite_subalgebra N (F n)" using assms unfolding risk_neutral_prob_def martingale_def subalgebra_def using filtration filtration_def risk_neutral_prob_def prob_space.subalgebra_sigma_finite in_stock by metis have "disc_filtr_prob_space N F" proof - have "prob_space N" using assms unfolding risk_neutral_prob_def by simp moreover have "disc_filtr N F" using assms subalgebra_filtration by (metis (no_types, lifting) filtration disc_filtr_def filtration_def) ultimately show ?thesis by (simp add: disc_filtr_prob_space_axioms_def disc_filtr_prob_space_def) qed fix asset assume "asset \ support_set pf" hence "asset \ stocks Mkt" using assms by auto have "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable M" using assms readable by (meson \asset \ stocks Mkt\ borel_adapt_stoch_proc_borel_measurable discounted_adapted rfr_disc_equity_market.risk_neutral_prob_def rfr_disc_equity_market_axioms) hence b: "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable N" using assms Conditional_Expectation.measurable_from_subalg[of N M _ borel] by auto show "AEeq N (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) (discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n)" proof - have "AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) w = (real_cond_exp N (F n) (\z. pf asset (Suc n) z * discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "AE w in N. discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n) w = pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) (Suc n) w" by (simp add: discounted_value_def) show "discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n) \ borel_measurable N" proof - have "(\y. prices Mkt asset (Suc n) y * pf asset (Suc n) y) \ borel_measurable N" using assms \asset\ support_set pf\ by (simp add:borel_measurable_integrable) thus ?thesis unfolding discounted_value_def using discount_factor_borel_measurable[of r "Suc n" N] by simp qed show "(\z. pf asset (Suc n) z * discounted_value r (prices Mkt asset) (Suc n) z) \ borel_measurable N" proof - have "pf asset (Suc n) \ borel_measurable M" using assms \asset\ support_set pf\ unfolding trading_strategy_def using borel_predict_stoch_proc_borel_measurable[of "pf asset"] by auto hence a: "pf asset (Suc n) \ borel_measurable N" using assms Conditional_Expectation.measurable_from_subalg[of N M _ borel] by blast show ?thesis using a b by simp qed qed also have "AE w in N. (real_cond_exp N (F n) (\z. pf asset (Suc n) z * discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w = pf asset (Suc n) w * (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w" proof (rule sigma_finite_subalgebra.real_cond_exp_mult) show "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable N" using b by simp show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "pf asset (Suc n) \ borel_measurable (F n)" using assms \asset\ support_set pf\ unfolding trading_strategy_def predict_stoch_proc_def by auto show "integrable N (\z. pf asset (Suc n) z * discounted_value r (prices Mkt asset) (Suc n) z)" proof - have "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" using assms \asset \ support_set pf\ by auto hence "integrable N (discounted_value r (\m w. prices Mkt asset m w * pf asset m w) (Suc n))" using assms unfolding risk_neutral_prob_def using acceptable_rate by (auto simp add:discounted_integrable subalgebra_def) thus ?thesis by (smt (verit, ccfv_SIG) Bochner_Integration.integrable_cong discounted_value_def mult.assoc mult.commute) qed qed also have "AE w in N. pf asset (Suc n) w * (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w = pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) n w" proof - have "AEeq N (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) (\z. discounted_value r (\m y. prices Mkt asset m y) n z)" proof - have "martingale N F (discounted_value r (prices Mkt asset))" using assms \asset \ stocks Mkt\ unfolding risk_neutral_prob_def by simp moreover have "filtrated_prob_space N F" using \disc_filtr_prob_space N F\ using assms(2) disc_filtr_prob_space.axioms(1) filtrated_prob_space.intro filtrated_prob_space_axioms.intro filtration prob_space_axioms by (metis assms(3) subalgebra_filtration) ultimately show ?thesis using martingaleAE[of N F "discounted_value r (prices Mkt asset)" n "Suc n"] assms by simp qed thus ?thesis by auto qed also have "AE w in N. pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) n w = discounted_value r (\m y. pf asset (Suc m) y * prices Mkt asset m y) n w" by (simp add: discounted_value_def) also have "AE w in N. discounted_value r (\m y. pf asset (Suc m) y * prices Mkt asset m y) n w = discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n w" by (simp add: discounted_value_def) finally show "AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) w = (\x. discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n x) w" . qed qed lemma (in rfr_disc_equity_market) self_fin_trad_strat_mart: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "trading_strategy pf" and "self_financing Mkt pf" and "stock_portfolio Mkt pf" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" shows "martingale N F (discounted_value r (cls_val_process Mkt pf))" (*unfolding martingale_def*) proof (rule disc_martingale_charact) show nsigfin: "\n. sigma_finite_subalgebra N (F n)" using filt_equiv_prob_space_subalgebra assms using filtration filtration_def risk_neutral_prob_def subalgebra_sigma_finite by fastforce show "filtration N F" using assms by (simp add:filt_equiv_filtration) have "borel_adapt_stoch_proc F (discounted_value r (cls_val_process Mkt pf))" using assms discounted_adapted cls_val_process_adapted[of pf] stock_portfolio_def by (metis (mono_tags, opaque_lifting) support_adapt_def readable subsetCE) thus "\m. discounted_value r (cls_val_process Mkt pf) m \ borel_measurable (F m)" unfolding adapt_stoch_proc_def by simp show "\t. integrable N (discounted_value r (cls_val_process Mkt pf) t)" proof fix t have "integrable N (cls_val_process Mkt pf t)" using assms by (simp add: integrable_self_fin_uvp) thus "integrable N (discounted_value r (cls_val_process Mkt pf) t)" using assms discounted_integrable acceptable_rate by (metis filt_equiv_space) qed show "\n. AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" proof fix n show "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" proof - { fix w assume "w\ space M" have "discounted_value r (cls_val_process Mkt pf) (Suc n) w = discount_factor r (Suc n) w * (\x\support_set pf. prices Mkt x (Suc n) w * pf x (Suc n) w)" unfolding discounted_value_def cls_val_process_def using assms unfolding trading_strategy_def by simp also have "... = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" by (metis (no_types, lifting) mult.assoc sum.cong sum_distrib_left) finally have "discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" . } hence space: "\w\ space M. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" by simp hence nspace: "\w\ space N. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" using assms by (simp add:filt_equiv_space) have sup_disc: "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" using assms by (simp add:stocks_portfolio_risk_neutral filt_equiv_imp_subalgebra stock_portfolio_def) have "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = real_cond_exp N (F n) (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w" proof (rule sigma_finite_subalgebra.real_cond_exp_cong') show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "\w\space N. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w" using nspace by (metis (no_types, lifting) discounted_value_def mult.assoc sum.cong) show "(discounted_value r (cls_val_process Mkt pf) (Suc n)) \ borel_measurable N" using assms using \\t. integrable N (discounted_value r (cls_val_process Mkt pf) t)\ by blast qed also have "AE w in N. real_cond_exp N (F n) (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w = (\x\ support_set pf. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w)" proof (rule sigma_finite_subalgebra.real_cond_exp_bsum) show "sigma_finite_subalgebra N (F n)" using filt_equiv_prob_space_subalgebra assms using filtration filtration_def risk_neutral_prob_def subalgebra_sigma_finite by fastforce fix asset assume "asset \ support_set pf" show "integrable N (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))" proof (rule discounted_integrable) show "integrable N (\y. prices Mkt asset (Suc n) y * pf asset (Suc n) y)" using assms \asset\ support_set pf\ by simp show "space N = space M" using assms by (metis filt_equiv_space) show "-1 < r" using acceptable_rate by simp qed qed also have "AE w in N. (\x\ support_set pf. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w) = (\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w)" proof (rule AE_sum) show "finite (support_set pf)" using assms(3) portfolio_def trading_strategy_def by auto show "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" using sup_disc by simp qed also have "AE w in N. (\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (cls_val_process Mkt pf) n w" proof fix w assume "w\ space N" have "(\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (\m y. (\x\ support_set pf. prices Mkt x m y * pf x (Suc m) y)) n w" using discounted_sum assms(3) portfolio_def trading_strategy_def by (simp add: discounted_value_def sum_distrib_left) also have "... = discounted_value r (val_process Mkt pf) n w" unfolding val_process_def by (simp add: portfolio_def) also have "... = discounted_value r (cls_val_process Mkt pf) n w" using assms by (simp add:self_financingE discounted_cong) finally show "(\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (cls_val_process Mkt pf) n w" . qed finally show "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" . qed qed qed lemma (in disc_filtr_prob_space) finite_integrable_vp: assumes "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "prob_space N" and "filt_equiv F M N" and "trading_strategy pf" and "\n. \ asset \ support_set pf. prices Mkt asset n \ borel_measurable M" shows "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms by simp have "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset n `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(2) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms by blast moreover have "(\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset n ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset n \ borel_measurable M" using assms \asset \ support_set pf\ by simp moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed lemma (in disc_filtr_prob_space) finite_integrable_uvp: assumes "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "prob_space N" and "filt_equiv F M N" and "trading_strategy pf" and "\n. \ asset \ support_set pf. prices Mkt asset n \ borel_measurable M" shows "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms by simp have "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset (Suc n) `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(2) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms by blast moreover have "(\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset (Suc n) ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset (Suc n) \ borel_measurable M" using assms using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed lemma (in rfr_disc_equity_market) self_fin_trad_strat_mart_finite: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "trading_strategy pf" and "self_financing Mkt pf" and "support_set pf \ stocks Mkt" and "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" shows "martingale N F (discounted_value r (cls_val_process Mkt pf))" proof (rule self_fin_trad_strat_mart, (simp add:assms)+) show "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms unfolding risk_neutral_prob_def by auto have "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset n `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(7) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms(6) by blast moreover have "(\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset n ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset n \ borel_measurable M" using assms readable using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed show "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms unfolding risk_neutral_prob_def by auto have "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset (Suc n) `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(7) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms(6) by blast moreover have "(\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset (Suc n) ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset (Suc n) \ borel_measurable M" using assms readable using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed show "stock_portfolio Mkt pf" using assms stock_portfolio_def by (simp add: stock_portfolio_def trading_strategy_def) qed lemma (in rfr_disc_equity_market) replicating_expectation: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "replicating_portfolio pf pyf matur" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" and "viable_market Mkt" and "sets (F 0) = {{}, space M}" and "pyf \ borel_measurable (F matur)" shows "fair_price Mkt (prob_space.expectation N (discounted_value r (\m. pyf) matur)) pyf matur" proof - have fn: "filtrated_prob_space N F" using assms by (simp add: \pyf \ borel_measurable (F matur)\ filtrated_prob_space_axioms.intro filtrated_prob_space_def risk_neutral_prob_def filt_equiv_filtration) have "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using assms(3) disc_equity_market.replicating_portfolio_def disc_equity_market_axioms discounted_adapted filtrated_prob_space.borel_adapt_stoch_proc_borel_measurable fn cls_val_process_adapted by (metis (no_types, opaque_lifting) support_adapt_def readable stock_portfolio_def subsetCE) have "discounted_value r (\m. pyf) matur \ borel_measurable N" proof - have "(\m. pyf) matur \ borel_measurable (F matur)" using assms by simp hence "(\m. pyf) matur \ borel_measurable M" using filtration filtrationE1 measurable_from_subalg by blast hence "(\m. pyf) matur \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) thus ?thesis by (simp add:discounted_measurable) qed have mpyf: "AE w in M. cls_val_process Mkt pf matur w = pyf w" using assms unfolding replicating_portfolio_def by simp have "AE w in N. cls_val_process Mkt pf matur w = pyf w" proof (rule filt_equiv_borel_AE_eq) show "filt_equiv F M N" using assms by simp show "pyf \ borel_measurable (F matur)" using assms by simp show "AE w in M. cls_val_process Mkt pf matur w = pyf w" using mpyf by simp show "cls_val_process Mkt pf matur \ borel_measurable (F matur)" using assms(3) price_structure_def replicating_price_process by (meson support_adapt_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms readable stock_portfolio_def subsetCE) qed hence disc:"AE w in N. discounted_value r (cls_val_process Mkt pf) matur w = discounted_value r (\m. pyf) matur w" by (simp add:discounted_AE_cong) have "AEeq N (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) (real_cond_exp N (F 0) (discounted_value r (\m. pyf) matur))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (F 0)" using filtrated_prob_space.axioms(1) filtrated_prob_space.filtration fn filtrationE1 prob_space.subalgebra_sigma_finite by blast show "AEeq N (discounted_value r (cls_val_process Mkt pf) matur) (discounted_value r (\m. pyf) matur)" using disc by simp show "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using \discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N\ . show "discounted_value r (\m. pyf) matur \ borel_measurable N" using \discounted_value r (\m. pyf) matur \ borel_measurable N\ . qed have "martingale N F (discounted_value r (cls_val_process Mkt pf))" using assms unfolding replicating_portfolio_def using self_fin_trad_strat_mart[of N pf] by (simp add: stock_portfolio_def) hence "AEeq N (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) (discounted_value r (cls_val_process Mkt pf) 0)" using martingaleAE[of N F "discounted_value r (cls_val_process Mkt pf)" 0 matur] fn by simp also have "AE w in N. (discounted_value r (cls_val_process Mkt pf) 0 w) = initial_value pf" proof fix w assume "w\ space N" have "discounted_value r (cls_val_process Mkt pf) 0 w = cls_val_process Mkt pf 0 w" by (simp add:discounted_init) also have "... = val_process Mkt pf 0 w" unfolding cls_val_process_def using assms unfolding replicating_portfolio_def stock_portfolio_def by simp also have "... = initial_value pf" using assms unfolding replicating_portfolio_def using \w\ space N\ by (metis (no_types, lifting) support_adapt_def filt_equiv_space initial_valueI readable stock_portfolio_def subsetCE) finally show "discounted_value r (cls_val_process Mkt pf) 0 w = initial_value pf" . qed finally have "AE w in N. (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) w = initial_value pf" . moreover have "\w\ space N. (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) w = prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur)" proof (rule prob_space.trivial_subalg_cond_expect_eq) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "subalgebra N (F 0)" using \prob_space N\ filtrated_prob_space.filtration fn filtrationE1 by blast show "sets (F 0) = {{}, space N}" using assms by (simp add:filt_equiv_space) show "integrable N (discounted_value r (cls_val_process Mkt pf) matur)" proof (rule discounted_integrable) show "space N = space M" using assms by (simp add:filt_equiv_space) show "integrable N (cls_val_process Mkt pf matur)" using assms unfolding replicating_portfolio_def by (simp add: integrable_self_fin_uvp) show "-1 < r" using acceptable_rate by simp qed qed ultimately have "AE w in N. prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = initial_value pf" by simp hence "prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = initial_value pf" using assms unfolding risk_neutral_prob_def using prob_space.emeasure_space_1[of N] AE_eq_cst[of _ _ N] by simp moreover have "prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = prob_space.expectation N (discounted_value r (\m. pyf) matur)" proof (rule integral_cong_AE) show "AEeq N (discounted_value r (cls_val_process Mkt pf) matur) (discounted_value r (\m. pyf) matur)" using disc by simp show "discounted_value r (\m. pyf) matur \ borel_measurable N" using \discounted_value r (\m. pyf) matur \ borel_measurable N\ . show "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using \discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N\ . qed ultimately have "prob_space.expectation N (discounted_value r (\m. pyf) matur) = initial_value pf" by simp thus ?thesis using assms by (metis (full_types) support_adapt_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms readable replicating_fair_price stock_portfolio_def subsetCE) qed lemma (in rfr_disc_equity_market) replicating_expectation_finite: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "replicating_portfolio pf pyf matur" and "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "viable_market Mkt" and "sets (F 0) = {{}, space M}" and "pyf \ borel_measurable (F matur)" shows "fair_price Mkt (prob_space.expectation N (discounted_value r (\m. pyf) matur)) pyf matur" proof - have "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule finite_integrable_vp, (auto simp add:assms)) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "\n asset. asset \ support_set pf \ random_variable borel (prices Mkt asset n)" proof- fix n fix asset assume "asset \ support_set pf" show "random_variable borel (prices Mkt asset n)" using assms unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def using readable by (meson \asset \ support_set pf\ adapt_stoch_proc_borel_measurable subsetCE) qed qed moreover have "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule finite_integrable_uvp, (auto simp add:assms)) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "\n asset. asset \ support_set pf \ random_variable borel (prices Mkt asset n)" proof- fix n fix asset assume "asset \ support_set pf" show "random_variable borel (prices Mkt asset n)" using assms unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def using readable by (meson \asset \ support_set pf\ adapt_stoch_proc_borel_measurable subsetCE) qed qed ultimately show ?thesis using assms replicating_expectation by simp qed end \ No newline at end of file diff --git a/thys/Factored_Transition_System_Bounding/Dependency.thy b/thys/Factored_Transition_System_Bounding/Dependency.thy --- a/thys/Factored_Transition_System_Bounding/Dependency.thy +++ b/thys/Factored_Transition_System_Bounding/Dependency.thy @@ -1,421 +1,421 @@ theory Dependency imports Main "HOL-Library.Finite_Map" FactoredSystem ActionSeqProcess RelUtils begin section \Dependency\ text \ State variable dependency analysis may be used to find structure in a factored system and find useful projections, for example on variable sets which are closed under mutual dependency. [Abdulaziz et al., p.13] In the following the dependency predicate (`dep`) is formalized and some dependency related propositions are proven. Dependency between variables `v1`, `v2` w.r.t to an action set @{term "\"} is given if one of the following holds: (1) `v1` and `v2` are equal (2) an action @{term "(p, e) \ \"} exists where @{term "v1 \ \(p)"} and @{term "v2 \ \(e)"} (meaning that it is a necessary condition that `p v1` is given if the action has effect `e v2`). (3) or, an action @{term "(p, e) \ \"} exists s.t. @{term "v1 v2 \ \(e)"} This notion is extended to sets of variables `vs1`, `vs2` (`dep\_var\_set`): `vs1` and `vs2` are dependent iff `vs1` and `vs2` are disjoint and if dependent `v1`, `v2` exist where @{term "v1 \ vs1"}, @{term "v2 \ vs2"}. [Abdulaziz et al., Definition 7, p.13][Abdulaziz et al., HOL4 Definition 5, p.14] \ subsection \Dependent Variables and Variable Sets\ \ \NOTE name shortened to 'dep'.\ definition dep where "dep PROB v1 v2 \ (\a. a \ PROB \ ( ((v1 \ fmdom' (fst a)) \ (v2 \ fmdom' (snd a))) \ ((v1 \ fmdom' (snd a) \ v2 \ fmdom' (snd a))) ) ) \ (v1 = v2)" \ \NOTE name shortened to 'dep\_var\_set'.\ definition dep_var_set where "dep_var_set PROB vs1 vs2 \ (disjnt vs1 vs2) \ (\ v1 v2. (v1 \ vs1) \ (v2 \ vs2) \ (dep PROB v1 v2) )" lemma dep_var_set_self_empty: fixes PROB vs assumes "dep_var_set PROB vs vs" shows "(vs = {})" using assms unfolding dep_var_set_def proof - obtain v1 v2 where "v1 \ vs" "v2 \ vs" "disjnt vs vs" "dep PROB v1 v2" using assms unfolding dep_var_set_def by blast then show ?thesis by force qed lemma DEP_REFL: fixes PROB shows "reflexive (\v v'. dep PROB v v')" unfolding dep_def reflexive_def by presburger \ \NOTE added lemma.\ lemma NEQ_DEP_IMP_IN_DOM_i: fixes a v assumes "a \ PROB" "v \ fmdom' (fst a)" shows "v \ prob_dom PROB" proof - have "v \ fmdom' (fst a)" using assms(2) by simp moreover have "fmdom' (fst a) \ prob_dom PROB" using assms(1) unfolding prob_dom_def action_dom_def using case_prod_beta' by auto ultimately show ?thesis by blast qed \ \NOTE added lemma.\ lemma NEQ_DEP_IMP_IN_DOM_ii: fixes a v assumes "a \ PROB" "v \ fmdom' (snd a)" shows "v \ prob_dom PROB" proof - have "v \ fmdom' (snd a)" using assms(2) by simp moreover have "fmdom' (snd a) \ prob_dom PROB" using assms(1) unfolding prob_dom_def action_dom_def using case_prod_beta' by auto ultimately show ?thesis by blast qed lemma NEQ_DEP_IMP_IN_DOM: fixes PROB :: "(('a, 'b) fmap \ ('a, 'b) fmap) set" and v v' assumes "\(v = v')" "(dep PROB v v')" shows "(v \ (prob_dom PROB) \ v' \ (prob_dom PROB))" using assms unfolding dep_def using FDOM_pre_subset_prob_dom_pair FDOM_eff_subset_prob_dom_pair proof - obtain a where 1: "a \ PROB" "(v \ fmdom' (fst a) \ v' \ fmdom' (snd a) \ v \ fmdom' (snd a) \ v' \ fmdom' (snd a))" using assms unfolding dep_def by blast then consider (i) "v \ fmdom' (fst a) \ v' \ fmdom' (snd a)" | (ii) "v \ fmdom' (snd a) \ v' \ fmdom' (snd a)" by blast then show ?thesis proof (cases) case i then have "v \ fmdom' (fst a)" "v' \ fmdom' (snd a)" by simp+ then have "v \ prob_dom PROB" "v' \ prob_dom PROB" using 1 NEQ_DEP_IMP_IN_DOM_i NEQ_DEP_IMP_IN_DOM_ii by metis+ then show ?thesis by simp next case ii then have "v \ fmdom' (snd a)" "v' \ fmdom' (snd a)" by simp+ then have "v \ prob_dom PROB" "v' \ prob_dom PROB" using 1 NEQ_DEP_IMP_IN_DOM_ii by metis+ then show ?thesis by simp qed qed lemma dep_sos_imp_mem_dep: fixes PROB S vs assumes "(dep_var_set PROB (\ S) vs)" shows "(\vs'. vs' \ S \ dep_var_set PROB vs' vs)" proof - obtain v1 v2 where obtain_v1_v2: "v1 \ \S" "v2 \ vs" "disjnt (\S) vs" "dep PROB v1 v2" using assms dep_var_set_def[of PROB "\ S" vs] by blast moreover { fix vs' assume "vs' \ S" moreover have "vs' \ (\S)" using calculation Union_upper by blast ultimately have "disjnt vs' vs" using obtain_v1_v2(3) disjnt_subset1 by blast } ultimately show ?thesis unfolding dep_var_set_def by blast qed lemma dep_union_imp_or_dep: fixes PROB vs vs' vs'' assumes "(dep_var_set PROB vs (vs' \ vs''))" shows "(dep_var_set PROB vs vs' \ dep_var_set PROB vs vs'')" proof - obtain v1 v2 where obtain_v1_v2: "v1 \ vs" "v2 \ vs' \ vs''" "disjnt vs (vs' \ vs'')" "dep PROB v1 v2" using assms dep_var_set_def[of PROB vs "(vs' \ vs'')"] by blast \ \NOTE The proofs for the cases introduced here yield the goal's left and right side respectively.\ consider (i) "v2 \ vs'" | (ii) "v2 \ vs''" using obtain_v1_v2(2) by blast then show ?thesis proof (cases) case i have "vs' \ vs' \ vs''" by auto moreover have "disjnt (vs' \ vs'') vs" using obtain_v1_v2(3) disjnt_sym by blast ultimately have "disjnt vs vs'" using disjnt_subset1 disjnt_sym by blast then have "dep_var_set PROB vs vs'" unfolding dep_var_set_def using obtain_v1_v2(1, 4) i by blast then show ?thesis by simp next case ii then have "vs'' \ vs' \ vs''" by simp moreover have "disjnt (vs' \ vs'') vs" using obtain_v1_v2(3) disjnt_sym by fast ultimately have "disjnt vs vs''" using disjnt_subset1 disjnt_sym by metis then have "dep_var_set PROB vs vs''" unfolding dep_var_set_def using obtain_v1_v2(1, 4) ii by blast then show ?thesis by simp qed qed \ \NOTE This is symmetrical to `dep\_sos\_imp\_mem\_dep` w.r.t to `vs` and @{term "\ S"}.\ lemma dep_biunion_imp_or_dep: fixes PROB vs S assumes "(dep_var_set PROB vs (\S))" shows "(\vs'. vs' \ S \ dep_var_set PROB vs vs')" proof - obtain v1 v2 where obtain_v1_v2: "v1 \ vs" "v2 \ (\S)" "disjnt vs (\S)" "dep PROB v1 v2" using assms dep_var_set_def[of PROB vs "\ S"] by blast moreover { fix vs' assume "vs' \ S" then have "vs' \ (\S)" using calculation Union_upper by blast moreover have "disjnt (\S) vs" using obtain_v1_v2(3) disjnt_sym by blast ultimately have "disjnt vs vs'" using obtain_v1_v2(3) disjnt_subset1 disjnt_sym by metis } ultimately show ?thesis unfolding dep_var_set_def by blast qed subsection "Transitive Closure of Dependent Variables and Variable Sets" definition dep_tc where "dep_tc PROB = TC (\v1' v2'. dep PROB v1' v2')" \ \NOTE type of `PROB` had to be fixed for MP on `NEQ\_DEP\_IMP\_IN\_DOM`.\ lemma dep_tc_imp_in_dom: fixes PROB :: "(('a, 'b) fmap \ ('a, 'b) fmap) set" and v1 v2 assumes "\(v1 = v2)" "(dep_tc PROB v1 v2)" shows "(v1 \ prob_dom PROB)" proof - have "TC (dep PROB) v1 v2" using assms(2) unfolding dep_tc_def by simp then have "dep PROB v1 v2 \ (\y. v1 \ y \ y \ v2 \ dep PROB v1 y \ TC (dep PROB) y v2)" using TC_CASES1_NEQ[where R = "(\v1' v2'. dep PROB v1' v2')" and x = v1 and z = v2] by (simp add: TC_equiv_tranclp) \ \NOTE Split on the disjunction yielded by the previous step.\ then consider (i) "dep PROB v1 v2" | (ii) "(\y. v1 \ y \ y \ v2 \ dep PROB v1 y \ TC (dep PROB) y v2)" by fast then show ?thesis proof (cases) case i { consider (II) "(\a. a \ PROB \ ( v1 \ fmdom' (fst a) \ v2 \ fmdom' (snd a) \ v1 \ fmdom' (snd a) \ v2 \ fmdom' (snd a)))" | (III) "v1 = v2" using i unfolding dep_def by blast then have ?thesis proof (cases) case II then obtain a where 1: "a \ PROB" "(v1 \ fmdom' (fst a) \ v2 \ fmdom' (snd a) \ v1 \ fmdom' (snd a) \ v2 \ fmdom' (snd a))" by blast then have "v1 \ fmdom' (fst a) \ fmdom' (snd a)" by blast then have 2: "v1 \ action_dom (fst a) (snd a)" unfolding action_dom_def by blast then have "action_dom (fst a) (snd a) \ prob_dom PROB" using 1(1) exec_as_proj_valid_2 by fast then have "v1 \ prob_dom PROB" using 1 2 by fast then show ?thesis by simp next case III then show ?thesis using assms(1) by simp qed } then show ?thesis by simp next case ii then obtain y where "v1 \ y" "y \ v2" "dep PROB v1 y" "TC (dep PROB) y v2" using ii by blast then show ?thesis using NEQ_DEP_IMP_IN_DOM by metis qed qed lemma not_dep_disj_imp_not_dep: fixes PROB vs_1 vs_2 vs_3 assumes "((vs_1 \ vs_2) = {})" "(vs_3 \ vs_2)" "\(dep_var_set PROB vs_1 vs_2)" shows "\(dep_var_set PROB vs_1 vs_3)" using assms subset_eq unfolding dep_var_set_def disjnt_def by blast lemma dep_slist_imp_mem_dep: fixes PROB vs lvs assumes "(dep_var_set PROB (\ (set lvs)) vs)" shows "(\vs'. ListMem vs' lvs \ dep_var_set PROB vs' vs)" proof - obtain v1 v2 where obtain_v1_v2: "v1 \ \(set lvs)" "v2 \ vs" "disjnt (\(set lvs)) vs" "dep PROB v1 v2" using assms dep_var_set_def[of PROB "\ (set lvs)" vs] by blast then obtain vs' where obtain_vs': "vs' \ set lvs" "v1 \ vs'" by blast then have "ListMem vs' lvs" using ListMem_iff by fast moreover { have "disjnt vs' vs" using obtain_v1_v2(3) obtain_vs'(1) by auto then have "dep_var_set PROB vs' vs" unfolding dep_var_set_def using obtain_v1_v2(1, 2, 4) obtain_vs'(2) by blast } ultimately show ?thesis by blast qed lemma n_bigunion_le_sum_3: fixes PROB vs svs assumes "(\ vs'. vs' \ svs \ \(dep_var_set PROB vs' vs))" shows "\(dep_var_set PROB (\svs) vs)" proof - { assume "(dep_var_set PROB (\svs) vs)" then obtain v1 v2 where obtain_vs: "v1 \ \svs" "v2 \ vs" "disjnt (\svs) vs" "dep PROB v1 v2" unfolding dep_var_set_def by blast then obtain vs' where obtain_vs': "v1 \ vs'" "vs' \ svs" by blast then have a: "disjnt vs' vs" using obtain_vs(3) obtain_vs'(2) disjnt_subset1 by blast then have "\v1 v2. \(v1 \ vs') \ \(v2 \ vs) \ \disjnt vs' vs \ \dep PROB v1 v2" using assms obtain_vs'(2) dep_var_set_def by fast then have False using a obtain_vs'(1) obtain_vs(2, 4) by blast } then show ?thesis by blast qed lemma disj_not_dep_vset_union_imp_or: fixes PROB a vs vs' assumes "(a \ PROB)" "(disjnt vs vs')" "(\(dep_var_set PROB vs' vs) \ \(dep_var_set PROB vs vs'))" "(varset_action a (vs \ vs'))" shows "(varset_action a vs \ varset_action a vs')" using assms unfolding varset_action_def dep_var_set_def dep_def proof - assume a1: "fmdom' (snd a) \ vs \ vs'" assume "disjnt vs vs'" assume " \ (disjnt vs' vs \ (\v1 v2. v1 \ vs' \ v2 \ vs \ ((\a. a \ PROB \ (v1 \ fmdom' (fst a) \ v2 \ fmdom' (snd a) \ v1 \ fmdom' (snd a) \ v2 \ fmdom' (snd a))) \ v1 = v2))) \ \ (disjnt vs vs' \ (\v1 v2. v1 \ vs \ v2 \ vs' \ ((\a. a \ PROB \ (v1 \ fmdom' (fst a) \ v2 \ fmdom' (snd a) \ v1 \ fmdom' (snd a) \ v2 \ fmdom' (snd a))) \ v1 = v2))) " then have f2: "\aa ab. aa \ vs \ ab \ vs' \ aa \ fmdom' (snd a) \ ab \ fmdom' (snd a)" using \a \ PROB\ \disjnt vs vs'\ disjnt_sym by blast obtain aa :: "'a set \ 'a set \ 'a" where f3: "\A Aa a Ab Ac. (A \ Aa \ aa A Aa \ A) \ (aa A Aa \ Aa \ A \ Aa) \ ((a::'a) \ Ab \ \ Ab \ Ac \ a \ Ac)" - by moura + by (atomize_elim, (subst choice_iff[symmetric])+, blast) then have "\A. fmdom' (snd a) \ A \ aa (fmdom' (snd a)) A \ vs \ aa (fmdom' (snd a)) A \ vs'" using a1 by (meson Un_iff) then show "fmdom' (snd a) \ vs \ fmdom' (snd a) \ vs'" using f3 f2 by meson qed end \ No newline at end of file diff --git a/thys/HOLCF-Prelude/examples/Sieve_Primes.thy b/thys/HOLCF-Prelude/examples/Sieve_Primes.thy --- a/thys/HOLCF-Prelude/examples/Sieve_Primes.thy +++ b/thys/HOLCF-Prelude/examples/Sieve_Primes.thy @@ -1,278 +1,277 @@ theory Sieve_Primes imports "HOL-Computational_Algebra.Primes" "../Num_Class" "../HOLCF_Prelude" (*FIXME: import order matters, if HOLCF_Prelude is not imported last, constants like "filter" etc. refer to List.thy; I'm afraid however that the current order only works by accident*) begin section \The Sieve of Eratosthenes\ declare [[coercion int]] declare [[coercion_enabled]] text \ This example proves that the well-known Haskell two-liner that lazily calculates the list of all primes does indeed do so. This proof is using coinduction. \ text \ We need to hide some constants again since we imported something from HOL not via @{theory "HOLCF-Prelude.HOLCF_Main"}. \ no_notation Rings.divide (infixl "div" 70) and Rings.modulo (infixl "mod" 70) no_notation Set.member ("(:)") and Set.member ("(_/ : _)" [51, 51] 50) text \This is the implementation. We also need a modulus operator.\ fixrec sieve :: "[Integer] \ [Integer]" where "sieve\(p : xs) = p : (sieve\(filter\(\ x. neg\(eq\(mod\x\p)\0))\xs))" fixrec primes :: "[Integer]" where "primes = sieve\[2..]" text \Simplification rules for modI:\ definition MkI' :: "int \ Integer" where "MkI' x = MkI\x" lemma MkI'_simps [simp]: shows "MkI' 0 = 0" and "MkI' 1 = 1" and "MkI' (numeral k) = numeral k" unfolding MkI'_def zero_Integer_def one_Integer_def numeral_Integer_eq by rule+ lemma modI_numeral_numeral [simp]: "mod\(numeral i)\(numeral j) = MkI' (Rings.modulo (numeral i) (numeral j))" unfolding numeral_Integer_eq mod_Integer_def MkI'_def by simp text \Some lemmas demonstrating evaluation of our list:\ lemma "primes !! 0 = 2" unfolding primes.simps apply (simp only: enumFrom_intsFrom_conv) apply (subst intsFrom.simps) apply simp done lemma "primes !! 1 = 3" unfolding primes.simps apply (simp only: enumFrom_intsFrom_conv) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply simp done lemma "primes !! 2 = 5" unfolding primes.simps apply (simp only: enumFrom_intsFrom_conv) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply simp done lemma "primes !! 3 = 7" unfolding primes.simps apply (simp only: enumFrom_intsFrom_conv) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (subst intsFrom.simps) apply (simp del: filter_FF filter_TT) (* FIXME: remove these two from the default simpset *) done text \Auxiliary lemmas about prime numbers\ lemma find_next_prime_nat: fixes n :: nat assumes "prime n" shows "\ n'. n' > n \ prime n' \ (\k. n < k \ k < n' \ \ prime k)" using ex_least_nat_le[of "\ k . k > n \ prime k"] by (metis bigger_prime not_prime_0) text \Simplification for andalso:\ lemma andAlso_Def[simp]: "((Def x) andalso (Def y)) = Def (x \ y)" by (metis Def_bool2 Def_bool4 andalso_thms(1) andalso_thms(2)) text \This defines the bisimulation and proves it to be a list bisimulation.\ definition prim_bisim: "prim_bisim x1 x2 = (\ n . prime n \ x1 = sieve\(filter\(\ (MkI\i). Def ((\d. d > 1 \ d < n \ \ (d dvd i))))\[MkI\n..]) \ x2 = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\n..])" lemma prim_bisim_is_bisim: "list_bisim prim_bisim" proof - { fix xs ys assume "prim_bisim xs ys" then obtain n :: nat where "prime n" and "n > 1" and "xs = sieve\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\n..])" (is "_ = sieve\?xs") and "ys = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\n..]" - (* sledgehammer *) proof - assume a1: "\n. \prime n; 1 < n; xs = sieve\ (Data_List.filter\ (\ (MkI\i). Def (\d>1. d < int n \ \ d dvd i))\ [MkI\(int n)..]); ys = Data_List.filter\ (\ (MkI\i). Def (prime (nat \i\)))\ [MkI\(int n)..]\ \ thesis" - obtain ii where + obtain ii where f2: "\is isa. (\ prim_bisim is isa \ prime (ii is isa) \ is = sieve\ (Data_List.filter\ (\ (MkI\i). Def (\ia>1. ia < ii is isa \ \ ia dvd i))\ [MkI\(ii is isa)..]) \ isa = Data_List.filter\ (\ (MkI\i). Def (prime (nat \i\)))\ [MkI\(ii is isa)..]) \ ((\i. \ prime i \ is \ sieve\ (Data_List.filter\ (\ (MkI\ia). Def (\ib>1. ib < i \ \ ib dvd ia))\ [MkI\i..]) \ isa \ Data_List.filter\ (\ (MkI\i). Def (prime (nat \i\)))\ [MkI\i..]) \ prim_bisim is isa)" - using prim_bisim by moura + unfolding prim_bisim by (atomize_elim, (subst choice_iff[symmetric])+, blast) then have f3: "prime (ii xs ys)" using \prim_bisim xs ys\ by blast then obtain nn :: "int \ nat" where f4: "int (nn (ii xs ys)) = ii xs ys" by (metis (no_types) prime_gt_0_int zero_less_imp_eq_int) then have "prime (nn (ii xs ys))" using f3 by (metis (no_types) prime_nat_int_transfer) then show ?thesis using f4 f2 a1 \prim_bisim xs ys\ prime_gt_1_nat by presburger qed obtain n' where "n' > n" and "prime n'" and not_prime: "\k. n < k \ k < n' \ \ prime k" using find_next_prime_nat[OF \prime n\] by auto { fix k :: int assume "n < k" and "k < n'" have "k > 1" using \n < k\ \n > 1\ by auto then obtain k' :: nat where "k = int k'" by (cases k) auto then obtain p where "prime p" and "p dvd k" using \k > 1\ \k = int k'\ by (metis (full_types) less_numeral_extra(4) of_nat_1 of_nat_dvd_iff prime_factor_nat prime_nat_int_transfer) then have "p < n'" using \k < n'\ \k > 1\ using zdvd_imp_le [of p k] by simp then have "p \ n" using \prime p\ not_prime using not_le prime_gt_0_int zero_less_imp_eq_int by (metis of_nat_less_iff prime_nat_int_transfer) then have "\d::int>1. d \ n \ d dvd k" using \p dvd k\ \prime p\ of_nat_le_iff prime_gt_1_nat prime_gt_1_int by auto } then have between_have_divisors: "\k::int. n < k \ k < n' \ \d::int>1. d \ n \ d dvd k". { fix i { assume small: "\d::int>1. d \ n \ \ d dvd i" fix d assume "1 < d" and "d dvd i" and "d < n'" with small have "d > n" by auto obtain d'::int where "d' > 1" and "d' \ n" and "d' dvd d" using between_have_divisors[OF \n < d\ \d < n'\] by auto with \d dvd i\ small have False by (metis (full_types) dvd_trans) } then have "(\d::int. d > 1 \ d \ n \ \ (d dvd i)) = (\d::int. d > 1 \ d < n' \ \ (d dvd i))" using \n' > n\ by auto } then have between_not_relvant: "\ i. (\d::int. d > 1 \ d \ n \ \ (d dvd i)) = (\d::int. d > 1 \ d < n' \ \ (d dvd i))" . from \prime n\ have "\d::int >1. d < n \ \ d dvd n" unfolding prime_int_altdef using int_one_le_iff_zero_less le_less by (simp add: prime_int_not_dvd) then obtain xs' where "?xs = (MkI\n) : xs'" and "xs' = filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\(n+1)..]" by (subst (asm) intsFrom.simps[unfolded enumFrom_intsFrom_conv[symmetric]], simp add: one_Integer_def TT_def[symmetric] add.commute) { have "filter\(\ x. neg\(eq\(mod\x\(MkI\n))\0))\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\(n+1)..]) = filter\(\ (MkI\x). Def (\ n dvd x))\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\(n+1)..])" apply (rule filter_cong[rule_format]) apply (rename_tac x, case_tac x) apply (simp) apply (auto simp add: zero_Integer_def) apply (rule FF_def) apply (simp add: TT_def) by (metis dvd_eq_mod_eq_0) also have "... = filter\(\ (MkI\i). Def ((\ n dvd i) \ (\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\(n+1)..]" by (auto intro!: filter_cong[rule_format]) also have "... = filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d \ n \ \ (d dvd i))))\[MkI\(n+1)..]" apply (rule filter_cong[rule_format]) apply (rename_tac x, case_tac x) using \n > 1\ apply auto done also have "... = filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d \ n \ \ (d dvd i))))\[MkI\(int n+1)..]" by (metis (no_types, lifting) of_nat_1 of_nat_add) also have "... = filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d \ n \ \ (d dvd i))))\[MkI\n'..]" apply (rule filter_fast_forward[of n n']) using \n' > n\ apply (auto simp add: between_have_divisors) done also have "... = filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n' \ \ (d dvd i))))\[MkI\n'..]" by (auto intro: filter_cong[rule_format] simp add: between_not_relvant) also note calculation } note tmp = this { have "xs = sieve\?xs" by fact also have "... = sieve\((MkI\n) : xs')" using \?xs = _\ by simp also have "... = sieve\((MkI\n) : filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\(n+1)..])" using \xs' = _\ by simp also have "... = (MkI\n) : sieve\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n' \ \ (d dvd i))))\[MkI\n'..])" using tmp by simp also note calculation } moreover { have "ys = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\n..]" by fact also have "... = (MkI\n) : filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\(int n+1)..]" using \prime n\ by (subst intsFrom.simps[unfolded enumFrom_intsFrom_conv[symmetric]])(simp add: one_Integer_def TT_def[symmetric] add.commute) also have "... = (MkI\n) : filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\n'..]" apply (subst filter_fast_forward[of n n']) using \n' > n\ and not_prime apply auto apply (metis (full_types) \\k. \int n < k; k < int n'\ \ \d>1. d \ int n \ d dvd k\ le_less not_le prime_gt_0_int prime_int_not_dvd zdvd_imp_le) done also note calculation } moreover note \prime n'\ ultimately have "\ p xs' ys'. xs = p : xs' \ ys = p : ys' \ prim_bisim xs' ys'" unfolding prim_bisim using prime_nat_int_transfer by blast } then show ?thesis unfolding list.bisim_def by metis qed text \Now we apply coinduction:\ lemma sieve_produces_primes: fixes n :: nat assumes "prime n" shows "sieve\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < n \ \ (d dvd i))))\[MkI\n..]) = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\n..]" using assms apply - apply (rule list.coinduct[OF prim_bisim_is_bisim], auto simp add: prim_bisim) using prime_nat_int_transfer by blast text \And finally show the correctness of primes.\ theorem primes: shows "primes = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\2..]" proof - have "primes = sieve\[2 ..]" by (rule primes.simps) also have "... = sieve\(filter\(\ (MkI\i). Def ((\d::int. d > 1 \ d < (int 2) \ \ (d dvd i))))\[MkI\(int 2)..])" unfolding numeral_Integer_eq by (subst filter_TT, auto) also have "... = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\(int 2)..]" by (rule sieve_produces_primes[OF two_is_prime_nat]) also have "... = filter\(\ (MkI\i). Def (prime (nat \i\)))\[MkI\2..]" by simp finally show ?thesis . qed end diff --git a/thys/IsaGeoCoq/Tarski_Neutral.thy b/thys/IsaGeoCoq/Tarski_Neutral.thy --- a/thys/IsaGeoCoq/Tarski_Neutral.thy +++ b/thys/IsaGeoCoq/Tarski_Neutral.thy @@ -1,28006 +1,28006 @@ (* IsageoCoq Port part of GeoCoq 3.4.0 (https://geocoq.github.io/GeoCoq/) in Isabelle/Hol (Isabelle2021) Copyright (C) 2021 Roland Coghetto roland_coghetto (at) hotmail.com License: LGPL This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, MA 02110-1301 USA *) theory Tarski_Neutral imports Main begin section "Tarski's axiom system for neutral geometry" subsection "Tarski's axiom system for neutral geometry: dimensionless" locale Tarski_neutral_dimensionless = fixes Bet :: "'p \ 'p \ 'p \ bool" fixes Cong :: "'p \ 'p \ 'p \ 'p \ bool" assumes cong_pseudo_reflexivity: "\ a b. Cong a b b a" and cong_inner_transitivity: "\ a b p q r s. Cong a b p q \ Cong a b r s \ Cong p q r s" and cong_identity: "\ a b c. Cong a b c c \ a = b" and segment_construction: "\ a b c q. \x. (Bet q a x \ Cong a x b c)" and five_segment: "\ a b c a' b' c'. a \ b \ Bet a b c \ Bet a' b' c'\ Cong a b a' b' \ Cong b c b' c' \ Cong a d a' d' \ Cong b d b' d' \ Cong c d c' d'" and between_identity: "\ a b. Bet a b a \ a = b" and inner_pasch: "\ a b c p q. Bet a p c \ Bet b q c \ (\ x. Bet p x b \ Bet q x a)" and lower_dim: "\ a b c. (\ Bet a b c \ \ Bet b c a \ \ Bet c a b)" subsection "Tarski's axiom system for neutral geometry: 2D" locale Tarski_2D = Tarski_neutral_dimensionless + assumes upper_dim: "\ a b c p q. p \ q \ Cong a p a q \ Cong b p b q \ Cong c p c q \ (Bet a b c \ Bet b c a \ Bet c a b)" section "Definitions" subsection "Tarski's axiom system for neutral geometry: dimensionless" context Tarski_neutral_dimensionless begin subsubsection "Congruence" definition OFSC :: "['p,'p,'p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ _ OFSC _ _ _ _" [99,99,99,99,99,99,99,99] 50) where "A B C D OFSC A' B' C' D' \ Bet A B C \ Bet A' B' C' \ Cong A B A' B' \ Cong B C B' C' \ Cong A D A' D' \ Cong B D B' D'" definition Cong3 :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ Cong3 _ _ _" [99,99,99,99,99,99] 50) where "A B C Cong3 A' B' C' \ Cong A B A' B' \ Cong A C A' C' \ Cong B C B' C'" subsubsection "Betweenness" definition Col :: "['p,'p,'p] \ bool" ("Col _ _ _" [99,99,99] 50) where "Col A B C \ Bet A B C \ Bet B C A \ Bet C A B" definition Bet4 :: "['p,'p,'p,'p] \ bool" ("Bet4 _ _ _ _" [99,99,99,99] 50) where "Bet4 A1 A2 A3 A4 \ Bet A1 A2 A3 \ Bet A2 A3 A4 \ Bet A1 A3 A4 \ Bet A1 A2 A4" definition BetS :: "['p,'p,'p] \ bool" ("BetS _ _ _" [99,99,99] 50) where "BetS A B C \ Bet A B C \ A \ B \ B \ C" subsubsection "Collinearity" definition FSC :: "['p,'p,'p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ _ FSC _ _ _ _" [99,99,99,99,99,99,99,99] 50) where "A B C D FSC A' B' C' D' \ Col A B C \ A B C Cong3 A' B' C' \ Cong A D A' D' \ Cong B D B' D'" subsubsection "Congruence and Betweenness" definition IFSC :: "['p,'p,'p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ _ IFSC _ _ _ _" [99,99,99,99,99,99,99,99] 50) where "A B C D IFSC A' B' C' D' \ Bet A B C \ Bet A' B' C' \ Cong A C A' C' \ Cong B C B' C' \ Cong A D A' D' \ Cong C D C' D'" subsubsection "Between transivitity LE" definition Le :: "['p,'p,'p,'p] \ bool" ("_ _ Le _ _" [99,99,99,99] 50) where "A B Le C D \ \ E. (Bet C E D \ Cong A B C E)" definition Lt :: "['p,'p,'p,'p] \ bool" ("_ _ Lt _ _" [99,99,99,99] 50) where "A B Lt C D \ A B Le C D \ \ Cong A B C D" definition Ge :: "['p,'p,'p,'p] \ bool" ("_ _Ge _ _" [99,99,99,99] 50) where "A B Ge C D \ C D Le A B" definition Gt :: "['p,'p,'p,'p] \ bool" ("_ _ Gt _ _" [99,99,99,99] 50) where "A B Gt C D \ C D Lt A B" subsubsection "Out lines" definition Out :: "['p,'p,'p] \ bool" ("_ Out _ _" [99,99,99] 50) where "P Out A B \ A \ P \ B \ P \ (Bet P A B \ Bet P B A)" subsubsection "Midpoint" definition Midpoint :: "['p,'p,'p] \ bool" ("_ Midpoint _ _" [99,99,99] 50) where "M Midpoint A B \ Bet A M B \ Cong A M M B" subsubsection "Orthogonality" definition Per :: "['p,'p,'p] \ bool" ("Per _ _ _" [99,99,99] 50) where "Per A B C \ \ C'::'p. (B Midpoint C C' \ Cong A C A C')" definition PerpAt :: "['p,'p,'p,'p,'p] \ bool" ("_ PerpAt _ _ _ _ " [99,99,99,99,99] 50) where "X PerpAt A B C D \ A \ B \ C \ D \ Col X A B \ Col X C D \ (\ U V. ((Col U A B \ Col V C D) \ Per U X V))" definition Perp :: "['p,'p,'p,'p] \ bool" ("_ _ Perp _ _" [99,99,99,99] 50) where "A B Perp C D \ \ X::'p. X PerpAt A B C D" subsubsection "Coplanar" definition Coplanar :: "['p,'p,'p,'p] \ bool" ("Coplanar _ _ _ _" [99,99,99,99] 50) where "Coplanar A B C D \ \ X. (Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" definition TS :: "['p,'p,'p,'p] \ bool" ("_ _ TS _ _" [99,99,99,99] 50) where "A B TS P Q \ \ Col P A B \ \ Col Q A B \ (\ T::'p. Col T A B \ Bet P T Q)" definition ReflectL :: "['p,'p,'p,'p] \ bool" ("_ _ ReflectL _ _" [99,99,99,99] 50) where "P' P ReflectL A B \ (\ X. X Midpoint P P' \ Col A B X) \ (A B Perp P P' \ P = P')" definition Reflect :: "['p,'p,'p,'p] \ bool" ("_ _ Reflect _ _" [99,99,99,99] 50) where "P' P Reflect A B \ (A \ B \ P' P ReflectL A B) \ (A = B \ A Midpoint P P')" definition InAngle :: "['p,'p,'p,'p] \ bool" ("_ InAngle _ _ _" [99,99,99,99] 50) where "P InAngle A B C \ A \ B \ C \ B \ P \ B \ (\ X. Bet A X C \ (X = B \ B Out X P))" definition ParStrict:: "['p,'p,'p,'p] \ bool" ("_ _ ParStrict _ _" [99,99,99,99] 50) where "A B ParStrict C D \ Coplanar A B C D \ \ (\ X. Col X A B \ Col X C D)" definition Par:: "['p,'p,'p,'p] \ bool" ("_ _ Par _ _" [99,99,99,99] 50) where "A B Par C D \ A B ParStrict C D \ (A \ B \ C \ D \ Col A C D \ Col B C D)" definition Plg:: "['p,'p,'p,'p] \ bool" ("Plg _ _ _ _" [99,99,99,99] 50) where "Plg A B C D \ (A \ C \ B \ D) \ (\ M. M Midpoint A C \ M Midpoint B D)" definition ParallelogramStrict:: "['p,'p,'p,'p] \ bool" ("ParallelogramStrict _ _ _ _" [99,99,99,99] 50) where "ParallelogramStrict A B A' B' \ A A' TS B B' \ A B Par A' B' \ Cong A B A' B'" definition ParallelogramFlat:: "['p,'p,'p,'p] \ bool" ("ParallelogramFlat _ _ _ _" [99,99,99,99] 50) where "ParallelogramFlat A B A' B' \ Col A B A' \ Col A B B' \ Cong A B A' B' \ Cong A B' A' B \ (A \ A' \ B \ B')" definition Parallelogram:: "['p,'p,'p,'p] \ bool" ("Parallelogram _ _ _ _" [99,99,99,99] 50) where "Parallelogram A B A' B' \ ParallelogramStrict A B A' B' \ ParallelogramFlat A B A' B'" definition Rhombus:: "['p,'p,'p,'p] \ bool" ("Rhombus _ _ _ _" [99,99,99,99] 50) where "Rhombus A B C D \ Plg A B C D \ Cong A B B C" definition Rectangle:: "['p,'p,'p,'p] \ bool" ("Rectangle _ _ _ _" [99,99,99,99] 50) where "Rectangle A B C D \ Plg A B C D \ Cong A C B D" definition Square:: "['p,'p,'p,'p] \ bool" ("Square _ _ _ _" [99,99,99,99] 50) where "Square A B C D \ Rectangle A B C D \ Cong A B B C" definition Lambert:: "['p,'p,'p,'p] \ bool" ("Lambert _ _ _ _" [99,99,99,99] 50) where "Lambert A B C D \ A \ B \ B \ C \ C \ D \ A \ D \ Per B A D \ Per A D C \ Per A B C \ Coplanar A B C D" subsubsection "Plane" definition OS :: "['p,'p,'p,'p] \ bool" ("_ _ OS _ _" [99,99,99,99] 50) where "A B OS P Q \ \ R::'p. A B TS P R \ A B TS Q R" definition TSP :: "['p,'p,'p,'p,'p] \ bool" ("_ _ _TSP _ _" [99,99,99,99,99] 50) where "A B C TSP P Q \ (\ Coplanar A B C P) \ (\ Coplanar A B C Q) \ (\ T. Coplanar A B C T \ Bet P T Q)" definition OSP :: "['p,'p,'p,'p,'p] \ bool" ("_ _ _ OSP _ _" [99,99,99,99,99] 50) where "A B C OSP P Q \ \ R. ((A B C TSP P R) \ (A B C TSP Q R))" definition Saccheri:: "['p,'p,'p,'p] \ bool" ("Saccheri _ _ _ _" [99,99,99,99] 50) where "Saccheri A B C D \ Per B A D \ Per A D C \ Cong A B C D \ A D OS B C" subsubsection "Line reflexivity 2D" definition ReflectLAt :: "['p,'p,'p,'p,'p] \ bool" ("_ ReflectLAt _ _ _ _" [99,99,99,99,99] 50) where "M ReflectLAt P' P A B \ (M Midpoint P P' \ Col A B M) \ (A B Perp P P' \ P = P')" definition ReflectAt :: "['p,'p,'p,'p,'p] \ bool" ("_ ReflectAt _ _ _ _" [99,99,99,99,99] 50) where "M ReflectAt P' P A B \ (A \ B \ M ReflectLAt P' P A B) \ (A = B \ A = M \ M Midpoint P P')" subsubsection "Line reflexivity" definition upper_dim_axiom :: "bool" ("UpperDimAxiom" [] 50) where "upper_dim_axiom \ \ A B C P Q. P \ Q \ Cong A P A Q \ Cong B P B Q \ Cong C P C Q \ (Bet A B C \ Bet B C A \ Bet C A B)" definition all_coplanar_axiom :: "bool" ("AllCoplanarAxiom" [] 50) where "AllCoplanarAxiom \ \ A B C P Q. P \ Q \ Cong A P A Q \ Cong B P B Q \ Cong C P C Q \ (Bet A B C \ Bet B C A \ Bet C A B)" subsubsection "Angles" definition CongA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ CongA _ _ _" [99,99,99,99,99,99] 50) where "A B C CongA D E F \ A \ B \ C \ B \ D \ E \ F \ E \ (\ A' C' D' F'. Bet B A A' \ Cong A A' E D \ Bet B C C' \ Cong C C' E F \ Bet E D D' \ Cong D D' B A \ Bet E F F' \ Cong F F' B C \ Cong A' C' D' F')" definition LeA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ LeA _ _ _" [99,99,99,99,99,99] 50) where "A B C LeA D E F \ \ P. (P InAngle D E F \ A B C CongA D E P)" definition LtA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ LtA _ _ _" [99,99,99,99,99,99] 50) where "A B C LtA D E F \ A B C LeA D E F \ \ A B C CongA D E F" definition GtA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ GtA _ _ _" [99,99,99,99,99,99] 50) where "A B C GtA D E F \ D E F LtA A B C" definition Acute :: "['p,'p,'p] \ bool" ("Acute _ _ _" [99,99,99] 50) where "Acute A B C \ \ A' B' C'. (Per A' B' C' \ A B C LtA A' B' C')" definition Obtuse :: "['p,'p,'p] \ bool" ("Obtuse _ _ _" [99,99,99] 50) where "Obtuse A B C \ \ A' B' C'. (Per A' B' C' \ A' B' C' LtA A B C)" definition OrthAt :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ OrthAt _ _ _ _ _" [99,99,99,99,99,99] 50) where "X OrthAt A B C U V \ \ Col A B C \ U \ V \ Coplanar A B C X \ Col U V X \ (\ P Q. (Coplanar A B C P \ Col U V Q) \ Per P X Q)" definition Orth :: "['p,'p,'p,'p,'p] \ bool" ("_ _ _ Orth _ _" [99,99,99,99,99] 50) where "A B C Orth U V \ \ X. X OrthAt A B C U V" definition SuppA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ SuppA _ _ _ " [99,99,99,99,99,99] 50) where "A B C SuppA D E F \ A \ B \ (\ A'. Bet A B A' \ D E F CongA C B A')" subsubsection "Sum of angles" definition SumA :: "['p,'p,'p,'p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ _ _ _ SumA _ _ _" [99,99,99,99,99,99,99,99,99] 50) where "A B C D E F SumA G H I \ \ J. (C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I)" definition TriSumA :: "['p,'p,'p,'p,'p,'p] \ bool" ("_ _ _ TriSumA _ _ _" [99,99,99,99,99,99] 50) where "A B C TriSumA D E F \ \ G H I. (A B C B C A SumA G H I \ G H I C A B SumA D E F)" definition SAMS :: "['p,'p,'p,'p,'p,'p] \ bool" ("SAMS _ _ _ _ _ _" [99,99,99,99,99,99] 50) where "SAMS A B C D E F \ (A \ B \ (E Out D F \ \ Bet A B C)) \ (\ J. (C B J CongA D E F \ \ (B C OS A J) \ \ (A B TS C J) \ Coplanar A B C J))" subsubsection "Parallelism" definition Inter :: "['p,'p,'p,'p,'p] \ bool" ("_ Inter _ _ _ _" [99,99,99,99,99] 50) where "X Inter A1 A2 B1 B2 \ B1 \ B2 \ (\ P::'p. (Col P B1 B2 \ \ Col P A1 A2)) \ Col A1 A2 X \ Col B1 B2 X" subsubsection "Perpendicularity" definition Perp2 :: "['p,'p,'p,'p,'p] \ bool" ("_ Perp2 _ _ _ _" [99,99,99,99,99] 50) where "P Perp2 A B C D \ \ X Y. (Col P X Y \ X Y Perp A B \ X Y Perp C D)" subsubsection "Lentgh" definition QCong:: "(['p,'p] \ bool) \ bool" ("QCong _" [99] 50) where "QCong l \ \ A B. (\ X Y. (Cong A B X Y \ l X Y))" definition TarskiLen:: "['p,'p,(['p,'p] \ bool)] \ bool" ("TarskiLen _ _ _" [99,99,99] 50) where "TarskiLen A B l \ QCong l \ l A B" definition QCongNull :: "(['p,'p] \ bool) \ bool" ("QCongNull _" [99] 50) where "QCongNull l \ QCong l \ (\ A. l A A)" subsubsection "Equivalence Class of Angles" definition QCongA :: "(['p, 'p, 'p] \ bool) \ bool" ("QCongA _" [99] 50) where "QCongA a \ \ A B C. (A \ B \ C \ B \ (\ X Y Z. A B C CongA X Y Z \ a X Y Z))" definition Ang :: "['p,'p,'p, (['p, 'p, 'p] \ bool) ] \ bool" ("_ _ _ Ang _" [99,99,99,99] 50) where "A B C Ang a \ QCongA a \ a A B C" definition QCongAAcute :: "(['p, 'p, 'p] \ bool) \ bool" ("QCongAACute _" [99] 50) where "QCongAAcute a \ \ A B C. (Acute A B C \ (\ X Y Z. (A B C CongA X Y Z \ a X Y Z)))" definition AngAcute :: "['p,'p,'p, (['p,'p,'p] \ bool)] \ bool" ("_ _ _ AngAcute _" [99,99,99,99] 50) where "A B C AngAcute a \ ((QCongAAcute a) \ (a A B C))" definition QCongANullAcute :: "(['p,'p,'p] \ bool) \ bool" ("QCongANullAcute _" [99] 50) where "QCongANullAcute a \ QCongAAcute a \ (\ A B C. (a A B C \ B Out A C))" definition QCongAnNull :: "(['p,'p,'p] \ bool) \ bool" ("QCongAnNull _" [99] 50) where "QCongAnNull a \ QCongA a \ (\ A B C. (a A B C \ \ B Out A C))" definition QCongAnFlat :: "(['p,'p,'p] \ bool) \ bool" ("QCongAnFlat _" [99] 50) where "QCongAnFlat a \ QCongA a \ (\ A B C. (a A B C \ \ Bet A B C))" definition IsNullAngaP :: "(['p,'p,'p] \ bool) \ bool" ("IsNullAngaP _" [99] 50) where "IsNullAngaP a\ QCongAAcute a \ (\ A B C. (a A B C \ B Out A C))" definition QCongANull :: "(['p,'p,'p] \ bool) \ bool" ("QCongANull _" [99] 50) where "QCongANull a \ QCongA a \ (\ A B C. (a A B C \ B Out A C))" definition AngFlat :: "(['p, 'p, 'p] \ bool) \ bool" ("AngFlat _" [99] 50) where "AngFlat a \ QCongA a \ (\ A B C. (a A B C \ Bet A B C))" subsection "Parallel's definition Postulate" definition tarski_s_parallel_postulate :: "bool" ("TarskiSParallelPostulate") where "tarski_s_parallel_postulate \ \ A B C D T. (Bet A D T \ Bet B D C \ A \ D) \ (\ X Y. Bet A B X \ Bet A C Y \ Bet X T Y)" definition euclid_5 :: "bool" ("Euclid5") where "euclid_5 \ \ P Q R S T U. (BetS P T Q \ BetS R T S \ BetS Q U R \ \ Col P Q S \ Cong P T Q T \ Cong R T S T) \ (\ I. BetS S Q I \ BetS P U I)" definition euclid_s_parallel_postulate :: "bool" ("EuclidSParallelPostulate") where "euclid_s_parallel_postulate \ \ A B C D P Q R. (B C OS A D \ SAMS A B C B C D \ A B C B C D SumA P Q R \ \ Bet P Q R) \ (\ Y. B Out A Y \ C Out D Y)" definition playfair_s_postulate :: "bool" ("PlayfairSPostulate") where "playfair_s_postulate \ \ A1 A2 B1 B2 C1 C2 P. (A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2) \ (Col C1 B1 B2 \ Col C2 B1 B2)" section "Propositions" subsection "Congruence properties" lemma cong_reflexivity: shows "Cong A B A B" using cong_inner_transitivity cong_pseudo_reflexivity by blast lemma cong_symmetry: assumes "Cong A B C D" shows "Cong C D A B" using assms cong_inner_transitivity cong_reflexivity by blast lemma cong_transitivity: assumes "Cong A B C D" and "Cong C D E F" shows "Cong A B E F" by (meson assms(1) assms(2) cong_inner_transitivity cong_pseudo_reflexivity) lemma cong_left_commutativity: assumes "Cong A B C D" shows "Cong B A C D" using assms cong_inner_transitivity cong_pseudo_reflexivity by blast lemma cong_right_commutativity: assumes "Cong A B C D" shows "Cong A B D C" using assms cong_left_commutativity cong_symmetry by blast lemma cong_3421: assumes "Cong A B C D" shows "Cong C D B A" using assms cong_left_commutativity cong_symmetry by blast lemma cong_4312: assumes "Cong A B C D" shows "Cong D C A B" using assms cong_left_commutativity cong_symmetry by blast lemma cong_4321: assumes "Cong A B C D" shows "Cong D C B A" using assms cong_3421 cong_left_commutativity by blast lemma cong_trivial_identity: shows "Cong A A B B" using cong_identity segment_construction by blast lemma cong_reverse_identity: assumes "Cong A A C D" shows "C = D" using assms cong_3421 cong_identity by blast lemma cong_commutativity: assumes "Cong A B C D" shows "Cong B A D C" using assms cong_3421 by blast lemma not_cong_2134: assumes " \ Cong A B C D" shows "\ Cong B A C D" using assms cong_left_commutativity by blast lemma not_cong_1243: assumes "\ Cong A B C D" shows "\ Cong A B D C" using assms cong_right_commutativity by blast lemma not_cong_2143: assumes "\ Cong A B C D" shows "\ Cong B A D C" using assms cong_commutativity by blast lemma not_cong_3412: assumes "\ Cong A B C D" shows "\ Cong C D A B" using assms cong_symmetry by blast lemma not_cong_4312: assumes "\ Cong A B C D" shows "\ Cong D C A B" using assms cong_3421 by blast lemma not_cong_3421: assumes "\ Cong A B C D" shows "\ Cong C D B A" using assms cong_4312 by blast lemma not_cong_4321: assumes "\ Cong A B C D" shows "\ Cong D C B A" using assms cong_4321 by blast lemma five_segment_with_def: assumes "A B C D OFSC A' B' C' D'" and "A \ B" shows "Cong C D C' D'" using assms(1) assms(2) OFSC_def five_segment by blast lemma cong_diff: assumes "A \ B" and "Cong A B C D" shows "C \ D" using assms(1) assms(2) cong_identity by blast lemma cong_diff_2: assumes "B \ A" and "Cong A B C D" shows "C \ D" using assms(1) assms(2) cong_identity by blast lemma cong_diff_3: assumes "C \ D" and "Cong A B C D" shows "A \ B" using assms(1) assms(2) cong_reverse_identity by blast lemma cong_diff_4: assumes "D \ C" and "Cong A B C D" shows "A \ B" using assms(1) assms(2) cong_reverse_identity by blast lemma cong_3_sym: assumes "A B C Cong3 A' B' C'" shows "A' B' C' Cong3 A B C" using assms Cong3_def not_cong_3412 by blast lemma cong_3_swap: assumes "A B C Cong3 A' B' C'" shows "B A C Cong3 B' A' C'" using assms Cong3_def cong_commutativity by blast lemma cong_3_swap_2: assumes "A B C Cong3 A' B' C'" shows "A C B Cong3 A' C' B'" using assms Cong3_def cong_commutativity by blast lemma cong3_transitivity: assumes "A0 B0 C0 Cong3 A1 B1 C1" and "A1 B1 C1 Cong3 A2 B2 C2" shows "A0 B0 C0 Cong3 A2 B2 C2" by (meson assms(1) assms(2) Cong3_def cong_inner_transitivity not_cong_3412) lemma eq_dec_points: shows "A = B \ \ A = B" by simp lemma distinct: assumes "P \ Q" shows "R \ P \ R \ Q" using assms by simp lemma l2_11: assumes "Bet A B C" and "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "Cong A C A' C'" by (smt assms(1) assms(2) assms(3) assms(4) cong_right_commutativity cong_symmetry cong_trivial_identity five_segment) lemma bet_cong3: assumes "Bet A B C" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C'" by (meson assms(1) assms(2) Cong3_def l2_11 not_cong_3412 segment_construction) lemma construction_uniqueness: assumes "Q \ A" and "Bet Q A X" and "Cong A X B C" and "Bet Q A Y" and "Cong A Y B C" shows "X = Y" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) cong_identity cong_inner_transitivity cong_reflexivity five_segment) lemma Cong_cases: assumes "Cong A B C D \ Cong A B D C \ Cong B A C D \ Cong B A D C \ Cong C D A B \ Cong C D B A \ Cong D C A B \ Cong D C B A" shows "Cong A B C D" using assms not_cong_3421 not_cong_4321 by blast lemma Cong_perm : assumes "Cong A B C D" shows "Cong A B C D \ Cong A B D C \ Cong B A C D \ Cong B A D C \ Cong C D A B \ Cong C D B A \ Cong D C A B \ Cong D C B A" using assms not_cong_1243 not_cong_3412 by blast subsection "Betweeness properties" lemma bet_col: assumes "Bet A B C" shows "Col A B C" by (simp add: assms Col_def) lemma between_trivial: shows "Bet A B B" using cong_identity segment_construction by blast lemma between_symmetry: assumes "Bet A B C" shows "Bet C B A" using assms between_identity between_trivial inner_pasch by blast lemma Bet_cases: assumes "Bet A B C \ Bet C B A" shows "Bet A B C" using assms between_symmetry by blast lemma Bet_perm: assumes "Bet A B C" shows "Bet A B C \ Bet C B A" using assms Bet_cases by blast lemma between_trivial2: shows "Bet A A B" using Bet_perm between_trivial by blast lemma between_equality: assumes "Bet A B C" and "Bet B A C" shows "A = B" using assms(1) assms(2) between_identity inner_pasch by blast lemma between_equality_2: assumes "Bet A B C" and "Bet A C B" shows "B = C" using assms(1) assms(2) between_equality between_symmetry by blast lemma between_exchange3: assumes "Bet A B C" and "Bet A C D" shows "Bet B C D" by (metis Bet_perm assms(1) assms(2) between_identity inner_pasch) lemma bet_neq12__neq: assumes "Bet A B C" and "A \ B" shows "A \ C" using assms(1) assms(2) between_identity by blast lemma bet_neq21__neq: assumes "Bet A B C" and "B \ A" shows "A \ C" using assms(1) assms(2) between_identity by blast lemma bet_neq23__neq: assumes "Bet A B C" and "B \ C" shows "A \ C" using assms(1) assms(2) between_identity by blast lemma bet_neq32__neq: assumes "Bet A B C" and "C \ B" shows "A \ C" using assms(1) assms(2) between_identity by blast lemma not_bet_distincts: assumes "\ Bet A B C" shows "A \ B \ B \ C" using assms between_trivial between_trivial2 by blast lemma between_inner_transitivity: assumes "Bet A B D" and "Bet B C D" shows "Bet A B C" using assms(1) assms(2) Bet_perm between_exchange3 by blast lemma outer_transitivity_between2: assumes "Bet A B C" and "Bet B C D" and "B \ C" shows "Bet A C D" proof - obtain X where "Bet A C X \ Cong C X C D" using segment_construction by blast thus ?thesis using assms(1) assms(2) assms(3) between_exchange3 cong_inner_transitivity construction_uniqueness by blast qed lemma between_exchange2: assumes "Bet A B D" and "Bet B C D" shows "Bet A C D" using assms(1) assms(2) between_inner_transitivity outer_transitivity_between2 by blast lemma outer_transitivity_between: assumes "Bet A B C" and "Bet B C D" and "B \ C" shows "Bet A B D" using assms(1) assms(2) assms(3) between_symmetry outer_transitivity_between2 by blast lemma between_exchange4: assumes "Bet A B C" and "Bet A C D" shows "Bet A B D" using assms(1) assms(2) between_exchange2 between_symmetry by blast lemma l3_9_4: assumes "Bet4 A1 A2 A3 A4" shows "Bet4 A4 A3 A2 A1" using assms Bet4_def Bet_cases by blast lemma l3_17: assumes "Bet A B C" and "Bet A' B' C" and "Bet A P A'" shows "(\ Q. Bet P Q C \ Bet B Q B')" proof - obtain X where "Bet B' X A \ Bet P X C" using Bet_perm assms(2) assms(3) inner_pasch by blast moreover then obtain Y where "Bet X Y C \ Bet B Y B'" using Bet_perm assms(1) inner_pasch by blast ultimately show ?thesis using between_exchange2 by blast qed lemma lower_dim_ex: "\ A B C. \ (Bet A B C \ Bet B C A \ Bet C A B)" using lower_dim by auto lemma two_distinct_points: "\ X::'p. \ Y::'p. X \ Y" using lower_dim_ex not_bet_distincts by blast lemma point_construction_different: "\ C. Bet A B C \ B \ C" using Tarski_neutral_dimensionless.two_distinct_points Tarski_neutral_dimensionless_axioms cong_reverse_identity segment_construction by blast lemma another_point: "\ B::'p. A \ B" using point_construction_different by blast lemma Cong_stability: assumes "\ \ Cong A B C D" shows "Cong A B C D" using assms by simp lemma l2_11_b: assumes "Bet A B C" and "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "Cong A C A' C'" using assms(1) assms(2) assms(3) assms(4) l2_11 by auto lemma cong_dec_eq_dec_b: assumes "\ A \ B" shows "A = B" using assms(1) by simp lemma BetSEq: assumes "BetS A B C" shows "Bet A B C \ A \ B \ A \ C \ B \ C" using assms BetS_def between_identity by auto subsection "Collinearity" subsubsection "Collinearity and betweenness" lemma l4_2: assumes "A B C D IFSC A' B' C' D'" shows "Cong B D B' D'" proof cases assume "A = C" thus ?thesis by (metis IFSC_def Tarski_neutral_dimensionless.between_identity Tarski_neutral_dimensionless_axioms assms cong_diff_3) next assume H1: "A \ C" have H2: "Bet A B C \ Bet A' B' C' \ Cong A C A' C' \ Cong B C B' C' \ Cong A D A' D' \ Cong C D C' D'" using IFSC_def assms by auto obtain E where P1: "Bet A C E \ Cong C E A C" using segment_construction by blast have P1A: "Bet A C E" using P1 by simp have P1B: "Cong C E A C" using P1 by simp obtain E' where P2: "Bet A' C' E' \ Cong C' E' C E" using segment_construction by blast have P2A: "Bet A' C' E'" using P2 by simp have P2B: "Cong C' E' C E" using P2 by simp then have "Cong C E C' E'" using not_cong_3412 by blast then have "Cong E D E' D'" using H1 H2 P1 P2 five_segment by blast thus ?thesis by (smt H1 H2 P1A P1B P2A P2B Tarski_neutral_dimensionless.cong_commutativity Tarski_neutral_dimensionless.cong_diff_3 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms between_inner_transitivity between_symmetry five_segment) qed lemma l4_3: assumes "Bet A B C" and "Bet A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong A B A' B'" proof - have "A B C A IFSC A' B' C' A'" using IFSC_def assms(1) assms(2) assms(3) assms(4) cong_trivial_identity not_cong_2143 by blast thus ?thesis using l4_2 not_cong_2143 by blast qed lemma l4_3_1: assumes "Bet A B C" and "Bet A' B' C'" and "Cong A B A' B'" and "Cong A C A' C'" shows "Cong B C B' C'" by (meson assms(1) assms(2) assms(3) assms(4) between_symmetry cong_4321 l4_3) lemma l4_5: assumes "Bet A B C" and "Cong A C A' C'" shows "\ B'. (Bet A' B' C' \ A B C Cong3 A' B' C')" proof - obtain X' where P1: "Bet C' A' X' \ A' \ X'" using point_construction_different by auto obtain B' where P2: "Bet X' A' B' \ Cong A' B' A B" using segment_construction by blast obtain C'' where P3: "Bet X' B' C'' \ Cong B' C'' B C" using segment_construction by blast then have P4: "Bet A' B' C''" using P2 between_exchange3 by blast then have "C'' = C'" by (smt P1 P2 P3 assms(1) assms(2) between_exchange4 between_symmetry cong_symmetry construction_uniqueness l2_11_b) then show ?thesis by (smt Cong3_def P1 P2 P3 Tarski_neutral_dimensionless.construction_uniqueness Tarski_neutral_dimensionless_axioms P4 assms(1) assms(2) between_exchange4 between_symmetry cong_commutativity cong_symmetry cong_trivial_identity five_segment not_bet_distincts) qed lemma l4_6: assumes "Bet A B C" and "A B C Cong3 A' B' C'" shows "Bet A' B' C'" proof - obtain x where P1: "Bet A' x C' \ A B C Cong3 A' x C'" using Cong3_def assms(1) assms(2) l4_5 by blast then have "A' x C' Cong3 A' B' C'" using assms(2) cong3_transitivity cong_3_sym by blast then have "A' x C' x IFSC A' x C' B'" by (meson Cong3_def Cong_perm IFSC_def P1 cong_reflexivity) then have "Cong x x x B'" using l4_2 by auto then show ?thesis using P1 cong_reverse_identity by blast qed lemma cong3_bet_eq: assumes "Bet A B C" and "A B C Cong3 A X C" shows "X = B" proof - have "A B C B IFSC A B C X" by (meson Cong3_def Cong_perm IFSC_def assms(1) assms(2) cong_reflexivity) then show ?thesis using cong_reverse_identity l4_2 by blast qed subsubsection "Collinearity" lemma col_permutation_1: assumes "Col A B C" shows "Col B C A" using assms(1) Col_def by blast lemma col_permutation_2: assumes "Col A B C" shows "Col C A B" using assms(1) col_permutation_1 by blast lemma col_permutation_3: assumes "Col A B C" shows "Col C B A" using assms(1) Bet_cases Col_def by auto lemma col_permutation_4: assumes "Col A B C" shows "Col B A C" using assms(1) Bet_perm Col_def by blast lemma col_permutation_5: assumes "Col A B C" shows "Col A C B" using assms(1) col_permutation_1 col_permutation_3 by blast lemma not_col_permutation_1: assumes "\ Col A B C" shows "\ Col B C A" using assms col_permutation_2 by blast lemma not_col_permutation_2: assumes "~ Col A B C" shows "~ Col C A B" using assms col_permutation_1 by blast lemma not_col_permutation_3: assumes "\ Col A B C" shows "\ Col C B A" using assms col_permutation_3 by blast lemma not_col_permutation_4: assumes "\ Col A B C" shows "\ Col B A C" using assms col_permutation_4 by blast lemma not_col_permutation_5: assumes "\ Col A B C" shows "\ Col A C B" using assms col_permutation_5 by blast lemma Col_cases: assumes "Col A B C \ Col A C B \ Col B A C \ Col B C A \ Col C A B \ Col C B A" shows "Col A B C" using assms not_col_permutation_4 not_col_permutation_5 by blast lemma Col_perm: assumes "Col A B C" shows "Col A B C \ Col A C B \ Col B A C \ Col B C A \ Col C A B \ Col C B A" using Col_cases assms by blast lemma col_trivial_1: "Col A A B" using bet_col not_bet_distincts by blast lemma col_trivial_2: "Col A B B" by (simp add: Col_def between_trivial2) lemma col_trivial_3: "Col A B A" by (simp add: Col_def between_trivial2) lemma l4_13: assumes "Col A B C" and "A B C Cong3 A' B' C'" shows "Col A' B' C'" by (metis Tarski_neutral_dimensionless.Col_def Tarski_neutral_dimensionless.cong_3_swap Tarski_neutral_dimensionless.cong_3_swap_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) l4_6) lemma l4_14R1: assumes "Bet A B C" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C'" by (simp add: assms(1) assms(2) bet_cong3) lemma l4_14R2: assumes "Bet B C A" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C'" by (meson assms(1) assms(2) between_symmetry cong_3_swap_2 l4_5) lemma l4_14R3: assumes "Bet C A B" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C'" by (meson assms(1) assms(2) between_symmetry cong_3_swap l4_14R1 not_cong_2143) lemma l4_14: assumes "Col A B C" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C'" using Col_def assms(1) assms(2) l4_14R1 l4_14R2 l4_14R3 by blast lemma l4_16R1: assumes "A B C D FSC A' B' C' D'" and "A \ B" and "Bet A B C" shows "Cong C D C' D'" proof - have "A B C Cong3 A' B' C'" using FSC_def assms(1) by blast then have "Bet A' B' C'" using assms(3) l4_6 by blast then have "A B C D OFSC A' B' C' D'" by (meson Cong3_def FSC_def OFSC_def assms(1) cong_3_sym l4_6) thus ?thesis using assms(2) five_segment_with_def by blast qed lemma l4_16R2: assumes "A B C D FSC A' B' C' D'" and "Bet B C A" shows "Cong C D C' D'" proof - have "A B C Cong3 A' B' C'" using FSC_def assms(1) by blast then have "Bet B' C' A'" using Bet_perm assms(2) cong_3_swap_2 l4_6 by blast then have "B C A D IFSC B' C' A' D'" by (meson Cong3_def FSC_def IFSC_def assms(1) assms(2) not_cong_2143) then show ?thesis using l4_2 by auto qed lemma l4_16R3: assumes "A B C D FSC A' B' C' D'" and "A \ B" and "Bet C A B" shows "Cong C D C' D'" proof - have "A B C Cong3 A' B' C'" using FSC_def assms(1) by blast then have "Bet C' A' B'" using assms(3) between_symmetry cong_3_swap l4_6 by blast thus ?thesis by (smt Cong3_def FSC_def assms(1) assms(2) assms(3) between_symmetry cong_commutativity five_segment) qed lemma l4_16: assumes "A B C D FSC A' B' C' D'" and "A \ B" shows "Cong C D C' D'" by (meson Col_def FSC_def assms(1) assms(2) l4_16R1 l4_16R2 l4_16R3) lemma l4_17: assumes "A \ B" and "Col A B C" and "Cong A P A Q" and "Cong B P B Q" shows "Cong C P C Q" proof - { assume "\ Bet B C A" then have "\p pa. Bet p pa C \ Cong pa P pa Q \ Cong p P p Q \ p \ pa" using Col_def assms(1) assms(2) assms(3) assms(4) between_symmetry by blast then have ?thesis using cong_reflexivity five_segment by blast } then show ?thesis by (meson IFSC_def assms(3) assms(4) cong_reflexivity l4_2) qed lemma l4_18: assumes "A \ B" and "Col A B C" and "Cong A C A C'" and "Cong B C B C'" shows "C = C'" using assms(1) assms(2) assms(3) assms(4) cong_diff_3 l4_17 by blast lemma l4_19: assumes "Bet A C B" and "Cong A C A C'" and "Cong B C B C'" shows "C = C'" by (metis Col_def assms(1) assms(2) assms(3) between_equality between_trivial cong_identity l4_18 not_cong_3421) lemma not_col_distincts: assumes "\ Col A B C" shows "\ Col A B C \ A \ B \ B \ C \ A \ C" using Col_def assms between_trivial by blast lemma NCol_cases: assumes "\ Col A B C \ \ Col A C B \ \ Col B A C \ \ Col B C A \ \ Col C A B \ \ Col C B A" shows "\ Col A B C" using assms not_col_permutation_2 not_col_permutation_3 by blast lemma NCol_perm: assumes "\ Col A B C" shows "\ Col A B C \ ~ Col A C B \ ~ Col B A C \ ~ Col B C A \ ~ Col C A B \ ~ Col C B A" using NCol_cases assms by blast lemma col_cong_3_cong_3_eq: assumes "A \ B" and "Col A B C" and "A B C Cong3 A' B' C1" and "A B C Cong3 A' B' C2" shows "C1 = C2" by (metis Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless.cong_diff Tarski_neutral_dimensionless.l4_18 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong_inner_transitivity l4_13) subsection "Between transitivity le" lemma l5_1: assumes "A \ B" and "Bet A B C" and "Bet A B D" shows "Bet A C D \ Bet A D C" proof - obtain C' where P1: "Bet A D C' \ Cong D C' C D" using segment_construction by blast obtain D' where P2: "Bet A C D' \ Cong C D' C D" using segment_construction by blast obtain B' where P3: "Bet A C' B' \ Cong C' B' C B" using segment_construction by blast obtain B'' where P4: "Bet A D' B'' \ Cong D' B'' D B" using segment_construction by blast then have P5: "Cong B C' B'' C" by (smt P1 P2 assms(3) between_exchange3 between_symmetry cong_4312 cong_inner_transitivity l2_11_b) then have "Cong B B' B'' B" by (meson Bet_cases P1 P2 P3 P4 assms(2) assms(3) between_exchange4 between_inner_transitivity l2_11_b) then have P6: "B'' = B'" by (meson P1 P2 P3 P4 assms(1) assms(2) assms(3) between_exchange4 cong_inner_transitivity construction_uniqueness not_cong_2134) have "Bet B C D'" using P2 assms(2) between_exchange3 by blast then have "B C D' C' FSC B' C' D C" by (smt Cong3_def FSC_def P1 P2 P3 P5 P6 bet_col between_exchange3 between_symmetry cong_3421 cong_pseudo_reflexivity cong_transitivity l2_11_b) then have P8: "Cong D' C' D C" using P3 P4 P6 cong_identity l4_16 by blast obtain E where P9: "Bet C E C' \ Bet D E D'" using P1 P2 between_trivial2 l3_17 by blast then have P10: "D E D' C IFSC D E D' C'" by (smt IFSC_def P1 P2 P8 Tarski_neutral_dimensionless.cong_reflexivity Tarski_neutral_dimensionless_axioms cong_3421 cong_inner_transitivity) then have "Cong E C E C'" using l4_2 by auto have P11: "C E C' D IFSC C E C' D'" by (smt IFSC_def P1 P2 Tarski_neutral_dimensionless.cong_reflexivity Tarski_neutral_dimensionless_axioms P8 P9 cong_3421 cong_inner_transitivity) then have "Cong E D E D'" using l4_2 by auto obtain P where "Bet C' C P \ Cong C P C D'" using segment_construction by blast obtain R where "Bet D' C R \ Cong C R C E" using segment_construction by blast obtain Q where "Bet P R Q \ Cong R Q R P" using segment_construction by blast have "D' C R P FSC P C E D'" by (meson Bet_perm Cong3_def FSC_def \Bet C E C' \ Bet D E D'\ \Bet C' C P \ Cong C P C D'\ \Bet D' C R \ Cong C R C E\ bet_col between_exchange3 cong_pseudo_reflexivity l2_11_b not_cong_4321) have "Cong R P E D'" by (metis Cong_cases \D' C R P FSC P C E D'\ \Bet C' C P \ Cong C P C D'\ \Bet D' C R \ Cong C R C E\ cong_diff_2 l4_16) have "Cong R Q E D" by (metis Cong_cases \Cong E D E D'\ \Cong R P E D'\ \Bet P R Q \ Cong R Q R P\ cong_transitivity) have "D' E D C FSC P R Q C" by (meson Bet_perm Cong3_def FSC_def \Cong R P E D'\ \Cong R Q E D\ \Bet C E C' \ Bet D E D'\ \Bet C' C P \ Cong C P C D'\ \Bet D' C R \ Cong C R C E\ \Bet P R Q \ Cong R Q R P\ bet_col l2_11_b not_cong_2143 not_cong_4321) have "Cong D C Q C" using \D' E D C FSC P R Q C\ \Cong E D E D'\ \Bet C E C' \ Bet D E D'\ cong_identity l4_16 l4_16R2 by blast have "Cong C P C Q" using P2 \Cong D C Q C\ \Bet C' C P \ Cong C P C D'\ cong_right_commutativity cong_transitivity by blast have "Bet A C D \ Bet A D C" proof cases assume "R = C" then show ?thesis by (metis P1 \Cong E C E C'\ \Bet D' C R \ Cong C R C E\ cong_diff_4) next assume "R \ C" { have "Cong D' P D' Q" proof - have "Col R C D'" by (simp add: \Bet D' C R \ Cong C R C E\ bet_col between_symmetry) have "Cong R P R Q" by (metis Tarski_neutral_dimensionless.Cong_cases Tarski_neutral_dimensionless_axioms \Bet P R Q \ Cong R Q R P\) have "Cong C P C Q" by (simp add: \Cong C P C Q\) then show ?thesis using \Col R C D'\ \Cong R P R Q\ \R \ C\ l4_17 by blast qed then have "Cong B P B Q" using \Cong C P C Q\ \Bet B C D'\ cong_diff_4 by (metis Col_def \Bet C' C P \ Cong C P C D'\ cong_reflexivity l4_17 not_cong_3412) have "Cong B' P B' Q" by (metis P2 P4 \B'' = B'\ \Cong C P C Q\ \Cong D' P D' Q\ \Bet C' C P \ Cong C P C D'\ between_exchange3 cong_diff_4 cong_identity cong_reflexivity five_segment) have "Cong C' P C' Q" proof - have "Bet B C' B'" using P1 P3 assms(3) between_exchange3 between_exchange4 by blast then show ?thesis by (metis Col_def \Cong B P B Q\ \Cong B' P B' Q\ between_equality l4_17 not_bet_distincts) qed have "Cong P P P Q" by (metis Tarski_neutral_dimensionless.cong_diff_2 Tarski_neutral_dimensionless_axioms \Cong C P C Q\ \Cong C' P C' Q\ \R \ C\ \Bet C E C' \ Bet D E D'\ \Bet C' C P \ Cong C P C D'\ \Bet D' C R \ Cong C R C E\ bet_col bet_neq12__neq l4_17) thus ?thesis by (metis P2 \Cong R P E D'\ \Cong R Q E D\ \Bet P R Q \ Cong R Q R P\ bet_neq12__neq cong_diff_4) } then have "R \ C \ Bet A C D \ Bet A D C" by blast qed thus ?thesis by simp qed lemma l5_2: assumes "A \ B" and "Bet A B C" and "Bet A B D" shows "Bet B C D \ Bet B D C" using assms(1) assms(2) assms(3) between_exchange3 l5_1 by blast lemma segment_construction_2: assumes "A \ Q" shows "\ X. ((Bet Q A X \ Bet Q X A) \ Cong Q X B C)" proof - obtain A' where P1: "Bet A Q A' \ Cong Q A' A Q" using segment_construction by blast obtain X where P2: "Bet A' Q X \ Cong Q X B C" using segment_construction by blast then show ?thesis by (metis P1 Tarski_neutral_dimensionless.cong_diff_4 Tarski_neutral_dimensionless_axioms between_symmetry l5_2) qed lemma l5_3: assumes "Bet A B D" and "Bet A C D" shows "Bet A B C \ Bet A C B" by (metis Bet_perm assms(1) assms(2) between_inner_transitivity l5_2 point_construction_different) lemma bet3__bet: assumes "Bet A B E" and "Bet A D E" and "Bet B C D" shows "Bet A C E" by (meson assms(1) assms(2) assms(3) between_exchange2 between_symmetry l5_3) lemma le_bet: assumes "C D Le A B" shows "\ X. (Bet A X B \ Cong A X C D)" by (meson Le_def assms cong_symmetry) lemma l5_5_1: assumes "A B Le C D" shows "\ X. (Bet A B X \ Cong A X C D)" proof - obtain P where P1: "Bet C P D \ Cong A B C P" using Le_def assms by blast obtain X where P2: "Bet A B X \ Cong B X P D" using segment_construction by blast then have "Cong A X C D" using P1 l2_11_b by blast then show ?thesis using P2 by blast qed lemma l5_5_2: assumes "\ X. (Bet A B X \ Cong A X C D)" shows "A B Le C D" proof - obtain P where P1: "Bet A B P \ Cong A P C D" using assms by blast obtain B' where P2: "Bet C B' D \ A B P Cong3 C B' D" using P1 l4_5 by blast then show ?thesis using Cong3_def Le_def by blast qed lemma l5_6: assumes "A B Le C D" and "Cong A B A' B'" and "Cong C D C' D'" shows "A' B' Le C' D'" by (meson Cong3_def Le_def assms(1) assms(2) assms(3) cong_inner_transitivity l4_5) lemma le_reflexivity: shows "A B Le A B" using between_trivial cong_reflexivity l5_5_2 by blast lemma le_transitivity: assumes "A B Le C D" and "C D Le E F" shows "A B Le E F" by (meson assms(1) assms(2) between_exchange4 cong_reflexivity l5_5_1 l5_5_2 l5_6 le_bet) lemma between_cong: assumes "Bet A C B" and "Cong A C A B" shows "C = B" by (smt assms(1) assms(2) between_trivial cong_inner_transitivity five_segment l4_19 l4_3_1) lemma cong3_symmetry: assumes "A B C Cong3 A' B' C'" shows "A' B' C' Cong3 A B C" by (simp add: assms cong_3_sym) lemma between_cong_2: assumes "Bet A D B" and "Bet A E B" and "Cong A D A E" shows "D = E" using l5_3 by (smt Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) cong_diff cong_inner_transitivity l4_3_1) lemma between_cong_3: assumes "A \ B" and "Bet A B D" and "Bet A B E" and "Cong B D B E" shows "D = E" by (meson assms(1) assms(2) assms(3) assms(4) cong_reflexivity construction_uniqueness) lemma le_anti_symmetry: assumes "A B Le C D" and "C D Le A B" shows "Cong A B C D" by (smt Le_def Tarski_neutral_dimensionless.between_exchange4 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_neq21__neq between_cong between_exchange3 cong_transitivity l5_5_1 not_cong_3421) lemma cong_dec: shows "Cong A B C D \ \ Cong A B C D" by simp lemma bet_dec: shows "Bet A B C \ \ Bet A B C" by simp lemma col_dec: shows "Col A B C \ \ Col A B C" by simp lemma le_trivial: shows "A A Le C D" using Le_def between_trivial2 cong_trivial_identity by blast lemma le_cases: shows "A B Le C D \ C D Le A B" by (metis (full_types) cong_reflexivity l5_5_2 l5_6 not_bet_distincts segment_construction_2) lemma le_zero: assumes "A B Le C C" shows "A = B" by (metis assms cong_diff_4 le_anti_symmetry le_trivial) lemma le_diff: assumes "A \ B" and "A B Le C D" shows "C \ D" using assms(1) assms(2) le_zero by blast lemma lt_diff: assumes "A B Lt C D" shows "C \ D" using Lt_def assms cong_trivial_identity le_zero by blast lemma bet_cong_eq: assumes "Bet A B C" and "Bet A C D" and "Cong B C A D" shows "C = D \ A = B" proof - have "Bet C B A" using Bet_perm assms(1) by blast then show ?thesis by (metis (no_types) Cong_perm Le_def assms(2) assms(3) between_cong cong_pseudo_reflexivity le_anti_symmetry) qed lemma cong__le: assumes "Cong A B C D" shows "A B Le C D" using Le_def assms between_trivial by blast lemma cong__le3412: assumes "Cong A B C D" shows "C D Le A B" using assms cong__le cong_symmetry by blast lemma le1221: shows "A B Le B A" by (simp add: cong__le cong_pseudo_reflexivity) lemma le_left_comm: assumes "A B Le C D" shows "B A Le C D" using assms le1221 le_transitivity by blast lemma le_right_comm: assumes "A B Le C D" shows "A B Le D C" by (meson assms cong_right_commutativity l5_5_1 l5_5_2) lemma le_comm: assumes "A B Le C D" shows "B A Le D C" using assms le_left_comm le_right_comm by blast lemma ge_left_comm: assumes "A B Ge C D" shows "B A Ge C D" by (meson Ge_def assms le_right_comm) lemma ge_right_comm: assumes "A B Ge C D" shows "A B Ge D C" using Ge_def assms le_left_comm by presburger lemma ge_comm0: assumes "A B Ge C D" shows "B A Ge D C" by (meson assms ge_left_comm ge_right_comm) lemma lt_right_comm: assumes "A B Lt C D" shows "A B Lt D C" using Lt_def assms le_right_comm not_cong_1243 by blast lemma lt_left_comm: assumes "A B Lt C D" shows "B A Lt C D" using Lt_def assms le_comm lt_right_comm not_cong_2143 by blast lemma lt_comm: assumes "A B Lt C D" shows "B A Lt D C" using assms lt_left_comm lt_right_comm by blast lemma gt_left_comm0: assumes "A B Gt C D" shows "B A Gt C D" by (meson Gt_def assms lt_right_comm) lemma gt_right_comm: assumes "A B Gt C D" shows "A B Gt D C" using Gt_def assms lt_left_comm by presburger lemma gt_comm: assumes "A B Gt C D" shows "B A Gt D C" by (meson assms gt_left_comm0 gt_right_comm) lemma cong2_lt__lt: assumes "A B Lt C D" and "Cong A B A' B'" and "Cong C D C' D'" shows "A' B' Lt C' D'" by (meson Lt_def assms(1) assms(2) assms(3) l5_6 le_anti_symmetry not_cong_3412) lemma fourth_point: assumes "A \ B" and "B \ C" and "Col A B P" and "Bet A B C" shows "Bet P A B \ Bet A P B \ Bet B P C \ Bet B C P" by (metis Col_def Tarski_neutral_dimensionless.l5_2 Tarski_neutral_dimensionless_axioms assms(3) assms(4) between_symmetry) lemma third_point: assumes "Col A B P" shows "Bet P A B \ Bet A P B \ Bet A B P" using Col_def assms between_symmetry by blast lemma l5_12_a: assumes "Bet A B C" shows "A B Le A C \ B C Le A C" using assms between_symmetry cong_left_commutativity cong_reflexivity l5_5_2 le_left_comm by blast lemma bet__le1213: assumes "Bet A B C" shows "A B Le A C" using assms l5_12_a by blast lemma bet__le2313: assumes "Bet A B C" shows "B C Le A C" by (simp add: assms l5_12_a) lemma bet__lt1213: assumes "B \ C" and "Bet A B C" shows "A B Lt A C" using Lt_def assms(1) assms(2) bet__le1213 between_cong by blast lemma bet__lt2313: assumes "A \ B" and "Bet A B C" shows "B C Lt A C" using Lt_def assms(1) assms(2) bet__le2313 bet_cong_eq l5_1 by blast lemma l5_12_b: assumes "Col A B C" and "A B Le A C" and "B C Le A C" shows "Bet A B C" by (metis assms(1) assms(2) assms(3) between_cong col_permutation_5 l5_12_a le_anti_symmetry not_cong_2143 third_point) lemma bet_le_eq: assumes "Bet A B C" and "A C Le B C" shows "A = B" by (meson assms(1) assms(2) bet__le2313 bet_cong_eq l5_1 le_anti_symmetry) lemma or_lt_cong_gt: "A B Lt C D \ A B Gt C D \ Cong A B C D" by (meson Gt_def Lt_def cong_symmetry local.le_cases) lemma lt__le: assumes "A B Lt C D" shows "A B Le C D" using Lt_def assms by blast lemma le1234_lt__lt: assumes "A B Le C D" and "C D Lt E F" shows "A B Lt E F" by (meson Lt_def assms(1) assms(2) cong__le3412 le_anti_symmetry le_transitivity) lemma le3456_lt__lt: assumes "A B Lt C D" and "C D Le E F" shows "A B Lt E F" by (meson Lt_def assms(1) assms(2) cong2_lt__lt cong_reflexivity le1234_lt__lt) lemma lt_transitivity: assumes "A B Lt C D" and "C D Lt E F" shows "A B Lt E F" using Lt_def assms(1) assms(2) le1234_lt__lt by blast lemma not_and_lt: "\ (A B Lt C D \ C D Lt A B)" by (simp add: Lt_def le_anti_symmetry) lemma nlt: "\ A B Lt A B" using not_and_lt by blast lemma le__nlt: assumes "A B Le C D" shows "\ C D Lt A B" using assms le3456_lt__lt nlt by blast lemma cong__nlt: assumes "Cong A B C D" shows "\ A B Lt C D" by (simp add: Lt_def assms) lemma nlt__le: assumes "\ A B Lt C D" shows "C D Le A B" using Lt_def assms cong__le3412 local.le_cases by blast lemma lt__nle: assumes "A B Lt C D" shows "\ C D Le A B" using assms le__nlt by blast lemma nle__lt: assumes "\ A B Le C D" shows "C D Lt A B" using assms nlt__le by blast lemma lt1123: assumes "B \ C" shows "A A Lt B C" using assms le_diff nle__lt by blast lemma bet2_le2__le_R1: assumes "Bet a P b" and "Bet A Q B" and "P a Le Q A" and "P b Le Q B" and "B = Q" shows "a b Le A B" by (metis assms(3) assms(4) assms(5) le_comm le_diff) lemma bet2_le2__le_R2: assumes "Bet a Po b" and "Bet A PO B" and "Po a Le PO A" and "Po b Le PO B" and "A \ PO" and "B \ PO" shows "a b Le A B" proof - obtain b' where P1: "Bet A PO b' \ Cong PO b' b Po" using segment_construction by blast obtain a' where P2: "Bet B PO a' \ Cong PO a' a Po" using segment_construction by blast obtain a'' where P3: "Bet PO a'' A \ Cong Po a PO a''" using Le_def assms(3) by blast have P4: "a' = a''" by (meson Bet_cases P2 P3 assms(2) assms(6) between_inner_transitivity cong_right_commutativity construction_uniqueness not_cong_3412) have P5: "B a' Le B A" using Bet_cases P3 P4 assms(2) bet__le1213 between_exchange2 by blast obtain b'' where P6: "Bet PO b'' B \ Cong Po b PO b''" using Le_def assms(4) by blast then have "b' = b''" using P1 assms(2) assms(5) between_inner_transitivity cong_right_commutativity construction_uniqueness not_cong_3412 by blast then have "a' b' Le a' B" using Bet_cases P2 P6 bet__le1213 between_exchange2 by blast then have "a' b' Le A B" using P5 le_comm le_transitivity by blast thus ?thesis by (smt Cong_cases P1 P3 P4 Tarski_neutral_dimensionless.l5_6 Tarski_neutral_dimensionless_axioms assms(1) between_exchange3 between_symmetry cong_reflexivity l2_11_b) qed lemma bet2_le2__le: assumes "Bet a P b" and "Bet A Q B" and "P a Le Q A" and "P b Le Q B" shows "a b Le A B" proof cases assume "A = Q" thus ?thesis using assms(3) assms(4) le_diff by force next assume "\ A = Q" thus ?thesis using assms(1) assms(2) assms(3) assms(4) bet2_le2__le_R1 bet2_le2__le_R2 by blast qed lemma Le_cases: assumes "A B Le C D \ B A Le C D \ A B Le D C \ B A Le D C" shows "A B Le C D" using assms le_left_comm le_right_comm by blast lemma Lt_cases: assumes "A B Lt C D \ B A Lt C D \ A B Lt D C \ B A Lt D C" shows "A B Lt C D" using assms lt_comm lt_left_comm by blast subsection "Out lines" lemma bet_out: assumes "B \ A" and "Bet A B C" shows "A Out B C" using Out_def assms(1) assms(2) bet_neq12__neq by fastforce lemma bet_out_1: assumes "B \ A" and "Bet C B A" shows "A Out B C" by (simp add: assms(1) assms(2) bet_out between_symmetry) lemma out_dec: shows "P Out A B \ \ P Out A B" by simp lemma out_diff1: assumes "A Out B C" shows "B \ A" using Out_def assms by auto lemma out_diff2: assumes "A Out B C" shows "C \ A" using Out_def assms by auto lemma out_distinct: assumes "A Out B C" shows "B \ A \ C \ A" using assms out_diff1 out_diff2 by auto lemma out_col: assumes "A Out B C" shows "Col A B C" using Col_def Out_def assms between_symmetry by auto lemma l6_2: assumes "A \ P" and "B \ P" and "C \ P" and "Bet A P C" shows "Bet B P C \ P Out A B" by (smt Bet_cases Out_def assms(1) assms(2) assms(3) assms(4) between_inner_transitivity l5_2 outer_transitivity_between) lemma bet_out__bet: assumes "Bet A P C" and "P Out A B" shows "Bet B P C" by (metis Tarski_neutral_dimensionless.l6_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) not_bet_distincts out_diff1) lemma l6_3_1: assumes "P Out A B" shows "A \ P \ B \ P \ (\ C. (C \ P \ Bet A P C \ Bet B P C))" using assms bet_out__bet out_diff1 out_diff2 point_construction_different by fastforce lemma l6_3_2: assumes "A \ P" and "B \ P" and "\ C. (C \ P \ Bet A P C \ Bet B P C)" shows "P Out A B" using assms(1) assms(2) assms(3) l6_2 by blast lemma l6_4_1: assumes "P Out A B" and "Col A P B" shows "\ Bet A P B" using Out_def assms(1) between_equality between_symmetry by fastforce lemma l6_4_2: assumes "Col A P B" and "\ Bet A P B" shows "P Out A B" by (metis Out_def assms(1) assms(2) bet_out col_permutation_1 third_point) lemma out_trivial: assumes "A \ P" shows "P Out A A" by (simp add: assms bet_out_1 between_trivial2) lemma l6_6: assumes "P Out A B" shows "P Out B A" using Out_def assms by auto lemma l6_7: assumes "P Out A B" and "P Out B C" shows "P Out A C" by (smt Out_def assms(1) assms(2) between_exchange4 l5_1 l5_3) lemma bet_out_out_bet: assumes "Bet A B C" and "B Out A A'" and "B Out C C'" shows "Bet A' B C'" by (metis Out_def assms(1) assms(2) assms(3) bet_out__bet between_inner_transitivity outer_transitivity_between) lemma out2_bet_out: assumes "B Out A C" and "B Out X P" and "Bet A X C" shows "B Out A P \ B Out C P" by (smt Out_def Tarski_neutral_dimensionless.l6_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) between_exchange2 between_symmetry) lemma l6_11_uniqueness: assumes "A Out X R" and "Cong A X B C" and "A Out Y R" and "Cong A Y B C" shows "X = Y" by (metis Out_def assms(1) assms(2) assms(3) assms(4) between_cong cong_symmetry cong_transitivity l6_6 l6_7) lemma l6_11_existence: assumes "R \ A" and "B \ C" shows "\ X. (A Out X R \ Cong A X B C)" by (metis Out_def assms(1) assms(2) cong_reverse_identity segment_construction_2) lemma segment_construction_3: assumes "A \ B" and "X \ Y" shows "\ C. (A Out B C \ Cong A C X Y)" by (metis assms(1) assms(2) l6_11_existence l6_6) lemma l6_13_1: assumes "P Out A B" and "P A Le P B" shows "Bet P A B" by (metis Out_def assms(1) assms(2) bet__lt1213 le__nlt) lemma l6_13_2: assumes "P Out A B" and "Bet P A B" shows "P A Le P B" by (simp add: assms(2) bet__le1213) lemma l6_16_1: assumes "P \ Q" and "Col S P Q" and "Col X P Q" shows "Col X P S" by (smt Col_def assms(1) assms(2) assms(3) bet3__bet col_permutation_4 l5_1 l5_3 outer_transitivity_between third_point) lemma col_transitivity_1: assumes "P \ Q" and "Col P Q A" and "Col P Q B" shows "Col P A B" by (meson Tarski_neutral_dimensionless.l6_16_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) not_col_permutation_2) lemma col_transitivity_2: assumes "P \ Q" and "Col P Q A" and "Col P Q B" shows "Col Q A B" by (metis Tarski_neutral_dimensionless.col_transitivity_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) not_col_permutation_4) lemma l6_21: assumes "\ Col A B C" and "C \ D" and "Col A B P" and "Col A B Q" and "Col C D P" and "Col C D Q" shows "P = Q" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) col_transitivity_1 l6_16_1 not_col_distincts) lemma col2__eq: assumes "Col A X Y" and "Col B X Y" and "\ Col A X B" shows "X = Y" using assms(1) assms(2) assms(3) l6_16_1 by blast lemma not_col_exists: assumes "A \ B" shows "\ C. \ Col A B C" by (metis Col_def assms col_transitivity_2 lower_dim_ex) lemma col3: assumes "X \ Y" and "Col X Y A" and "Col X Y B" and "Col X Y C" shows "Col A B C" by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_2) lemma colx: assumes "A \ B" and "Col X Y A" and "Col X Y B" and "Col A B C" shows "Col X Y C" by (metis assms(1) assms(2) assms(3) assms(4) l6_21 not_col_distincts) lemma out2__bet: assumes "A Out B C" and "C Out A B" shows "Bet A B C" by (metis Out_def assms(1) assms(2) between_equality between_symmetry) lemma bet2_le2__le1346: assumes "Bet A B C" and "Bet A' B' C'" and "A B Le A' B'" and "B C Le B' C'" shows "A C Le A' C'" using Le_cases assms(1) assms(2) assms(3) assms(4) bet2_le2__le by blast lemma bet2_le2__le2356_R1: assumes "Bet A A C" and "Bet A' B' C'" and "A A Le A' B'" and "A' C' Le A C" shows "B' C' Le A C" using assms(2) assms(4) bet__le2313 le3456_lt__lt lt__nle nlt__le by blast lemma bet2_le2__le2356_R2: assumes "A \ B" and "Bet A B C" and "Bet A' B' C'" and "A B Le A' B'" and "A' C' Le A C" shows "B' C' Le B C" proof - have "A \ C" using assms(1) assms(2) bet_neq12__neq by blast obtain B0 where P1: "Bet A B B0 \ Cong A B0 A' B'" using assms(4) l5_5_1 by blast then have P2: "A \ B0" using assms(1) bet_neq12__neq by blast obtain C0 where P3: "Bet A C0 C \ Cong A' C' A C0" using Le_def assms(5) by blast then have "A \ C0" using assms(1) assms(3) assms(4) bet_neq12__neq cong_diff le_diff by blast then have P4: "Bet A B0 C0" by (smt Out_def P1 P2 P3 assms(1) assms(2) assms(3) bet__le1213 between_exchange2 between_symmetry l5_1 l5_3 l5_6 l6_13_1 not_cong_3412) have K1: "B0 C0 Le B C0" using P1 P4 between_exchange3 l5_12_a by blast have K2: "B C0 Le B C" using P1 P3 P4 bet__le1213 between_exchange3 between_exchange4 by blast then have "Cong B0 C0 B' C'" using P1 P3 P4 assms(3) l4_3_1 not_cong_3412 by blast then show ?thesis by (meson K1 K2 cong__nlt le_transitivity nlt__le) qed lemma bet2_le2__le2356: assumes "Bet A B C" and "Bet A' B' C'" and "A B Le A' B'" and "A' C' Le A C" shows "B' C' Le B C" proof (cases) assume "A = B" then show ?thesis using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356_R1 by blast next assume "\ A = B" then show ?thesis using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356_R2 by blast qed lemma bet2_le2__le1245: assumes "Bet A B C" and "Bet A' B' C'" and "B C Le B' C'" and "A' C' Le A C" shows "A' B' Le A B" using assms(1) assms(2) assms(3) assms(4) bet2_le2__le2356 between_symmetry le_comm by blast lemma cong_preserves_bet: assumes "Bet B A' A0" and "Cong B A' E D'" and "Cong B A0 E D0" and "E Out D' D0" shows "Bet E D' D0" using Tarski_neutral_dimensionless.l6_13_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet__le1213 l5_6 by fastforce lemma out_cong_cong: assumes "B Out A A0" and "E Out D D0" and "Cong B A E D" and "Cong B A0 E D0" shows "Cong A A0 D D0" by (meson Out_def assms(1) assms(2) assms(3) assms(4) cong_4321 cong_symmetry l4_3_1 l5_6 l6_13_1 l6_13_2) lemma not_out_bet: assumes "Col A B C" and "\ B Out A C" shows "Bet A B C" using assms(1) assms(2) l6_4_2 by blast lemma or_bet_out: shows "Bet A B C \ B Out A C \ \ Col A B C" using not_out_bet by blast lemma not_bet_out: assumes "Col A B C" and "\ Bet A B C" shows "B Out A C" by (simp add: assms(1) assms(2) l6_4_2) lemma not_bet_and_out: shows "\ (Bet A B C \ B Out A C)" using bet_col l6_4_1 by blast lemma out_to_bet: assumes "Col A' B' C'" and "B Out A C \ B' Out A' C'" and "Bet A B C" shows "Bet A' B' C'" using assms(1) assms(2) assms(3) not_bet_and_out or_bet_out by blast lemma col_out2_col: assumes "Col A B C" and "B Out A AA" and "B Out C CC" shows "Col AA B CC" using l6_21 by (smt Tarski_neutral_dimensionless.out_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) col_trivial_2 not_col_permutation_1 out_diff1) lemma bet2_out_out: assumes "B \ A" and "B' \ A" and "A Out C C'" and "Bet A B C" and "Bet A B' C'" shows "A Out B B'" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) bet_out l6_6 l6_7) lemma bet2__out: assumes "A \ B" and "A \ B'" and "Bet A B C" and "Bet A B' C" shows "A Out B B'" using Out_def assms(1) assms(2) assms(3) assms(4) l5_3 by auto lemma out_bet_out_1: assumes "P Out A C" and "Bet A B C" shows "P Out A B" by (metis assms(1) assms(2) not_bet_and_out out2_bet_out out_trivial) lemma out_bet_out_2: assumes "P Out A C" and "Bet A B C" shows "P Out B C" using assms(1) assms(2) l6_6 l6_7 out_bet_out_1 by blast lemma out_bet__out: assumes "Bet P Q A" and "Q Out A B" shows "P Out A B" by (smt Out_def assms(1) assms(2) bet_out_1 bet_out__bet) lemma segment_reverse: assumes "Bet A B C " shows "\ B'. Bet A B' C \ Cong C B' A B" by (metis Bet_perm Cong_perm assms bet_cong_eq cong_reflexivity segment_construction_2) lemma diff_col_ex: shows "\ C. A \ C \ B \ C \ Col A B C" by (metis bet_col bet_neq12__neq point_construction_different) lemma diff_bet_ex3: assumes "Bet A B C" shows "\ D. A \ D \ B \ D \ C \ D \ Col A B D" by (metis (mono_tags, opaque_lifting) Col_def bet_out_1 between_trivial2 col_transitivity_1 l6_4_1 point_construction_different) lemma diff_col_ex3: assumes "Col A B C" shows "\ D. A \ D \ B \ D \ C \ D \ Col A B D" by (metis Bet_perm Col_def between_equality between_trivial2 point_construction_different) lemma Out_cases: assumes "A Out B C \ A Out C B" shows "A Out B C" using assms l6_6 by blast subsection "Midpoint" lemma midpoint_dec: "I Midpoint A B \ \ I Midpoint A B" by simp lemma is_midpoint_id: assumes "A Midpoint A B" shows "A = B" using Midpoint_def assms between_cong by blast lemma is_midpoint_id_2: assumes "A Midpoint B A" shows "A = B" using Midpoint_def assms cong_diff_2 by blast lemma l7_2: assumes "M Midpoint A B" shows "M Midpoint B A" using Bet_perm Cong_perm Midpoint_def assms by blast lemma l7_3: assumes "M Midpoint A A" shows "M = A" using Midpoint_def assms bet_neq23__neq by blast lemma l7_3_2: "A Midpoint A A" by (simp add: Midpoint_def between_trivial2 cong_reflexivity) lemma symmetric_point_construction: "\ P'. A Midpoint P P'" by (meson Midpoint_def cong__le cong__le3412 le_anti_symmetry segment_construction) lemma symmetric_point_uniqueness: assumes "P Midpoint A P1" and "P Midpoint A P2" shows "P1 = P2" by (metis Midpoint_def assms(1) assms(2) between_cong_3 cong_diff_4 cong_inner_transitivity) lemma l7_9: assumes "A Midpoint P X" and "A Midpoint Q X" shows "P = Q" using assms(1) assms(2) l7_2 symmetric_point_uniqueness by blast lemma l7_9_bis: assumes "A Midpoint P X" and "A Midpoint X Q" shows "P = Q" using assms(1) assms(2) l7_2 symmetric_point_uniqueness by blast lemma l7_13_R1: assumes "A \ P" and "A Midpoint P' P" and "A Midpoint Q' Q" shows "Cong P Q P' Q'" proof - obtain X where P1: "Bet P' P X \ Cong P X Q A" using segment_construction by blast obtain X' where P2: "Bet X P' X' \ Cong P' X' Q A" using segment_construction by blast obtain Y where P3: "Bet Q' Q Y \ Cong Q Y P A" using segment_construction by blast obtain Y' where P4: "Bet Y Q' Y' \ Cong Q' Y' P A" using segment_construction by blast have P5: "Bet Y A Q'" by (meson Midpoint_def P3 P4 assms(3) bet3__bet between_symmetry l5_3) have P6: "Bet P' A X" using Midpoint_def P1 assms(2) between_exchange4 by blast have P7: "Bet A P X" using Midpoint_def P1 assms(2) between_exchange3 by blast have P8: "Bet Y Q A" using Midpoint_def P3 assms(3) between_exchange3 between_symmetry by blast have P9: "Bet A Q' Y'" using P4 P5 between_exchange3 by blast have P10: "Bet X' P' A" using P2 P6 between_exchange3 between_symmetry by blast have P11: "Bet X A X'" using P10 P2 P6 between_symmetry outer_transitivity_between2 by blast have P12: "Bet Y A Y'" using P4 P5 between_exchange4 by blast have P13: "Cong A X Y A" using P1 P3 P7 P8 l2_11_b not_cong_4321 by blast have P14: "Cong A Y' X' A" proof - have Q1: "Cong Q' Y' P' A" using Midpoint_def P4 assms(2) cong_transitivity not_cong_3421 by blast have "Cong A Q' X' P'" by (meson Midpoint_def P2 assms(3) cong_transitivity not_cong_3421) then show ?thesis using P10 P9 Q1 l2_11_b by blast qed have P15: "Cong A Y A Y'" proof - have "Cong Q Y Q' Y'" using P3 P4 cong_transitivity not_cong_3412 by blast then show ?thesis using Bet_perm Cong_perm Midpoint_def P8 P9 assms(3) l2_11_b by blast qed have P16: "Cong X A Y' A" using Cong_cases P13 P15 cong_transitivity by blast have P17: "Cong A X' A Y" using P14 P15 cong_transitivity not_cong_3421 by blast have P18: "X A X' Y' FSC Y' A Y X" proof - have Q3: "Col X A X'" by (simp add: Col_def P11) have "Cong X X' Y' Y" using Bet_cases P11 P12 P16 P17 l2_11_b by blast then show ?thesis by (simp add: Cong3_def FSC_def P16 P17 Q3 cong_4321 cong_pseudo_reflexivity) qed then have "Y Q A X IFSC Y' Q' A X'" by (smt IFSC_def Midpoint_def P14 P15 P16 P7 P8 P9 assms(1) assms(3) bet_neq12__neq between_symmetry cong_4321 cong_inner_transitivity cong_right_commutativity l4_16) then have "X P A Q IFSC X' P' A Q'" by (meson IFSC_def Midpoint_def P10 P7 assms(2) between_symmetry cong_4312 l4_2) then show ?thesis using l4_2 by auto qed lemma l7_13: assumes "A Midpoint P' P" and "A Midpoint Q' Q" shows "Cong P Q P' Q'" proof (cases) assume "A = P" then show ?thesis using Midpoint_def assms(1) assms(2) cong_3421 is_midpoint_id_2 by blast next show ?thesis by (metis Tarski_neutral_dimensionless.l7_13_R1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) cong_trivial_identity is_midpoint_id_2 not_cong_2143) qed lemma l7_15: assumes "A Midpoint P P'" and "A Midpoint Q Q'" and "A Midpoint R R'" and "Bet P Q R" shows "Bet P' Q' R'" proof - have "P Q R Cong3 P' Q' R'" using Cong3_def assms(1) assms(2) assms(3) l7_13 l7_2 by blast then show ?thesis using assms(4) l4_6 by blast qed lemma l7_16: assumes "A Midpoint P P'" and "A Midpoint Q Q'" and "A Midpoint R R'" and "A Midpoint S S'" and "Cong P Q R S" shows "Cong P' Q' R' S'" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) cong_transitivity l7_13 not_cong_3412) lemma symmetry_preserves_midpoint: assumes "Z Midpoint A D" and "Z Midpoint B E" and "Z Midpoint C F" and "B Midpoint A C" shows "E Midpoint D F" by (meson Midpoint_def assms(1) assms(2) assms(3) assms(4) l7_15 l7_16) lemma Mid_cases: assumes "A Midpoint B C \ A Midpoint C B" shows "A Midpoint B C" using assms l7_2 by blast lemma Mid_perm: assumes "A Midpoint B C" shows "A Midpoint B C \ A Midpoint C B" by (simp add: assms l7_2) lemma l7_17: assumes "A Midpoint P P'" and "B Midpoint P P'" shows "A = B" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. p Midpoint pa (pp p pa)" by (meson symmetric_point_construction) then have "\p pa. Bet pa p (pp p pa)" by (meson Midpoint_def) then have f2: "\p. Bet p p p" by (meson between_inner_transitivity) have f3: "\p pa. Bet (pp pa p) pa p" using f1 Mid_perm Midpoint_def by blast have f4: "\p. pp p p = p" using f2 f1 by (metis Midpoint_def bet_cong_eq) have f5: "Bet (pp P P') P B" using f3 by (meson Midpoint_def assms(2) between_inner_transitivity) have f6: "A Midpoint P' P" using Mid_perm assms(1) by blast have f7: "Bet (pp P P') P A" using f3 Midpoint_def assms(1) between_inner_transitivity by blast have f8: "Bet P' A P" using f6 by (simp add: Midpoint_def) have "Cong P' A A P" using f6 Midpoint_def by blast then have "P' = P \ A = B" using f8 by (metis (no_types) Midpoint_def assms(2) bet_cong_eq between_inner_transitivity l5_2) then show ?thesis using f7 f6 f5 f4 f1 by (metis (no_types) Col_perm Mid_perm assms(2) bet_col l4_18 l5_2 l7_13) qed lemma l7_17_bis: assumes "A Midpoint P P'" and "B Midpoint P' P" shows "A = B" by (meson Tarski_neutral_dimensionless.l7_17 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2)) lemma l7_20: assumes "Col A M B" and "Cong M A M B" shows "A = B \ M Midpoint A B" by (metis Bet_cases Col_def Midpoint_def assms(1) assms(2) between_cong cong_left_commutativity not_cong_3412) lemma l7_20_bis: assumes "A \ B" and "Col A M B" and "Cong M A M B" shows "M Midpoint A B" using assms(1) assms(2) assms(3) l7_20 by blast lemma cong_col_mid: assumes "A \ C" and "Col A B C" and "Cong A B B C" shows "B Midpoint A C" using assms(1) assms(2) assms(3) cong_left_commutativity l7_20 by blast lemma l7_21_R1: assumes "\ Col A B C" and "B \ D" and "Cong A B C D" and "Cong B C D A" and "Col A P C" and "Col B P D" shows "P Midpoint A C" proof - obtain X where P1: "B D P Cong3 D B X" using Col_perm assms(6) cong_pseudo_reflexivity l4_14 by blast have P2: "Col D B X" using P1 assms(6) l4_13 not_col_permutation_5 by blast have P3: "B D P A FSC D B X C" using FSC_def P1 assms(3) assms(4) assms(6) not_col_permutation_5 not_cong_2143 not_cong_3412 by blast have P4: "B D P C FSC D B X A" by (simp add: FSC_def P1 assms(3) assms(4) assms(6) col_permutation_5 cong_4321) have "A P C Cong3 C X A" using Cong3_def Cong_perm P3 P4 assms(2) cong_pseudo_reflexivity l4_16 by blast then show ?thesis by (smt Cong3_def NCol_cases P2 assms(1) assms(2) assms(5) assms(6) colx cong_col_mid l4_13 not_col_distincts not_col_permutation_1 not_cong_1243) qed lemma l7_21: assumes "\ Col A B C" and "B \ D" and "Cong A B C D" and "Cong B C D A" and "Col A P C" and "Col B P D" shows "P Midpoint A C \ P Midpoint B D" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) col_transitivity_2 is_midpoint_id_2 l7_21_R1 not_col_distincts not_cong_3412) lemma l7_22_aux_R1: assumes "Bet A1 C C" and "Bet B1 C B2" and "Cong C A1 C B1" and "Cong C C C B2" and "M1 Midpoint A1 B1" and "M2 Midpoint A2 B2"and "C A1 Le C C" shows "Bet M1 C M2" by (metis assms(3) assms(5) assms(7) cong_diff_3 l7_3 le_diff not_bet_distincts) lemma l7_22_aux_R2: assumes "A2 \ C" and "Bet A1 C A2" and "Bet B1 C B2" and "Cong C A1 C B1" and "Cong C A2 C B2" and "M1 Midpoint A1 B1" and "M2 Midpoint A2 B2" and "C A1 Le C A2" shows "Bet M1 C M2" proof - obtain X where P1: "C Midpoint A2 X" using symmetric_point_construction by blast obtain X0 where P2: "C Midpoint B2 X0" using symmetric_point_construction by blast obtain X1 where P3: "C Midpoint M2 X1" using symmetric_point_construction by blast have P4: "X1 Midpoint X X0" using P1 P2 P3 assms(7) symmetry_preserves_midpoint by blast have P5: "C A1 Le C X" using Cong_perm Midpoint_def P1 assms(8) cong_reflexivity l5_6 by blast have P6: "Bet C A1 X" by (smt Midpoint_def P1 P5 assms(1) assms(2) bet2__out between_symmetry is_midpoint_id_2 l5_2 l6_13_1) have P7: "C B1 Le C X0" proof - have Q1: "Cong C A1 C B1" by (simp add: assms(4)) have "Cong C X C X0" using P1 P2 assms(5) l7_16 l7_3_2 by blast then show ?thesis using P5 Q1 l5_6 by blast qed have P8: "Bet C B1 X0" by (smt Midpoint_def P2 P7 assms(1) assms(3) assms(5) bet2__out between_symmetry cong_identity l5_2 l6_13_1) obtain Q where P9: "Bet X1 Q C \ Bet A1 Q B1" by (meson Bet_perm Midpoint_def P4 P6 P8 l3_17) have P10: "X A1 C X1 IFSC X0 B1 C X1" by (smt Cong_perm IFSC_def Midpoint_def P1 P2 P4 P6 P8 assms(4) assms(5) between_symmetry cong_reflexivity l7_16 l7_3_2) have P11: "Cong A1 X1 B1 X1" using P10 l4_2 by blast have P12: "Cong Q A1 Q B1" proof (cases) assume "C = X1" then show ?thesis using P9 assms(4) bet_neq12__neq by blast next assume Q1: "\ C = X1" have Q2: "Col C X1 Q" using Col_def P9 by blast have Q3: "Cong C A1 C B1" by (simp add: assms(4)) have "Cong X1 A1 X1 B1" using P11 not_cong_2143 by blast then show ?thesis using Q1 Q2 Q3 l4_17 by blast qed have P13: "Q Midpoint A1 B1" by (simp add: Midpoint_def P12 P9 cong_left_commutativity) then show ?thesis using Midpoint_def P3 P9 assms(6) between_inner_transitivity between_symmetry l7_17 by blast qed lemma l7_22_aux: assumes "Bet A1 C A2" and "Bet B1 C B2" and "Cong C A1 C B1" and "Cong C A2 C B2" and "M1 Midpoint A1 B1" and "M2 Midpoint A2 B2" and "C A1 Le C A2" shows "Bet M1 C M2" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l7_22_aux_R1 l7_22_aux_R2) lemma l7_22: assumes "Bet A1 C A2" and "Bet B1 C B2" and "Cong C A1 C B1" and "Cong C A2 C B2" and "M1 Midpoint A1 B1" and "M2 Midpoint A2 B2" shows "Bet M1 C M2" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) between_symmetry l7_22_aux local.le_cases) lemma bet_col1: assumes "Bet A B D" and "Bet A C D" shows "Col A B C" using Bet_perm Col_def assms(1) assms(2) l5_3 by blast lemma l7_25_R1: assumes "Cong C A C B" and "Col A B C" shows "\ X. X Midpoint A B" using assms(1) assms(2) l7_20 l7_3_2 not_col_permutation_5 by blast lemma l7_25_R2: assumes "Cong C A C B" and "\ Col A B C" shows "\ X. X Midpoint A B" proof - obtain P where P1: "Bet C A P \ A \ P" using point_construction_different by auto obtain Q where P2: "Bet C B Q \ Cong B Q A P" using segment_construction by blast obtain R where P3: "Bet A R Q \ Bet B R P" using P1 P2 between_symmetry inner_pasch by blast obtain X where P4: "Bet A X B \ Bet R X C" using P1 P3 inner_pasch by blast have "Cong X A X B" proof - have Q1: "Cong R A R B \ Cong X A X B" proof (cases) assume "R = C" then show ?thesis using P4 bet_neq12__neq by blast next assume Q2: "\ R = C" have "Col R C X" using Col_perm P4 bet_col by blast then show ?thesis using Q2 assms(1) l4_17 by blast qed have "Cong R A R B" proof - have Q3: "C A P B OFSC C B Q A" by (simp add: OFSC_def P1 P2 assms(1) cong_pseudo_reflexivity cong_symmetry) have Q4: "Cong P B Q A" using Q3 assms(2) five_segment_with_def not_col_distincts by blast obtain R' where Q5: "Bet A R' Q \ B R P Cong3 A R' Q" using Cong_perm P3 Q4 l4_5 by blast have Q6: "B R P A IFSC A R' Q B" by (meson Cong3_def IFSC_def OFSC_def P3 Q3 Q5 not_cong_2143) have Q7: "B R P Q IFSC A R' Q P" using IFSC_def P2 Q6 cong_pseudo_reflexivity by auto have Q8: "Cong R A R' B" using Q6 l4_2 by auto have Q9: "Cong R Q R' P" using Q7 l4_2 by auto have Q10: "A R Q Cong3 B R' P" using Cong3_def Q4 Q8 Q9 cong_commutativity not_cong_4321 by blast have Q11: "Col B R' P" using P3 Q10 bet_col l4_13 by blast have "R = R'" proof - have R1: "B \ P" using P1 assms(1) between_cong by blast then have R2: "A \ Q" using Q4 cong_diff_2 by blast have R3: "B \ Q" using P1 P2 cong_diff_3 by blast then have R4: "B \ R" by (metis Cong3_def P1 Q11 Q5 assms(2) bet_col cong_diff_3 l6_21 not_col_distincts) have R5: "\ Col A Q B" by (metis P2 R3 assms(2) bet_col col_permutation_3 col_trivial_2 l6_21) have R6: "B \ P" by (simp add: R1) have R7: "Col A Q R" using NCol_cases P3 bet_col by blast have R8: "Col A Q R'" using Q5 bet_col col_permutation_5 by blast have R9: "Col B P R" using NCol_cases P3 bet_col by blast have "Col B P R'" using Col_perm Q11 by blast then show ?thesis using R5 R6 R7 R8 R9 l6_21 by blast qed then show ?thesis using Q8 by blast qed then show ?thesis using Q1 by blast qed then show ?thesis using P4 assms(2) bet_col l7_20_bis not_col_distincts by blast qed lemma l7_25: assumes "Cong C A C B" shows "\ X. X Midpoint A B" using assms l7_25_R1 l7_25_R2 by blast lemma midpoint_distinct_1: assumes "A \ B" and "I Midpoint A B" shows "I \ A \ I \ B" using assms(1) assms(2) is_midpoint_id is_midpoint_id_2 by blast lemma midpoint_distinct_2: assumes "I \ A" and "I Midpoint A B" shows "A \ B \ I \ B" using assms(1) assms(2) is_midpoint_id_2 l7_3 by blast lemma midpoint_distinct_3: assumes "I \ B" and "I Midpoint A B" shows "A \ B \ I \ A" using assms(1) assms(2) is_midpoint_id l7_3 by blast lemma midpoint_def: assumes "Bet A B C" and "Cong A B B C" shows "B Midpoint A C" using Midpoint_def assms(1) assms(2) by blast lemma midpoint_bet: assumes "B Midpoint A C" shows "Bet A B C" using Midpoint_def assms by blast lemma midpoint_col: assumes "M Midpoint A B" shows "Col M A B" using assms bet_col col_permutation_4 midpoint_bet by blast lemma midpoint_cong: assumes "B Midpoint A C" shows "Cong A B B C" using Midpoint_def assms by blast lemma midpoint_out: assumes "A \ C" and "B Midpoint A C" shows "A Out B C" using assms(1) assms(2) bet_out midpoint_bet midpoint_distinct_1 by blast lemma midpoint_out_1: assumes "A \ C" and "B Midpoint A C" shows "C Out A B" by (metis Tarski_neutral_dimensionless.midpoint_bet Tarski_neutral_dimensionless.midpoint_distinct_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_out_1 l6_6) lemma midpoint_not_midpoint: assumes "A \ B" and "I Midpoint A B" shows "\ B Midpoint A I" using assms(1) assms(2) between_equality_2 midpoint_bet midpoint_distinct_1 by blast lemma swap_diff: assumes "A \ B" shows "B \ A" using assms by auto lemma cong_cong_half_1: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "Cong A B A' B'" shows "Cong A M A' M'" proof - obtain M'' where P1: "Bet A' M'' B' \ A M B Cong3 A' M'' B'" using assms(1) assms(3) l4_5 midpoint_bet by blast have P2: "M'' Midpoint A' B'" by (meson Cong3_def P1 assms(1) cong_inner_transitivity midpoint_cong midpoint_def) have P3: "M' = M''" using P2 assms(2) l7_17 by auto then show ?thesis using Cong3_def P1 by blast qed lemma cong_cong_half_2: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "Cong A B A' B'" shows "Cong B M B' M'" using assms(1) assms(2) assms(3) cong_cong_half_1 l7_2 not_cong_2143 by blast lemma cong_mid2__cong: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "Cong A M A' M'" shows "Cong A B A' B'" by (meson assms(1) assms(2) assms(3) cong_inner_transitivity l2_11_b midpoint_bet midpoint_cong) lemma mid__lt: assumes "A \ B" and "M Midpoint A B" shows "A M Lt A B" using assms(1) assms(2) bet__lt1213 midpoint_bet midpoint_distinct_1 by blast lemma le_mid2__le13: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "A M Le A' M'" shows "A B Le A' B'" by (smt Tarski_neutral_dimensionless.cong_mid2__cong Tarski_neutral_dimensionless.l7_13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) bet2_le2__le2356 l5_6 l7_3_2 le_anti_symmetry le_comm local.le_cases midpoint_bet) lemma le_mid2__le12: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "A B Le A' B'" shows "A M Le A' M'" by (meson assms(1) assms(2) assms(3) cong__le3412 cong_cong_half_1 le_anti_symmetry le_mid2__le13 local.le_cases) lemma lt_mid2__lt13: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "A M Lt A' M'" shows "A B Lt A' B'" by (meson Tarski_neutral_dimensionless.le_mid2__le12 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) lt__nle nlt__le) lemma lt_mid2__lt12: assumes "M Midpoint A B" and "M' Midpoint A' B'" and "A B Lt A' B'" shows "A M Lt A' M'" by (meson Tarski_neutral_dimensionless.le_mid2__le13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) le__nlt nle__lt) lemma midpoint_preserves_out: assumes "A Out B C" and "M Midpoint A A'" and "M Midpoint B B'" and "M Midpoint C C'" shows "A' Out B' C'" by (smt Out_def assms(1) assms(2) assms(3) assms(4) l6_4_2 l7_15 l7_2 not_bet_and_out not_col_distincts) lemma col_cong_bet: assumes "Col A B D" and "Cong A B C D" and "Bet A C B" shows "Bet C A D \ Bet C B D" by (smt Col_def assms(1) assms(2) assms(3) bet_cong_eq between_inner_transitivity col_transitivity_2 cong_4321 l6_2 not_bet_and_out not_cong_4312 third_point) lemma col_cong2_bet1: assumes "Col A B D" and "Bet A C B" and "Cong A B C D" and "Cong A C B D" shows "Bet C B D" by (metis assms(1) assms(2) assms(3) assms(4) bet__le1213 bet_cong_eq between_symmetry col_cong_bet cong__le cong_left_commutativity l5_12_b l5_6 outer_transitivity_between2) lemma col_cong2_bet2: assumes "Col A B D" and "Bet A C B" and "Cong A B C D" and "Cong A D B C" shows "Bet C A D" by (metis assms(1) assms(2) assms(3) assms(4) bet_cong_eq col_cong_bet cong_identity not_bet_distincts not_cong_3421 outer_transitivity_between2) lemma col_cong2_bet3: assumes "Col A B D" and "Bet A B C" and "Cong A B C D" and "Cong A C B D" shows "Bet B C D" by (metis assms(1) assms(2) assms(3) assms(4) bet__le1213 bet__le2313 bet_col col_transitivity_2 cong_diff_3 cong_reflexivity l5_12_b l5_6 not_bet_distincts) lemma col_cong2_bet4: assumes "Col A B C" and "Bet A B D" and "Cong A B C D" and "Cong A D B C" shows "Bet B D C" using assms(1) assms(2) assms(3) assms(4) col_cong2_bet3 cong_right_commutativity by blast lemma col_bet2_cong1: assumes "Col A B D" and "Bet A C B" and "Cong A B C D" and "Bet C B D" shows "Cong A C D B" by (meson assms(2) assms(3) assms(4) between_symmetry cong_pseudo_reflexivity cong_right_commutativity l4_3) lemma col_bet2_cong2: assumes "Col A B D" and "Bet A C B" and "Cong A B C D" and "Bet C A D" shows "Cong D A B C" by (meson assms(2) assms(3) assms(4) between_symmetry cong_commutativity cong_pseudo_reflexivity cong_symmetry l4_3) lemma bet2_lt2__lt: assumes "Bet a Po b" and "Bet A PO B" and "Po a Lt PO A" and "Po b Lt PO B" shows "a b Lt A B" by (metis Lt_cases Tarski_neutral_dimensionless.nle__lt Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet2_le2__le1245 le__nlt lt__le) lemma bet2_lt_le__lt: assumes "Bet a Po b" and "Bet A PO B" and "Cong Po a PO A" and "Po b Lt PO B" shows "a b Lt A B" by (smt Lt_def Tarski_neutral_dimensionless.l4_3_1 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) bet2_le2__le cong__le not_cong_2143) subsection "Orthogonality" lemma per_dec: "Per A B C \ \ Per A B C" by simp lemma l8_2: assumes "Per A B C" shows "Per C B A" proof - obtain C' where P1: "B Midpoint C C' \ Cong A C A C'" using Per_def assms by blast obtain A' where P2: "B Midpoint A A'" using symmetric_point_construction by blast have "Cong C' A C A'" using Mid_perm P1 P2 l7_13 by blast thus ?thesis using P1 P2 Per_def cong_4321 cong_inner_transitivity by blast qed lemma Per_cases: assumes "Per A B C \ Per C B A" shows "Per A B C" using assms l8_2 by blast lemma Per_perm : assumes "Per A B C" shows "Per A B C \ Per C B A" by (simp add: assms l8_2) lemma l8_3 : assumes "Per A B C" and "A \ B" and "Col B A A'" shows "Per A' B C" by (smt Per_def assms(1) assms(2) assms(3) l4_17 l7_13 l7_2 l7_3_2) lemma l8_4: assumes "Per A B C" and "B Midpoint C C'" shows "Per A B C'" by (metis Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) l8_3 midpoint_col midpoint_distinct_1) lemma l8_5: "Per A B B" using Per_def cong_reflexivity l7_3_2 by blast lemma l8_6: assumes "Per A B C" and "Per A' B C" and "Bet A C A'" shows "B = C" by (metis Per_def assms(1) assms(2) assms(3) l4_19 midpoint_distinct_3 symmetric_point_uniqueness) lemma l8_7: assumes "Per A B C" and "Per A C B" shows "B = C" proof - obtain C' where P1: "B Midpoint C C' \ Cong A C A C'" using Per_def assms(1) by blast obtain A' where P2: "C Midpoint A A'" using Per_def assms(2) l8_2 by blast have "Per C' C A" by (metis P1 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(2) bet_col l8_2 midpoint_bet midpoint_distinct_3) then have "Cong A C' A' C'" using Cong_perm P2 Per_def symmetric_point_uniqueness by blast then have "Cong A' C A' C'" using P1 P2 cong_inner_transitivity midpoint_cong not_cong_2134 by blast then have Q4: "Per A' B C" using P1 Per_def by blast have "Bet A' C A" using Mid_perm P2 midpoint_bet by blast thus ?thesis using Q4 assms(1) l8_6 by blast qed lemma l8_8: assumes "Per A B A" shows "A = B" using Tarski_neutral_dimensionless.l8_6 Tarski_neutral_dimensionless_axioms assms between_trivial2 by fastforce lemma per_distinct: assumes "Per A B C" and "A \ B" shows "A \ C" using assms(1) assms(2) l8_8 by blast lemma per_distinct_1: assumes "Per A B C" and "B \ C" shows "A \ C" using assms(1) assms(2) l8_8 by blast lemma l8_9: assumes "Per A B C" and "Col A B C" shows "A = B \ C = B" using Col_cases assms(1) assms(2) l8_3 l8_8 by blast lemma l8_10: assumes "Per A B C" and "A B C Cong3 A' B' C'" shows "Per A' B' C'" proof - obtain D where P1: "B Midpoint C D \ Cong A C A D" using Per_def assms(1) by blast obtain D' where P2: "Bet C' B' D' \ Cong B' D' B' C'" using segment_construction by blast have P3: "B' Midpoint C' D'" by (simp add: Midpoint_def P2 cong_4312) have "Cong A' C' A' D'" proof (cases) assume "C = B" thus ?thesis by (metis Cong3_def P3 assms(2) cong_diff_4 cong_reflexivity is_midpoint_id) next assume Q1: "\ C = B" have "C B D A OFSC C' B' D' A'" by (metis Cong3_def OFSC_def P1 P3 Tarski_neutral_dimensionless.cong_mid2__cong Tarski_neutral_dimensionless_axioms assms(2) cong_commutativity l4_3_1 midpoint_bet) thus ?thesis by (meson OFSC_def P1 Q1 cong_4321 cong_inner_transitivity five_segment_with_def) qed thus ?thesis using Per_def P3 by blast qed lemma col_col_per_per: assumes "A \ X" and "C \ X" and "Col U A X" and "Col V C X" and "Per A X C" shows "Per U X V" by (meson Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) not_col_permutation_3) lemma perp_in_dec: "X PerpAt A B C D \ \ X PerpAt A B C D" by simp lemma perp_distinct: assumes "A B Perp C D" shows "A \ B \ C \ D" using PerpAt_def Perp_def assms by auto lemma l8_12: assumes "X PerpAt A B C D" shows "X PerpAt C D A B" using Per_perm PerpAt_def assms by auto lemma per_col: assumes "B \ C" and "Per A B C" and "Col B C D" shows "Per A B D" by (metis Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) l8_2) lemma l8_13_2: assumes "A \ B" and "C \ D" and "Col X A B" and "Col X C D" and "\ U. \ V. Col U A B \ Col V C D \ U \ X \ V \ X \ Per U X V" shows "X PerpAt A B C D" proof - obtain pp :: 'p and ppa :: 'p where f1: "Col pp A B \ Col ppa C D \ pp \ X \ ppa \ X \ Per pp X ppa" using assms(5) by blast obtain ppb :: "'p \ 'p \ 'p \ 'p \ 'p \ 'p" and ppc :: "'p \ 'p \ 'p \ 'p \ 'p \ 'p" where "\x0 x1 x2 x3 x4. (\v5 v6. (Col v5 x3 x2 \ Col v6 x1 x0) \ \ Per v5 x4 v6) = ((Col (ppb x0 x1 x2 x3 x4) x3 x2 \ Col (ppc x0 x1 x2 x3 x4) x1 x0) \ \ Per (ppb x0 x1 x2 x3 x4) x4 (ppc x0 x1 x2 x3 x4))" by moura then have f2: "\p pa pb pc pd. (\ p PerpAt pa pb pc pd \ pa \ pb \ pc \ pd \ Col p pa pb \ Col p pc pd \ (\pe pf. (\ Col pe pa pb \ \ Col pf pc pd) \ Per pe p pf)) \ (p PerpAt pa pb pc pd \ pa = pb \ pc = pd \ \ Col p pa pb \ \ Col p pc pd \ (Col (ppb pd pc pb pa p) pa pb \ Col (ppc pd pc pb pa p) pc pd) \ \ Per (ppb pd pc pb pa p) p (ppc pd pc pb pa p))" using PerpAt_def by fastforce { assume "\ Col (ppb D C B A X) pp X" then have "\ Col (ppb D C B A X) A B \ \ Col (ppc D C B A X) C D \ Per (ppb D C B A X) X (ppc D C B A X)" using f1 by (meson assms(1) assms(3) col3 not_col_permutation_2) } moreover { assume "\ Col (ppc D C B A X) ppa X" then have "\ Col (ppb D C B A X) A B \ \ Col (ppc D C B A X) C D \ Per (ppb D C B A X) X (ppc D C B A X)" using f1 by (meson assms(2) assms(4) col3 not_col_permutation_2) } ultimately have "\ Col (ppb D C B A X) A B \ \ Col (ppc D C B A X) C D \ Per (ppb D C B A X) X (ppc D C B A X)" using f1 by (meson Tarski_neutral_dimensionless.col_col_per_per Tarski_neutral_dimensionless_axioms) then have "(X PerpAt A B C D \ A = B \ C = D \ \ Col X A B \ \ Col X C D \ Col (ppb D C B A X) A B \ Col (ppc D C B A X) C D \ \ Per (ppb D C B A X) X (ppc D C B A X)) \ (\ Col (ppb D C B A X) A B \ \ Col (ppc D C B A X) C D \ Per (ppb D C B A X) X (ppc D C B A X))" using f2 by presburger thus ?thesis using assms(1) assms(2) assms(3) assms(4) by blast qed lemma l8_14_1: "\ A B Perp A B" by (metis PerpAt_def Perp_def Tarski_neutral_dimensionless.col_trivial_1 Tarski_neutral_dimensionless.col_trivial_3 Tarski_neutral_dimensionless_axioms l8_8) lemma l8_14_2_1a: assumes "X PerpAt A B C D" shows "A B Perp C D" using Perp_def assms by blast lemma perp_in_distinct: assumes "X PerpAt A B C D" shows "A \ B \ C \ D" using PerpAt_def assms by blast lemma l8_14_2_1b: assumes "X PerpAt A B C D" and "Col Y A B" and "Col Y C D" shows "X = Y" by (metis PerpAt_def assms(1) assms(2) assms(3) l8_13_2 l8_14_1 l8_14_2_1a) lemma l8_14_2_1b_bis: assumes "A B Perp C D" and "Col X A B" and "Col X C D" shows "X PerpAt A B C D" using Perp_def assms(1) assms(2) assms(3) l8_14_2_1b by blast lemma l8_14_2_2: assumes "A B Perp C D" and "\ Y. (Col Y A B \ Col Y C D) \ X = Y" shows "X PerpAt A B C D" by (metis Tarski_neutral_dimensionless.PerpAt_def Tarski_neutral_dimensionless.Perp_def Tarski_neutral_dimensionless_axioms assms(1) assms(2)) lemma l8_14_3: assumes "X PerpAt A B C D" and "Y PerpAt A B C D" shows "X = Y" by (meson PerpAt_def assms(1) assms(2) l8_14_2_1b) lemma l8_15_1: assumes "Col A B X" and "A B Perp C X" shows "X PerpAt A B C X" using NCol_perm assms(1) assms(2) col_trivial_3 l8_14_2_1b_bis by blast lemma l8_15_2: assumes "Col A B X" and "X PerpAt A B C X" shows "A B Perp C X" using assms(2) l8_14_2_1a by blast lemma perp_in_per: assumes "B PerpAt A B B C" shows "Per A B C" by (meson NCol_cases PerpAt_def assms col_trivial_3) lemma perp_sym: assumes "A B Perp A B" shows "C D Perp C D" using assms l8_14_1 by auto lemma perp_col0: assumes "A B Perp C D" and "X \ Y" and "Col A B X" and "Col A B Y" shows "C D Perp X Y" proof - obtain X0 where P1: "X0 PerpAt A B C D" using Perp_def assms(1) by blast then have P2: " A \ B \ C \ D \ Col X0 A B \ Col X0 C D \ ((Col U A B \ Col V C D) \ Per U X0 V)" using PerpAt_def by blast have Q1: "C \ D" using P2 by blast have Q2: "X \ Y" using assms(2) by blast have Q3: "Col X0 C D" using P2 by blast have Q4: "Col X0 X Y" proof - have "\p pa. Col p pa Y \ Col p pa X \ Col p pa X0 \ p \ pa" by (metis (no_types) Col_cases P2 assms(3) assms(4)) thus ?thesis using col3 by blast qed have "X0 PerpAt C D X Y" proof - have "\ U V. (Col U C D \ Col V X Y) \ Per U X0 V" by (metis Col_perm P1 Per_perm Q2 Tarski_neutral_dimensionless.PerpAt_def Tarski_neutral_dimensionless_axioms assms(3) assms(4) colx) thus ?thesis using Q1 Q2 Q3 Q4 PerpAt_def by blast qed thus ?thesis using Perp_def by auto qed lemma per_perp_in: assumes "A \ B" and "B \ C" and "Per A B C" shows "B PerpAt A B B C" by (metis Col_def assms(1) assms(2) assms(3) between_trivial2 l8_13_2) lemma per_perp: assumes "A \ B" and "B \ C" and "Per A B C" shows "A B Perp B C" using Perp_def assms(1) assms(2) assms(3) per_perp_in by blast lemma perp_left_comm: assumes "A B Perp C D" shows "B A Perp C D" proof - obtain X where "X PerpAt A B C D" using Perp_def assms by blast then have "X PerpAt B A C D" using PerpAt_def col_permutation_5 by auto thus ?thesis using Perp_def by blast qed lemma perp_right_comm: assumes "A B Perp C D" shows "A B Perp D C" by (meson Perp_def assms l8_12 perp_left_comm) lemma perp_comm: assumes "A B Perp C D" shows "B A Perp D C" by (simp add: assms perp_left_comm perp_right_comm) lemma perp_in_sym: assumes "X PerpAt A B C D" shows "X PerpAt C D A B" by (simp add: assms l8_12) lemma perp_in_left_comm: assumes "X PerpAt A B C D" shows "X PerpAt B A C D" by (metis Col_cases PerpAt_def assms) lemma perp_in_right_comm: assumes "X PerpAt A B C D" shows "X PerpAt A B D C" using assms perp_in_left_comm perp_in_sym by blast lemma perp_in_comm: assumes "X PerpAt A B C D" shows "X PerpAt B A D C" by (simp add: assms perp_in_left_comm perp_in_right_comm) lemma Perp_cases: assumes "A B Perp C D \ B A Perp C D \ A B Perp D C \ B A Perp D C \ C D Perp A B \ C D Perp B A \ D C Perp A B \ D C Perp B A" shows "A B Perp C D" by (meson Perp_def assms perp_in_sym perp_left_comm) lemma Perp_perm : assumes "A B Perp C D" shows "A B Perp C D \ B A Perp C D \ A B Perp D C \ B A Perp D C \ C D Perp A B \ C D Perp B A \ D C Perp A B \ D C Perp B A" by (meson Perp_def assms perp_in_sym perp_left_comm) lemma Perp_in_cases: assumes "X PerpAt A B C D \ X PerpAt B A C D \ X PerpAt A B D C \ X PerpAt B A D C \ X PerpAt C D A B \ X PerpAt C D B A \ X PerpAt D C A B \ X PerpAt D C B A" shows "X PerpAt A B C D" using assms perp_in_left_comm perp_in_sym by blast lemma Perp_in_perm: assumes "X PerpAt A B C D" shows "X PerpAt A B C D \ X PerpAt B A C D \ X PerpAt A B D C \ X PerpAt B A D C \ X PerpAt C D A B \ X PerpAt C D B A \ X PerpAt D C A B \ X PerpAt D C B A" using Perp_in_cases assms by blast lemma perp_in_col: assumes "X PerpAt A B C D" shows "Col A B X \ Col C D X" using PerpAt_def assms col_permutation_2 by presburger lemma perp_perp_in: assumes "A B Perp C A" shows "A PerpAt A B C A" using assms l8_15_1 not_col_distincts by blast lemma perp_per_1: assumes "A B Perp C A" shows "Per B A C" using Perp_in_cases assms perp_in_per perp_perp_in by blast lemma perp_per_2: assumes "A B Perp A C" shows "Per B A C" by (simp add: Perp_perm assms perp_per_1) lemma perp_col: assumes "A \ E" and "A B Perp C D" and "Col A B E" shows "A E Perp C D" using Perp_perm assms(1) assms(2) assms(3) col_trivial_3 perp_col0 by blast lemma perp_col2: assumes "A B Perp X Y" and "C \ D" and "Col A B C" and "Col A B D" shows "C D Perp X Y" using Perp_perm assms(1) assms(2) assms(3) assms(4) perp_col0 by blast lemma perp_col4: assumes "P \ Q" and "R \ S" and "Col A B P" and "Col A B Q" and "Col C D R" and "Col C D S" and "A B Perp C D" shows "P Q Perp R S" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) perp_col0 by blast lemma perp_not_eq_1: assumes "A B Perp C D" shows "A \ B" using assms perp_distinct by auto lemma perp_not_eq_2: assumes "A B Perp C D" shows "C \ D" using assms perp_distinct by auto lemma diff_per_diff: assumes "A \ B" and "Cong A P B R" and "Per B A P" and "Per A B R" shows "P \ R" using assms(1) assms(3) assms(4) l8_2 l8_7 by blast lemma per_not_colp: assumes "A \ B" and "A \ P" and "B \ R" and "Per B A P" and "Per A B R" shows "\ Col P A R" by (metis Per_cases Tarski_neutral_dimensionless.col_permutation_4 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4) assms(5) l8_3 l8_7) lemma per_not_col: assumes "A \ B" and "B \ C" and "Per A B C" shows "\ Col A B C" using assms(1) assms(2) assms(3) l8_9 by auto lemma perp_not_col2: assumes "A B Perp C D" shows "\ Col A B C \ \ Col A B D" using assms l8_14_1 perp_col2 perp_distinct by blast lemma perp_not_col: assumes "A B Perp P A" shows "\ Col A B P" proof - have "A PerpAt A B P A" using assms perp_perp_in by auto then have "Per P A B" by (simp add: perp_in_per perp_in_sym) then have "\ Col B A P" by (metis NCol_perm Tarski_neutral_dimensionless.perp_not_eq_1 Tarski_neutral_dimensionless.perp_not_eq_2 Tarski_neutral_dimensionless_axioms assms per_not_col) thus ?thesis using Col_perm by blast qed lemma perp_in_col_perp_in: assumes "C \ E" and "Col C D E" and "P PerpAt A B C D" shows "P PerpAt A B C E" proof - have P2: "C \ D" using assms(3) perp_in_distinct by blast have P3: "Col P A B" using PerpAt_def assms(3) by auto have "Col P C D" using PerpAt_def assms(3) by blast then have "Col P C E" using P2 assms(2) col_trivial_2 colx by blast thus ?thesis by (smt P3 Perp_perm Tarski_neutral_dimensionless.l8_14_2_1b_bis Tarski_neutral_dimensionless.perp_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) l8_14_2_1a) qed lemma perp_col2_bis: assumes "A B Perp C D" and "Col C D P" and "Col C D Q" and "P \ Q" shows "A B Perp P Q" using Perp_cases assms(1) assms(2) assms(3) assms(4) perp_col0 by blast lemma perp_in_perp_bis_R1: assumes "X \ A" and "X PerpAt A B C D" shows "X B Perp C D \ A X Perp C D" by (metis assms(2) l8_14_2_1a perp_col perp_in_col) lemma perp_in_perp_bis: assumes "X PerpAt A B C D" shows "X B Perp C D \ A X Perp C D" by (metis assms l8_14_2_1a perp_in_perp_bis_R1) lemma col_per_perp: assumes "A \ B" and "B \ C" and (* "D \ B" and *) "D \ C" and "Col B C D" and "Per A B C" shows "C D Perp A B" by (metis Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) col_trivial_2 per_perp perp_col2_bis) lemma per_cong_mid_R1: assumes "B = H" and (* "B \ C" and *) "Bet A B C" and "Cong A H C H" and "Per H B C" shows "B Midpoint A C" using assms(1) assms(2) assms(3) midpoint_def not_cong_1243 by blast lemma per_cong_mid_R2: assumes (*"B \ H" and *) "B \ C" and "Bet A B C" and "Cong A H C H" and "Per H B C" shows "B Midpoint A C" proof - have P1: "Per C B H" using Per_cases assms(4) by blast have P2: "Per H B A" using assms(1) assms(2) assms(4) bet_col col_permutation_1 per_col by blast then have P3: "Per A B H" using Per_cases by blast obtain C' where P4: "B Midpoint C C' \ Cong H C H C'" using Per_def assms(4) by blast obtain H' where P5: "B Midpoint H H' \ Cong C H C H'" using P1 Per_def by blast obtain A' where P6: "B Midpoint A A' \ Cong H A H A'" using P2 Per_def by blast obtain H'' where P7: "B Midpoint H H'' \ Cong A H A H'" using P3 P5 Tarski_neutral_dimensionless.Per_def Tarski_neutral_dimensionless_axioms symmetric_point_uniqueness by fastforce then have P8: "H' = H''" using P5 symmetric_point_uniqueness by blast have "H B H' A IFSC H B H' C" proof - have Q1: "Bet H B H'" by (simp add: P7 P8 midpoint_bet) have Q2: "Cong H H' H H'" by (simp add: cong_reflexivity) have Q3: "Cong B H' B H'" by (simp add: cong_reflexivity) have Q4: "Cong H A H C" using assms(3) not_cong_2143 by blast have "Cong H' A H' C" using P5 P7 assms(3) cong_commutativity cong_inner_transitivity by blast thus ?thesis by (simp add: IFSC_def Q1 Q2 Q3 Q4) qed thus ?thesis using assms(1) assms(2) bet_col bet_neq23__neq l4_2 l7_20_bis by auto qed lemma per_cong_mid: assumes "B \ C" and "Bet A B C" and "Cong A H C H" and "Per H B C" shows "B Midpoint A C" using assms(1) assms(2) assms(3) assms(4) per_cong_mid_R1 per_cong_mid_R2 by blast lemma per_double_cong: assumes "Per A B C" and "B Midpoint C C'" shows "Cong A C A C'" using Mid_cases Per_def assms(1) assms(2) l7_9_bis by blast lemma cong_perp_or_mid_R1: assumes "Col A B X" and "A \ B" and "M Midpoint A B" and "Cong A X B X" shows "X = M \ \ Col A B X \ M PerpAt X M A B" using assms(1) assms(2) assms(3) assms(4) col_permutation_5 cong_commutativity l7_17_bis l7_2 l7_20 by blast lemma cong_perp_or_mid_R2: assumes "\ Col A B X" and "A \ B" and "M Midpoint A B" and "Cong A X B X" shows "X = M \ \ Col A B X \ M PerpAt X M A B" proof - have P1: "Col M A B" by (simp add: assms(3) midpoint_col) have "Per X M A" using Per_def assms(3) assms(4) cong_commutativity by blast thus ?thesis by (metis P1 assms(1) assms(2) assms(3) midpoint_distinct_1 not_col_permutation_4 per_perp_in perp_in_col_perp_in perp_in_right_comm) qed lemma cong_perp_or_mid: assumes "A \ B" and "M Midpoint A B" and "Cong A X B X" shows "X = M \ \ Col A B X \ M PerpAt X M A B" using assms(1) assms(2) assms(3) cong_perp_or_mid_R1 cong_perp_or_mid_R2 by blast lemma col_per2_cases: assumes "B \ C" and "B' \ C" and "C \ D" and "Col B C D" and "Per A B C" and "Per A B' C" shows "B = B' \ \ Col B' C D" by (meson Tarski_neutral_dimensionless.l8_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l6_16_1 per_col) lemma l8_16_1: assumes "Col A B X" and "Col A B U" and "A B Perp C X" shows "\ Col A B C \ Per C X U" by (metis assms(1) assms(2) assms(3) l8_5 perp_col0 perp_left_comm perp_not_col2 perp_per_2) lemma l8_16_2: assumes "Col A B X" and "Col A B U" and "U \ X" and "\ Col A B C" and "Per C X U" shows "A B Perp C X" proof - obtain X where "X PerpAt A B C X" by (metis (no_types) NCol_perm assms(1) assms(2) assms(3) assms(4) assms(5) l8_13_2 l8_2 not_col_distincts) thus ?thesis by (smt Perp_perm assms(1) assms(2) assms(3) assms(4) assms(5) col3 col_per_perp not_col_distincts per_col per_perp) qed lemma l8_18_uniqueness: assumes (*"\ Col A B C" and *) "Col A B X" and "A B Perp C X" and "Col A B Y" and "A B Perp C Y" shows "X = Y" using assms(1) assms(2) assms(3) assms(4) l8_16_1 l8_7 by blast lemma midpoint_distinct: assumes "\ Col A B C" and "Col A B X" and "X Midpoint C C'" shows "C \ C'" using assms(1) assms(2) assms(3) l7_3 by auto lemma l8_20_1_R1: assumes "A = B" shows "Per B A P" by (simp add: assms l8_2 l8_5) lemma l8_20_1_R2: assumes "A \ B" and "Per A B C" and "P Midpoint C' D" and "A Midpoint C' C" and "B Midpoint D C" shows "Per B A P" proof - obtain B' where P1: "A Midpoint B B'" using symmetric_point_construction by blast obtain D' where P2: "A Midpoint D D'" using symmetric_point_construction by blast obtain P' where P3: "A Midpoint P P'" using symmetric_point_construction by blast have P4: "Per B' B C" by (metis P1 Tarski_neutral_dimensionless.Per_cases Tarski_neutral_dimensionless.per_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) midpoint_col not_col_permutation_4) have P5: "Per B B' C'" proof - have "Per B' B C" by (simp add: P4) have "B' B C Cong3 B B' C'" by (meson Cong3_def P1 assms(4) l7_13 l7_2) thus ?thesis using P4 l8_10 by blast qed have P6: "B' Midpoint D' C'" by (meson P1 P2 assms(4) assms(5) l7_15 l7_16 l7_2 midpoint_bet midpoint_cong midpoint_def) have P7: "P' Midpoint C D'" using P2 P3 assms(3) assms(4) symmetry_preserves_midpoint by blast have P8: "A Midpoint P P'" by (simp add: P3) obtain D'' where P9: "B Midpoint C D'' \ Cong B' C B' D" using P4 assms(5) l7_2 per_double_cong by blast have P10: "D'' = D" using P9 assms(5) l7_9_bis by blast obtain D'' where P11: "B' Midpoint C' D'' \ Cong B C' B D''" using P5 Per_def by blast have P12: "D' = D''" by (meson P11 P6 Tarski_neutral_dimensionless.l7_9_bis Tarski_neutral_dimensionless_axioms) have P13: "P Midpoint C' D" using assms(3) by blast have P14: "Cong C D C' D'" using P2 assms(4) l7_13 l7_2 by blast have P15: "Cong C' D C D'" using P2 assms(4) cong_4321 l7_13 by blast have P16: "Cong P D P' D'" using P2 P8 cong_symmetry l7_13 by blast have P17: "Cong P D P' C" using P16 P7 cong_3421 cong_transitivity midpoint_cong by blast have P18: "C' P D B IFSC D' P' C B" by (metis Bet_cases IFSC_def P10 P11 P12 P13 P15 P17 P7 P9 cong_commutativity cong_right_commutativity l7_13 l7_3_2 midpoint_bet) then have "Cong B P B P'" using Tarski_neutral_dimensionless.l4_2 Tarski_neutral_dimensionless_axioms not_cong_2143 by fastforce thus ?thesis using P8 Per_def by blast qed lemma l8_20_1: assumes "Per A B C" and "P Midpoint C' D" and "A Midpoint C' C" and "B Midpoint D C" shows "Per B A P" using assms(1) assms(2) assms(3) assms(4) l8_20_1_R1 l8_20_1_R2 by fastforce lemma l8_20_2: assumes "P Midpoint C' D" and "A Midpoint C' C" and "B Midpoint D C" and "B \ C" shows "A \ P" using assms(1) assms(2) assms(3) assms(4) l7_3 symmetric_point_uniqueness by blast lemma perp_col1: assumes "C \ X" and "A B Perp C D" and "Col C D X" shows "A B Perp C X" using assms(1) assms(2) assms(3) col_trivial_3 perp_col2_bis by blast lemma l8_18_existence: assumes "\ Col A B C" shows "\ X. Col A B X \ A B Perp C X" proof - obtain Y where P1: "Bet B A Y \ Cong A Y A C" using segment_construction by blast then obtain P where P2: "P Midpoint C Y" using Mid_cases l7_25 by blast then have P3: "Per A P Y" using P1 Per_def l7_2 by blast obtain Z where P3: "Bet A Y Z \ Cong Y Z Y P" using segment_construction by blast obtain Q where P4: "Bet P Y Q \ Cong Y Q Y A" using segment_construction by blast obtain Q' where P5: "Bet Q Z Q' \ Cong Z Q' Q Z" using segment_construction by blast then have P6: "Z Midpoint Q Q'" using midpoint_def not_cong_3412 by blast obtain C' where P7: "Bet Q' Y C' \ Cong Y C' Y C" using segment_construction by blast obtain X where P8: "X Midpoint C C'" using Mid_cases P7 l7_25 by blast have P9: "A Y Z Q OFSC Q Y P A" by (simp add: OFSC_def P3 P4 between_symmetry cong_4321 cong_pseudo_reflexivity) have P10: "A \ Y" using P1 assms cong_reverse_identity not_col_distincts by blast then have P11: "Cong Z Q P A" using P9 five_segment_with_def by blast then have P12: "A P Y Cong3 Q Z Y" using Cong3_def P3 P4 not_cong_4321 by blast have P13: "Per Q Z Y" using Cong_perm P1 P12 P2 Per_def l8_10 l8_4 by blast then have P14: "Per Y Z Q" by (simp add: l8_2) have P15: "P \ Y" using NCol_cases P1 P2 assms bet_col l7_3_2 l7_9_bis by blast obtain Q'' where P16:"Z Midpoint Q Q'' \ Cong Y Q Y Q'" using P14 P6 per_double_cong by blast then have P17: "Q' = Q''" using P6 symmetric_point_uniqueness by blast have P18: "Bet Z Y X" proof - have "Bet Q Y C" using P15 P2 P4 between_symmetry midpoint_bet outer_transitivity_between2 by blast thus ?thesis using P16 P6 P7 P8 l7_22 not_cong_3412 by blast qed have P19: "Q \ Y" using P10 P4 cong_reverse_identity by blast have P20: "Per Y X C" proof - have "Bet C P Y" by (simp add: P2 midpoint_bet) thus ?thesis using P7 P8 Per_def not_cong_3412 by blast qed have P21: "Col P Y Q" by (simp add: Col_def P4) have P22: "Col P Y C" using P2 midpoint_col not_col_permutation_5 by blast have P23: "Col P Q C" using P15 P21 P22 col_transitivity_1 by blast have P24: "Col Y Q C" using P15 P21 P22 col_transitivity_2 by auto have P25: "Col A Y B" by (simp add: Col_def P1) have P26: "Col A Y Z" using P3 bet_col by blast have P27: "Col A B Z" using P10 P25 P26 col_transitivity_1 by blast have P28: "Col Y B Z" using P10 P25 P26 col_transitivity_2 by blast have P29: "Col Q Y P" using P21 not_col_permutation_3 by blast have P30: "Q \ C" using P15 P2 P4 between_equality_2 between_symmetry midpoint_bet by blast have P31: "Col Y B Z" using P28 by auto have P32: "Col Y Q' C'" by (simp add: P7 bet_col col_permutation_4) have P33: "Q \ Q'" using P11 P15 P22 P25 P5 assms bet_neq12__neq col_transitivity_1 cong_reverse_identity by blast have P34: "C \ C'" by (smt P15 P18 P3 P31 P8 assms bet_col col3 col_permutation_2 col_permutation_3 cong_3421 cong_diff midpoint_distinct_3) have P35: "Q Y C Z OFSC Q' Y C' Z" by (meson OFSC_def P15 P16 P2 P4 P5 P7 between_symmetry cong_3421 cong_reflexivity midpoint_bet not_cong_3412 outer_transitivity_between2) then have P36: "Cong C Z C' Z" using P19 five_segment_with_def by blast have P37: "Col Z Y X" by (simp add: P18 bet_col) have P38: "Y \ Z" using P15 P3 cong_reverse_identity by blast then have P40: "X \ Y" by (metis (mono_tags, opaque_lifting) Col_perm Cong_perm P14 P24 P25 P27 P36 P8 Per_def assms colx per_not_colp) have "Col A B X" using Col_perm P26 P31 P37 P38 col3 by blast thus ?thesis by (metis P18 P20 P27 P37 P40 Tarski_neutral_dimensionless.per_col Tarski_neutral_dimensionless_axioms assms between_equality col_permutation_3 l5_2 l8_16_2 l8_2) qed lemma l8_21_aux: assumes "\ Col A B C" shows "\ P. \ T. (A B Perp P A \ Col A B T \ Bet C T P)" proof - obtain X where P1: "Col A B X \ A B Perp C X" using assms l8_18_existence by blast have P2: "X PerpAt A B C X" by (simp add: P1 l8_15_1) have P3: "Per A X C" by (meson P1 Per_perm Tarski_neutral_dimensionless.l8_16_1 Tarski_neutral_dimensionless_axioms col_trivial_3) obtain C' where P4: "X Midpoint C C' \ Cong A C A C'" using P3 Per_def by blast obtain C'' where P5: "A Midpoint C C''" using symmetric_point_construction by blast obtain P where P6: "P Midpoint C' C''" by (metis Cong_perm P4 P5 Tarski_neutral_dimensionless.Midpoint_def Tarski_neutral_dimensionless_axioms cong_inner_transitivity l7_25) have P7: "Per X A P" by (smt P3 P4 P5 P6 l7_2 l8_20_1_R2 l8_4 midpoint_distinct_3 symmetric_point_uniqueness) have P8: "X \ C" using P1 assms by auto have P9: "A \ P" using P4 P5 P6 P8 l7_9 midpoint_distinct_2 by blast obtain T where P10: "Bet P T C \ Bet A T X" by (meson Mid_perm Midpoint_def P4 P5 P6 l3_17) have "A B Perp P A \ Col A B T \ Bet C T P" proof cases assume "A = X" thus ?thesis by (metis Bet_perm Col_def P1 P10 P9 between_identity col_trivial_3 perp_col2_bis) next assume "A \ X" thus ?thesis by (metis Bet_perm Col_def P1 P10 P7 P9 Perp_perm col_transitivity_2 col_trivial_1 l8_3 per_perp perp_not_col2) qed thus ?thesis by blast qed lemma l8_21: assumes "A \ B" shows "\ P T. A B Perp P A \ Col A B T \ Bet C T P" by (meson assms between_trivial2 l8_21_aux not_col_exists) lemma per_cong: assumes "A \ B" and "A \ P" and "Per B A P" and "Per A B R" and "Cong A P B R" and "Col A B X" and "Bet P X R" shows "Cong A R P B" proof - have P1: "Per P A B" using Per_cases assms(3) by blast obtain Q where P2: "R Midpoint B Q" using symmetric_point_construction by auto have P3: "B \ R" using assms(2) assms(5) cong_identity by blast have P4: "Per A B Q" by (metis P2 P3 assms(1) assms(4) bet_neq23__neq col_permutation_4 midpoint_bet midpoint_col per_perp_in perp_in_col_perp_in perp_in_per) have P5: "Per P A X" using P1 assms(1) assms(6) per_col by blast have P6: "B \ Q" using P2 P3 l7_3 by blast have P7: "Per R B X" by (metis assms(1) assms(4) assms(6) l8_2 not_col_permutation_4 per_col) have P8: "X \ A" using P3 assms(1) assms(2) assms(3) assms(4) assms(7) bet_col per_not_colp by blast obtain P' where P9: "A Midpoint P P'" using Per_def assms(3) by blast obtain R' where P10: "Bet P' X R' \ Cong X R' X R" using segment_construction by blast obtain M where P11: "M Midpoint R R'" by (meson P10 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms l7_25) have P12: "Per X M R" using P10 P11 Per_def cong_symmetry by blast have P13: "Cong X P X P'" using P9 assms(1) assms(3) assms(6) cong_left_commutativity l4_17 midpoint_cong per_double_cong by blast have P14: "X \ P'" using P13 P8 P9 cong_identity l7_3 by blast have P15: "P \ P'" using P9 assms(2) midpoint_distinct_2 by blast have P16: "\ Col X P P'" using P13 P15 P8 P9 l7_17 l7_20 not_col_permutation_4 by blast have P17: "Bet A X M" using P10 P11 P13 P9 assms(7) cong_symmetry l7_22 by blast have P18: "X \ R" using P3 P7 per_distinct_1 by blast have P19: "X \ R'" using P10 P18 cong_diff_3 by blast have P20: "X \ M" by (metis Col_def P10 P11 P16 P18 P19 assms(7) col_transitivity_1 midpoint_col) have P21: "M = B" by (smt Col_def P12 P17 P20 P8 Per_perm assms(1) assms(4) assms(6) col_transitivity_2 l8_3 l8_7) have "P X R P' OFSC P' X R' P" by (simp add: OFSC_def P10 P13 assms(7) cong_commutativity cong_pseudo_reflexivity cong_symmetry) then have "Cong R P' R' P" using P13 P14 cong_diff_3 five_segment_with_def by blast then have "P' A P R IFSC R' B R P" by (metis Bet_perm Cong_perm Midpoint_def P11 P21 P9 Tarski_neutral_dimensionless.IFSC_def Tarski_neutral_dimensionless_axioms assms(5) cong_mid2__cong cong_pseudo_reflexivity) thus ?thesis using l4_2 not_cong_1243 by blast qed lemma perp_cong: assumes "A \ B" and "A \ P" and "A B Perp P A" and "A B Perp R B" and "Cong A P B R" and "Col A B X" and "Bet P X R" shows "Cong A R P B" using Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) per_cong perp_per_1 by blast lemma perp_exists: assumes "A \ B" shows "\ X. PO X Perp A B" proof cases assume "Col A B PO" then obtain C where P1: "A \ C \ B \ C \ PO \ C \ Col A B C" using diff_col_ex3 by blast then obtain P T where P2: "PO C Perp P PO \ Col PO C T \ Bet PO T P" using l8_21 by blast then have "PO P Perp A B" by (metis P1 Perp_perm \Col A B PO\ assms col3 col_trivial_2 col_trivial_3 perp_col2) thus ?thesis by blast next assume "\ Col A B PO" thus ?thesis using l8_18_existence using assms col_trivial_2 col_trivial_3 l8_18_existence perp_col0 by blast qed lemma perp_vector: assumes "A \ B" shows "\ X Y. A B Perp X Y" using assms l8_21 by blast lemma midpoint_existence_aux: assumes "A \ B" and "A B Perp Q B" and "A B Perp P A" and "Col A B T" and "Bet Q T P" and "A P Le B Q" shows "\ X. X Midpoint A B" proof - obtain R where P1: "Bet B R Q \ Cong A P B R" using Le_def assms(6) by blast obtain X where P2: "Bet T X B \ Bet R X P" using P1 assms(5) between_symmetry inner_pasch by blast have P3: "Col A B X" by (metis Col_def Out_cases P2 assms(4) between_equality l6_16_1 not_out_bet out_diff1) have P4: "B \ R" using P1 assms(3) cong_identity perp_not_eq_2 by blast have P5: "\ Col A B Q" using assms(2) col_trivial_2 l8_16_1 by blast have P6: "\ Col A B R" using Col_def P1 P4 P5 l6_16_1 by blast have P7: "P \ R" using P2 P3 P6 between_identity by blast have "\ X. X Midpoint A B" proof cases assume "A = P" thus ?thesis using assms(3) col_trivial_3 perp_not_col2 by blast next assume Q1: "\ A = P" have Q2: "A B Perp R B" by (metis P1 P4 Perp_perm Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) l5_1 perp_col1) then have Q3: "Cong A R P B" using P1 P2 P3 Q1 assms(1) assms(3) between_symmetry perp_cong by blast then have "X Midpoint A B \ X Midpoint P R" by (smt P1 P2 P3 P6 P7 bet_col cong_left_commutativity cong_symmetry l7_2 l7_21 not_col_permutation_1) thus ?thesis by blast qed thus ?thesis by blast qed lemma midpoint_existence: "\ X. X Midpoint A B" proof cases assume "A = B" thus ?thesis using l7_3_2 by blast next assume P1: "\ A = B" obtain Q where P2: "A B Perp B Q" by (metis P1 l8_21 perp_comm) obtain P T where P3: "A B Perp P A \ Col A B T \ Bet Q T P" using P2 l8_21_aux not_col_distincts perp_not_col2 by blast have P4: "A P Le B Q \ B Q Le A P" by (simp add: local.le_cases) have P5: "A P Le B Q \ (\ X. X Midpoint A B)" by (meson P1 P2 P3 Tarski_neutral_dimensionless.Perp_cases Tarski_neutral_dimensionless.midpoint_existence_aux Tarski_neutral_dimensionless_axioms) have P6: "B Q Le A P \ (\ X. X Midpoint A B)" proof - { assume H1: "B Q Le A P" have Q6: "B \ A" using P1 by auto have Q2: "B A Perp P A" by (simp add: P3 perp_left_comm) have Q3: "B A Perp Q B" using P2 Perp_perm by blast have Q4: "Col B A T" using Col_perm P3 by blast have Q5: "Bet P T Q" using Bet_perm P3 by blast obtain X where "X Midpoint B A" using H1 Q2 Q3 Q4 Q5 Q6 midpoint_existence_aux by blast then have "\ X. X Midpoint A B" using l7_2 by blast } thus ?thesis by simp qed thus ?thesis using P4 P5 by blast qed lemma perp_in_id: assumes "X PerpAt A B C A" shows "X = A" by (meson Col_cases assms col_trivial_3 l8_14_2_1b) lemma l8_22: assumes "A \ B" and "A \ P" and "Per B A P" and "Per A B R" and "Cong A P B R" and "Col A B X" and "Bet P X R" and "Cong A R P B" shows "X Midpoint A B \ X Midpoint P R" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) bet_col cong_commutativity cong_diff cong_right_commutativity l7_21 not_col_permutation_5 per_not_colp) lemma l8_22_bis: assumes "A \ B" and "A \ P" and "A B Perp P A" and "A B Perp R B" and "Cong A P B R" and "Col A B X" and "Bet P X R" shows "Cong A R P B \ X Midpoint A B \ X Midpoint P R" by (metis l8_22 Perp_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) perp_cong perp_per_2) lemma perp_in_perp: assumes "X PerpAt A B C D" shows "A B Perp C D" using assms l8_14_2_1a by auto lemma perp_proj: assumes "A B Perp C D" and "\ Col A C D" shows "\ X. Col A B X \ A X Perp C D" using assms(1) not_col_distincts by auto lemma l8_24 : assumes "P A Perp A B" and "Q B Perp A B" and "Col A B T" and "Bet P T Q" and "Bet B R Q" and "Cong A P B R" shows "\ X. X Midpoint A B \ X Midpoint P R" proof - obtain X where P1: "Bet T X B \ Bet R X P" using assms(4) assms(5) inner_pasch by blast have P2: "Col A B X" by (metis Out_cases P1 assms(3) bet_out_1 col_out2_col not_col_distincts out_trivial) have P3: "A \ B" using assms(1) col_trivial_2 l8_16_1 by blast have P4: "A \ P" using assms(1) col_trivial_1 l8_16_1 by blast have "\ X. X Midpoint A B \ X Midpoint P R" proof cases assume "Col A B P" thus ?thesis using Perp_perm assms(1) perp_not_col by blast next assume Q1: "\ Col A B P" have Q2: "B \ R" using P4 assms(6) cong_diff by blast have Q3: "Q \ B" using Q2 assms(5) between_identity by blast have Q4: "\ Col A B Q" by (metis assms(2) col_permutation_3 l8_14_1 perp_col1 perp_not_col) have Q5: "\ Col A B R" by (meson Q2 Q4 assms(5) bet_col col_transitivity_1 not_col_permutation_2) have Q6: "P \ R" using P1 P2 Q5 between_identity by blast have "\ X. X Midpoint A B \ X Midpoint P R" proof cases assume "A = P" thus ?thesis using P4 by blast next assume R0: "\ A = P" have R1: "A B Perp R B" by (metis Perp_cases Q2 Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) assms(5) bet_col col_transitivity_1 perp_col1) have R2: "Cong A R P B" using P1 P2 P3 Perp_perm R0 R1 assms(1) assms(6) between_symmetry perp_cong by blast have R3: "\ Col A P B" using Col_perm Q1 by blast have R4: "P \ R" by (simp add: Q6) have R5: "Cong A P B R" by (simp add: assms(6)) have R6: "Cong P B R A" using R2 not_cong_4312 by blast have R7: "Col A X B" using Col_perm P2 by blast have R8: "Col P X R" by (simp add: P1 bet_col between_symmetry) thus ?thesis using l7_21 using R3 R4 R5 R6 R7 by blast qed thus ?thesis by simp qed thus ?thesis by simp qed lemma col_per2__per: assumes "A \ B" and "Col A B C" and "Per A X P" and "Per B X P" shows "Per C X P" by (meson Per_def assms(1) assms(2) assms(3) assms(4) l4_17 per_double_cong) lemma perp_in_per_1: assumes "X PerpAt A B C D" shows "Per A X C" using PerpAt_def assms col_trivial_1 by auto lemma perp_in_per_2: assumes "X PerpAt A B C D" shows "Per A X D" using assms perp_in_per_1 perp_in_right_comm by blast lemma perp_in_per_3: assumes "X PerpAt A B C D" shows "Per B X C" using assms perp_in_comm perp_in_per_2 by blast lemma perp_in_per_4: assumes "X PerpAt A B C D" shows "Per B X D" using assms perp_in_per_3 perp_in_right_comm by blast subsection "Planes" subsubsection "Coplanar" lemma coplanar_perm_1: assumes "Coplanar A B C D" shows "Coplanar A B D C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_2: assumes "Coplanar A B C D" shows "Coplanar A C B D" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_3: assumes "Coplanar A B C D" shows "Coplanar A C D B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_4: assumes "Coplanar A B C D" shows "Coplanar A D B C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_5: assumes "Coplanar A B C D" shows "Coplanar A D C B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_6: assumes "Coplanar A B C D" shows "Coplanar B A C D" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_7: assumes "Coplanar A B C D" shows "Coplanar B A D C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_8: assumes "Coplanar A B C D" shows "Coplanar B C A D" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_9: assumes "Coplanar A B C D" shows "Coplanar B C D A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_10: assumes "Coplanar A B C D" shows "Coplanar B D A C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_11: assumes "Coplanar A B C D" shows "Coplanar B D C A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_12: assumes "Coplanar A B C D" shows "Coplanar C A B D" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_13: assumes "Coplanar A B C D" shows "Coplanar C A D B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_14: assumes "Coplanar A B C D" shows "Coplanar C B A D" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_15: assumes "Coplanar A B C D" shows "Coplanar C B D A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_16: assumes "Coplanar A B C D" shows "Coplanar C D A B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_17: assumes "Coplanar A B C D" shows "Coplanar C D B A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_18: assumes "Coplanar A B C D" shows "Coplanar D A B C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_19: assumes "Coplanar A B C D" shows "Coplanar D A C B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_20: assumes "Coplanar A B C D" shows "Coplanar D B A C" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_21: assumes "Coplanar A B C D" shows "Coplanar D B C A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_22: assumes "Coplanar A B C D" shows "Coplanar D C A B" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma coplanar_perm_23: assumes "Coplanar A B C D" shows "Coplanar D C B A" proof - obtain X where P1: "(Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using Coplanar_def assms by blast then show ?thesis using Coplanar_def col_permutation_4 by blast qed lemma ncoplanar_perm_1: assumes "\ Coplanar A B C D" shows "\ Coplanar A B D C" using assms coplanar_perm_1 by blast lemma ncoplanar_perm_2: assumes "\ Coplanar A B C D" shows "\ Coplanar A C B D" using assms coplanar_perm_2 by blast lemma ncoplanar_perm_3: assumes "\ Coplanar A B C D" shows "\ Coplanar A C D B" using assms coplanar_perm_4 by blast lemma ncoplanar_perm_4: assumes "\ Coplanar A B C D" shows "\ Coplanar A D B C" using assms coplanar_perm_3 by blast lemma ncoplanar_perm_5: assumes "\ Coplanar A B C D" shows "\ Coplanar A D C B" using assms coplanar_perm_5 by blast lemma ncoplanar_perm_6: assumes "\ Coplanar A B C D" shows "\ Coplanar B A C D" using assms coplanar_perm_6 by blast lemma ncoplanar_perm_7: assumes "\ Coplanar A B C D" shows "\ Coplanar B A D C" using assms coplanar_perm_7 by blast lemma ncoplanar_perm_8: assumes "\ Coplanar A B C D" shows "\ Coplanar B C A D" using assms coplanar_perm_12 by blast lemma ncoplanar_perm_9: assumes "\ Coplanar A B C D" shows "\ Coplanar B C D A" using assms coplanar_perm_18 by blast lemma ncoplanar_perm_10: assumes "\ Coplanar A B C D" shows "\ Coplanar B D A C" using assms coplanar_perm_13 by blast lemma ncoplanar_perm_11: assumes "\ Coplanar A B C D" shows "\ Coplanar B D C A" using assms coplanar_perm_19 by blast lemma ncoplanar_perm_12: assumes "\ Coplanar A B C D" shows "\ Coplanar C A B D" using assms coplanar_perm_8 by blast lemma ncoplanar_perm_13: assumes "\ Coplanar A B C D" shows "\ Coplanar C A D B" using assms coplanar_perm_10 by blast lemma ncoplanar_perm_14: assumes "\ Coplanar A B C D" shows "\ Coplanar C B A D" using assms coplanar_perm_14 by blast lemma ncoplanar_perm_15: assumes "\ Coplanar A B C D" shows "\ Coplanar C B D A" using assms coplanar_perm_20 by blast lemma ncoplanar_perm_16: assumes "\ Coplanar A B C D" shows "\ Coplanar C D A B" using assms coplanar_perm_16 by blast lemma ncoplanar_perm_17: assumes "\ Coplanar A B C D" shows "\ Coplanar C D B A" using assms coplanar_perm_22 by blast lemma ncoplanar_perm_18: assumes "\ Coplanar A B C D" shows "\ Coplanar D A B C" using assms coplanar_perm_9 by blast lemma ncoplanar_perm_19: assumes "\ Coplanar A B C D" shows "\ Coplanar D A C B" using assms coplanar_perm_11 by blast lemma ncoplanar_perm_20: assumes "\ Coplanar A B C D" shows "\ Coplanar D B A C" using assms coplanar_perm_15 by blast lemma ncoplanar_perm_21: assumes "\ Coplanar A B C D" shows "\ Coplanar D B C A" using assms coplanar_perm_21 by blast lemma ncoplanar_perm_22: assumes "\ Coplanar A B C D" shows "\ Coplanar D C A B" using assms coplanar_perm_17 by blast lemma ncoplanar_perm_23: assumes "\ Coplanar A B C D" shows "\ Coplanar D C B A" using assms coplanar_perm_23 by blast lemma coplanar_trivial: shows "Coplanar A A B C" using Coplanar_def NCol_cases col_trivial_1 by blast lemma col__coplanar: assumes "Col A B C" shows "Coplanar A B C D" using Coplanar_def assms not_col_distincts by blast lemma ncop__ncol: assumes "\ Coplanar A B C D" shows "\ Col A B C" using assms col__coplanar by blast lemma ncop__ncols: assumes "\ Coplanar A B C D" shows "\ Col A B C \ \ Col A B D \ \ Col A C D \ \ Col B C D" by (meson assms col__coplanar coplanar_perm_4 ncoplanar_perm_9) lemma bet__coplanar: assumes "Bet A B C" shows "Coplanar A B C D" using assms bet_col ncop__ncol by blast lemma out__coplanar: assumes "A Out B C" shows "Coplanar A B C D" using assms col__coplanar out_col by blast lemma midpoint__coplanar: assumes "A Midpoint B C" shows "Coplanar A B C D" using assms midpoint_col ncop__ncol by blast lemma perp__coplanar: assumes "A B Perp C D" shows "Coplanar A B C D" proof - obtain P where "P PerpAt A B C D" using Perp_def assms by blast then show ?thesis using Coplanar_def perp_in_col by blast qed lemma ts__coplanar: assumes "A B TS C D" shows "Coplanar A B C D" by (metis (full_types) Coplanar_def TS_def assms bet_col col_permutation_2 col_permutation_3) lemma reflectl__coplanar: assumes "A B ReflectL C D" shows "Coplanar A B C D" by (metis (no_types) ReflectL_def Tarski_neutral_dimensionless.perp__coplanar Tarski_neutral_dimensionless_axioms assms col__coplanar col_trivial_1 ncoplanar_perm_17) lemma reflect__coplanar: assumes "A B Reflect C D" shows "Coplanar A B C D" by (metis (no_types) Reflect_def Tarski_neutral_dimensionless.reflectl__coplanar Tarski_neutral_dimensionless_axioms assms col_trivial_2 ncop__ncols) lemma inangle__coplanar: assumes "A InAngle B C D" shows "Coplanar A B C D" proof - obtain X where P1: "Bet B X D \ (X = C \ C Out X A)" using InAngle_def assms by auto then show ?thesis by (meson Col_cases Coplanar_def bet_col ncop__ncols out_col) qed lemma pars__coplanar: assumes "A B ParStrict C D" shows "Coplanar A B C D" using ParStrict_def assms by auto lemma par__coplanar: assumes "A B Par C D" shows "Coplanar A B C D" using Par_def assms ncop__ncols pars__coplanar by blast lemma plg__coplanar: assumes "Plg A B C D" shows "Coplanar A B C D" proof - obtain M where "Bet A M C \ Bet B M D" by (meson Plg_def assms midpoint_bet) then show ?thesis by (metis InAngle_def bet_out_1 inangle__coplanar ncop__ncols not_col_distincts) qed lemma plgs__coplanar: assumes "ParallelogramStrict A B C D" shows "Coplanar A B C D" using ParallelogramStrict_def assms par__coplanar by blast lemma plgf__coplanar: assumes "ParallelogramFlat A B C D" shows "Coplanar A B C D" using ParallelogramFlat_def assms col__coplanar by auto lemma parallelogram__coplanar: assumes "Parallelogram A B C D" shows "Coplanar A B C D" using Parallelogram_def assms plgf__coplanar plgs__coplanar by auto lemma rhombus__coplanar: assumes "Rhombus A B C D" shows "Coplanar A B C D" using Rhombus_def assms plg__coplanar by blast lemma rectangle__coplanar: assumes "Rectangle A B C D" shows "Coplanar A B C D" using Rectangle_def assms plg__coplanar by blast lemma square__coplanar: assumes "Square A B C D" shows "Coplanar A B C D" using Square_def assms rectangle__coplanar by blast lemma lambert__coplanar: assumes "Lambert A B C D" shows "Coplanar A B C D" using Lambert_def assms by presburger subsubsection "Planes" lemma ts_distincts: assumes "A B TS P Q" shows "A \ B \ A \ P \ A \ Q \ B \ P \ B \ Q \ P \ Q" using TS_def assms bet_neq12__neq not_col_distincts by blast lemma l9_2: assumes "A B TS P Q" shows "A B TS Q P" using TS_def assms between_symmetry by blast lemma invert_two_sides: assumes "A B TS P Q" shows "B A TS P Q" using TS_def assms not_col_permutation_5 by blast lemma l9_3: assumes "P Q TS A C" and "Col M P Q" and "M Midpoint A C" and "Col R P Q" and "R Out A B" shows "P Q TS B C" proof - have P1: "\ Col A P Q" using TS_def assms(1) by blast have P2: "P \ Q" using P1 not_col_distincts by auto obtain T where P3: "Col T P Q \ Bet A T C" using assms(2) assms(3) midpoint_bet by blast have P4: "A \ C" using assms(1) ts_distincts by blast have P5: "T = M" by (smt P1 P3 Tarski_neutral_dimensionless.bet_col1 Tarski_neutral_dimensionless_axioms assms(2) assms(3) col_permutation_2 l6_21 midpoint_bet) have "P Q TS B C" proof cases assume "C = M" then show ?thesis using P4 assms(3) midpoint_distinct_1 by blast next assume P6: "\ C = M" have P7: "\ Col B P Q" by (metis P1 assms(4) assms(5) col_permutation_1 colx l6_3_1 out_col) have P97: "Bet R A B \ Bet R B A" using Out_def assms(5) by auto { assume Q1: "Bet R A B" obtain B' where Q2: "M Midpoint B B'" using symmetric_point_construction by blast obtain R' where Q3: "M Midpoint R R'" using symmetric_point_construction by blast have Q4: "Bet B' C R'" using Q1 Q2 Q3 assms(3) between_symmetry l7_15 by blast obtain X where Q5: "Bet M X R' \ Bet C X B" using Bet_perm Midpoint_def Q2 Q4 between_trivial2 l3_17 by blast have Q6: "Col X P Q" proof - have R1: "Col P M R" using P2 assms(2) assms(4) col_permutation_4 l6_16_1 by blast have R2: "Col Q M R" by (metis R1 assms(2) assms(4) col_permutation_5 l6_16_1 not_col_permutation_3) { assume "M = X" then have "Col X P Q" using assms(2) by blast } then have R3: "M = X \ Col X P Q" by simp { assume "M \ X" then have S1: "M \ R'" using Q5 bet_neq12__neq by blast have "M \ R" using Q3 S1 midpoint_distinct_1 by blast then have "Col X P Q" by (smt Col_perm Q3 Q5 R1 R2 S1 bet_out col_transitivity_2 midpoint_col out_col) } then have "M \ X \ Col X P Q" by simp then show ?thesis using R3 by blast qed have "Bet B X C" using Q5 between_symmetry by blast then have "P Q TS B C" using Q6 using P7 TS_def assms(1) by blast } then have P98: "Bet R A B \ P Q TS B C" by simp { assume S2: "Bet R B A" have S3: "Bet C M A" using Bet_perm P3 P5 by blast then obtain X where "Bet B X C \ Bet M X R" using S2 inner_pasch by blast then have "P Q TS B C" by (metis Col_def P7 TS_def assms(1) assms(2) assms(4) between_inner_transitivity between_trivial l6_16_1 not_col_permutation_5) } then have "Bet R B A \ P Q TS B C" by simp then show ?thesis using P97 P98 by blast qed then show ?thesis by blast qed lemma mid_preserves_col: assumes "Col A B C" and "M Midpoint A A'" and "M Midpoint B B'" and "M Midpoint C C'" shows "Col A' B' C'" using Col_def assms(1) assms(2) assms(3) assms(4) l7_15 by auto lemma per_mid_per: assumes (*"A \ B" and*) "Per X A B" and "M Midpoint A B" and "M Midpoint X Y" shows "Cong A X B Y \ Per Y B A" by (meson Cong3_def Mid_perm assms(1) assms(2) assms(3) l7_13 l8_10) lemma sym_preserve_diff: assumes "A \ B" and "M Midpoint A A'" and "M Midpoint B B'" shows "A'\ B'" using assms(1) assms(2) assms(3) l7_9 by blast lemma l9_4_1_aux_R1: assumes "R = S" and "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "M Midpoint R S" shows "\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C')" proof - have P1: "M = R" using assms(1) assms(8) l7_3 by blast have P2: "\ Col A P Q" using TS_def assms(3) by auto then have P3: "P \ Q" using not_col_distincts by blast obtain T where P4: "Col T P Q \ Bet A T C" using TS_def assms(3) by blast { assume "\ M = T" then have "M PerpAt M T A M" using perp_col2 by (metis P1 P4 assms(4) assms(5) not_col_permutation_3 perp_left_comm perp_perp_in) then have "M T Perp C M" using P1 P4 \M \ T\ assms(1) assms(4) assms(7) col_permutation_1 perp_col2 by blast then have "Per T M A" using \M PerpAt M T A M\ perp_in_per_3 by blast have "Per T M C" by (simp add: \M T Perp C M\ perp_per_1) have "M = T" proof - have "Per C M T" by (simp add: \Per T M C\ l8_2) then show ?thesis using l8_6 l8_2 using P4 \Per T M A\ by blast qed then have "False" using \M \ T\ by blast } then have Q0: "M = T" by blast have R1: "\ U C'. ((M Midpoint U C' \ M Out U A) \ M Out C C')" proof - { fix U C' assume Q1: "M Midpoint U C' \ M Out U A" have Q2: "C \ M" using P1 assms(1) assms(7) perp_not_eq_2 by blast have Q3: "C' \ M" using Q1 midpoint_not_midpoint out_diff1 by blast have Q4: "Bet U M C" using P4 Q0 Q1 bet_out__bet l6_6 by blast then have "M Out C C'" by (metis (full_types) Out_def Q1 Q2 Q3 l5_2 midpoint_bet) } then show ?thesis by blast qed have R2: "\ U C'. ((M Midpoint U C' \ M Out C C') \ M Out U A)" proof - { fix U C' assume Q1: "M Midpoint U C' \ M Out C C'" have Q2: "C \ M" using P1 assms(1) assms(7) perp_not_eq_2 by blast have Q3: "C' \ M" using Q1 l6_3_1 by blast have Q4: "Bet U M C" by (metis Out_def Q1 between_inner_transitivity midpoint_bet outer_transitivity_between) then have "M Out U A" by (metis P2 P4 Q0 Q1 Q2 Q3 l6_2 midpoint_distinct_1) } then show ?thesis by blast qed then show ?thesis using R1 P1 P2 assms by blast qed lemma l9_4_1_aux_R21: assumes "R \ S" and "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "M Midpoint R S" shows "\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C')" proof - obtain D where P1: "Bet R D A \ Cong S C R D" using Le_def assms(2) by blast have P2: "C \ S" using assms(7) perp_not_eq_2 by auto have P3: "R \ D" using P1 P2 cong_identity by blast have P4: "R S Perp A R" using assms(1) assms(4) assms(5) assms(6) not_col_permutation_2 perp_col2 by blast have "\ M. (M Midpoint S R \ M Midpoint C D)" proof - have Q1: "\ Col A P Q" using TS_def assms(3) by blast have Q2: "P \ Q" using Q1 not_col_distincts by blast obtain T where Q3: "Col T P Q \ Bet A T C" using TS_def assms(3) by blast have Q4: "C S Perp S R" by (metis NCol_perm assms(1) assms(4) assms(6) assms(7) perp_col0) have Q5: "A R Perp S R" using P4 Perp_perm by blast have Q6: "Col S R T" using Col_cases Q2 Q3 assms(4) assms(6) col3 by blast have Q7: "Bet C T A" using Bet_perm Q3 by blast have Q8: "Bet R D A" by (simp add: P1) have "Cong S C R D" by (simp add: P1) then show ?thesis using P1 Q4 Q5 Q6 Q7 l8_24 by blast qed then obtain M' where P5: "M' Midpoint S R \ M' Midpoint C D" by blast have P6: "M = M'" by (meson P5 assms(8) l7_17_bis) have L1: "\ U C'. (M Midpoint U C' \ R Out U A) \ S Out C C'" proof - { fix U C' assume R1: "M Midpoint U C' \ R Out U A" have R2: "C \ S" using P2 by auto have R3: "C' \ S" using P5 R1 P6 l7_9_bis out_diff1 by blast have R4: "Bet S C C' \ Bet S C' C" proof - have R5: "Bet R U A \ Bet R A U" using Out_def R1 by auto { assume "Bet R U A" then have "Bet R U D \ Bet R D U" using P1 l5_3 by blast then have "Bet S C C' \ Bet S C' C" using P5 P6 R1 l7_15 l7_2 by blast } then have R6: "Bet R U A \ Bet S C C' \ Bet S C' C" by simp have "Bet R A U \ Bet S C C' \ Bet S C' C" using P1 P5 P6 R1 between_exchange4 l7_15 l7_2 by blast then show ?thesis using R5 R6 by blast qed then have "S Out C C'" by (simp add: Out_def R2 R3) } then show ?thesis by simp qed have "\ U C'. (M Midpoint U C' \ S Out C C') \ R Out U A" proof - { fix U C' assume Q1: "M Midpoint U C' \ S Out C C'" then have Q2: "U \ R" using P5 P6 l7_9_bis out_diff2 by blast have Q3: "A \ R" using assms(5) perp_not_eq_2 by auto have Q4: "Bet S C C' \ Bet S C' C" using Out_def Q1 by auto { assume V0: "Bet S C C'" have V1: "R \ D" by (simp add: P3) then have V2: "Bet R D U" proof - have W1: "M Midpoint S R" using P5 P6 by blast have W2: "M Midpoint C D" by (simp add: P5 P6) have "M Midpoint C' U" by (simp add: Q1 l7_2) then show ?thesis using V0 P5 P6 l7_15 by blast qed have "Bet R D A" using P1 by auto then have "Bet R U A \ Bet R A U" using V1 V2 l5_1 by blast } then have Q5: "Bet S C C' \ Bet R U A \ Bet R A U" by simp { assume R1: "Bet S C' C" have "Bet R U A" using P1 P5 P6 Q1 R1 between_exchange4 l7_15 l7_2 by blast } then have "Bet S C' C \ Bet R U A \ Bet R A U" by simp then have "Bet R U A \ Bet R A U" using Q4 Q5 by blast then have "R Out U A" by (simp add: Out_def Q2 Q3) } then show ?thesis by simp qed then show ?thesis using L1 by blast qed lemma l9_4_1_aux: assumes "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "M Midpoint R S" shows "\ U C'. (M Midpoint U C' \ (R Out U A \ S Out C C'))" using l9_4_1_aux_R1 l9_4_1_aux_R21 assms by smt lemma per_col_eq: assumes "Per A B C" and "Col A B C" and "B \ C" shows "A = B" using assms(1) assms(2) assms(3) l8_9 by blast lemma l9_4_1: assumes "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "M Midpoint R S" shows "\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C')" proof - have P1: "S C Le R A \ R A Le S C" using local.le_cases by blast { assume Q1: "S C Le R A" { fix U C' assume "M Midpoint U C'" then have "(R Out U A \ S Out C C')" using Q1 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l9_4_1_aux by blast } then have "\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C')" by simp } then have P2: "S C Le R A \ (\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C'))" by simp { assume Q2: " R A Le S C" { fix U C' assume "M Midpoint U C'" then have "(R Out A U \ S Out C' C)" using Q2 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_2 l9_2 l9_4_1_aux by blast then have "(R Out U A \ S Out C C')" using l6_6 by blast } then have "\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C')" by simp } then have P3: "R A Le S C \ (\ U C'. M Midpoint U C' \ (R Out U A \ S Out C C'))" by simp then show ?thesis using P1 P2 by blast qed lemma mid_two_sides: assumes "M Midpoint A B" and "\ Col A B X" and "M Midpoint X Y" shows "A B TS X Y" proof - have f1: "\ Col Y A B" by (meson Mid_cases Tarski_neutral_dimensionless.mid_preserves_col Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) col_permutation_3) have "Bet X M Y" using assms(3) midpoint_bet by blast then show ?thesis using f1 by (metis (no_types) TS_def assms(1) assms(2) col_permutation_1 midpoint_col) qed lemma col_preserves_two_sides: assumes "C \ D" and "Col A B C" and "Col A B D" and "A B TS X Y" shows "C D TS X Y" proof - have P1: "\ Col X A B" using TS_def assms(4) by blast then have P2: "A \ B" using not_col_distincts by blast have P3: "\ Col X C D" by (metis Col_cases P1 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3)) have P4: "\ Col Y C D" by (metis Col_cases TS_def Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4)) then show ?thesis proof - obtain pp :: "'p \ 'p \ 'p \ 'p \ 'p" where "\x0 x1 x2 x3. (\v4. Col v4 x3 x2 \ Bet x1 v4 x0) = (Col (pp x0 x1 x2 x3) x3 x2 \ Bet x1 (pp x0 x1 x2 x3) x0)" by moura then have f1: "\ Col X A B \ \ Col Y A B \ Col (pp Y X B A) A B \ Bet X (pp Y X B A) Y" using TS_def assms(4) by presburger then have "Col (pp Y X B A) C D" by (meson P2 assms(2) assms(3) col3 not_col_permutation_3 not_col_permutation_4) then show ?thesis using f1 TS_def P3 P4 by blast qed qed lemma out_out_two_sides: assumes "A \ B" and "A B TS X Y" and "Col I A B" and "Col I X Y" and "I Out X U" and "I Out Y V" shows "A B TS U V" proof - have P0: "\ Col X A B" using TS_def assms(2) by blast then have P1: "\ Col V A B" by (smt assms(2) assms(3) assms(4) assms(6) col_out2_col col_transitivity_1 not_col_permutation_3 not_col_permutation_4 out_diff2 out_trivial ts_distincts) have P2: "\ Col U A B" by (metis P0 assms(3) assms(5) col_permutation_2 colx out_col out_distinct) obtain T where P3: "Col T A B \ Bet X T Y" using TS_def assms(2) by blast have "I = T" proof - have f1: "\p pa pb. \ Col p pa pb \ \ Col p pb pa \ \ Col pa p pb \ \ Col pa pb p \ \ Col pb p pa \ \ Col pb pa p \ Col p pa pb" using Col_cases by blast then have f2: "Col X Y I" using assms(4) by blast have f3: "Col B A I" using f1 assms(3) by blast have f4: "Col B A T" using f1 P3 by blast have f5: "\ Col X A B \ \ Col X B A \ \ Col A X B \ \ Col A B X \ \ Col B X A \ \ Col B A X" using f1 \\ Col X A B\ by blast have f6: "A \ B \ A \ X \ A \ Y \ B \ X \ B \ Y \ X \ Y" using assms(2) ts_distincts by presburger have "Col X Y T" using f1 by (meson P3 bet_col) then show ?thesis using f6 f5 f4 f3 f2 by (meson Tarski_neutral_dimensionless.l6_21 Tarski_neutral_dimensionless_axioms) qed then have "Bet U T V" using P3 assms(5) assms(6) bet_out_out_bet by blast then show ?thesis using P1 P2 P3 TS_def by blast qed lemma l9_4_2_aux_R1: assumes "R = S " and "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "R Out U A" and "S Out V C" shows "P Q TS U V" proof - have "\ Col A P Q" using TS_def assms(3) by auto then have P2: "P \ Q" using not_col_distincts by blast obtain T where P3: "Col T P Q \ Bet A T C" using TS_def assms(3) by blast have "R = T" using assms(1) assms(5) assms(6) assms(7) col_permutation_1 l8_16_1 l8_6 by (meson P3) then show ?thesis by (smt P2 P3 assms(1) assms(3) assms(8) assms(9) bet_col col_transitivity_2 l6_6 not_col_distincts out_out_two_sides) qed lemma l9_4_2_aux_R2: assumes "R \ S" and "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "R Out U A" and "S Out V C" shows "P Q TS U V" proof - have P1: "P \ Q" using assms(7) perp_distinct by auto have P2: "R S TS A C" using assms(1) assms(3) assms(4) assms(6) col_permutation_1 col_preserves_two_sides by blast have P3: "Col R S P" using P1 assms(4) assms(6) col2__eq not_col_permutation_1 by blast have P4: "Col R S Q" by (metis P3 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(4) assms(6) col_trivial_2) have P5: "R S Perp A R" using NCol_perm assms(1) assms(4) assms(5) assms(6) perp_col2 by blast have P6: "R S Perp C S" using assms(1) assms(4) assms(6) assms(7) col_permutation_1 perp_col2 by blast have P7: "\ Col A R S" using P2 TS_def by blast obtain T where P8: "Col T R S \ Bet A T C" using P2 TS_def by blast obtain C' where P9: "Bet R C' A \ Cong S C R C'" using Le_def assms(2) by blast have "\ X. X Midpoint S R \ X Midpoint C C'" proof - have Q1: "C S Perp S R" using P6 Perp_perm by blast have Q2: "A R Perp S R" using P5 Perp_perm by blast have Q3: "Col S R T" using Col_cases P8 by blast have Q4: "Bet C T A" using Bet_perm P8 by blast have Q5: "Bet R C' A" by (simp add: P9) have "Cong S C R C'" by (simp add: P9) then show ?thesis using Q1 Q2 Q3 Q4 Q5 l8_24 by blast qed then obtain M where P10: "M Midpoint S R \ M Midpoint C C'" by blast obtain U' where P11: "M Midpoint U U'" using symmetric_point_construction by blast have P12: "R \ U" using assms(8) out_diff1 by blast have P13: "R S TS U U'" by (smt P10 P11 P12 P7 assms(8) col_transitivity_2 invert_two_sides mid_two_sides not_col_permutation_3 not_col_permutation_4 out_col) have P14: "R S TS V U" proof - have Q1: "Col M R S" using P10 midpoint_col not_col_permutation_5 by blast have Q2: "M Midpoint U' U" by (meson P11 Tarski_neutral_dimensionless.Mid_cases Tarski_neutral_dimensionless_axioms) have "S Out U' V" by (meson P10 P11 P2 P5 P6 Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(8) assms(9) l6_6 l6_7 l9_4_1_aux_R21 not_col_distincts) then show ?thesis using P13 Q1 Q2 col_trivial_3 l9_2 l9_3 by blast qed then show ?thesis using P1 P3 P4 col_preserves_two_sides l9_2 by blast qed lemma l9_4_2_aux: assumes "S C Le R A" and "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "R Out U A" and "S Out V C" shows "P Q TS U V" using l9_4_2_aux_R1 l9_4_2_aux_R2 by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8)) lemma l9_4_2: assumes "P Q TS A C" and "Col R P Q" and "P Q Perp A R" and "Col S P Q" and "P Q Perp C S" and "R Out U A" and "S Out V C" shows "P Q TS U V" proof - have P1: "S C Le R A \ R A Le S C" by (simp add: local.le_cases) have P2: "S C Le R A \ P Q TS U V" by (simp add: assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l9_4_2_aux) have "R A Le S C \ P Q TS U V" by (simp add: assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) l9_2 l9_4_2_aux) then show ?thesis using P1 P2 by blast qed lemma l9_5: assumes "P Q TS A C" and "Col R P Q" and "R Out A B" shows "P Q TS B C" proof - have P1: "P \ Q" using assms(1) ts_distincts by blast obtain A' where P2: "Col P Q A' \ P Q Perp A A'" by (metis NCol_perm Tarski_neutral_dimensionless.TS_def Tarski_neutral_dimensionless_axioms assms(1) l8_18_existence) obtain C' where P3: "Col P Q C' \ P Q Perp C C'" using Col_perm TS_def assms(1) l8_18_existence by blast obtain M where P5: "M Midpoint A' C'" using midpoint_existence by blast obtain D where S2: "M Midpoint A D" using symmetric_point_construction by auto have "\ B0. Col P Q B0 \ P Q Perp B B0" proof - have S1: "\ Col P Q B" by (metis P2 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless.perp_not_col2 Tarski_neutral_dimensionless_axioms assms(2) assms(3) col_permutation_1 l6_3_1 out_col) then show ?thesis by (simp add: l8_18_existence) qed then obtain B' where P99: "Col P Q B' \ P Q Perp B B'" by blast have "P Q TS B C" proof - have S3: "C' Out D C \ A' Out A A" using Out_cases P2 P3 P5 S2 assms(1) l9_4_1 not_col_permutation_1 by blast then have S4: "C' Out D C" using P2 Tarski_neutral_dimensionless.perp_not_eq_2 Tarski_neutral_dimensionless_axioms out_trivial by fastforce have S5: "P Q TS A D" using P2 P3 S3 S4 assms(1) col_permutation_2 l9_4_2 by blast { assume "A' \ C'" then have "Col M P Q" by (smt P2 P3 P5 col_trivial_2 l6_21 midpoint_col not_col_permutation_1) then have "P Q TS B D" using S2 S5 assms(2) assms(3) l9_3 by blast } then have "A' \ C' \ P Q TS B D" by simp then have S6: "P Q TS B D" by (metis P3 P5 S2 S5 assms(2) assms(3) l9_3 midpoint_distinct_2 not_col_permutation_1) have S7: "Col B' P Q" using Col_perm P99 by blast have S8: "P Q Perp B B'" using P99 by blast have S9: "Col C' P Q" using Col_cases P3 by auto have S10: "P Q Perp D C'" by (metis Col_perm P3 S4 l6_3_1 out_col perp_col1 perp_right_comm) have S11: "B' Out B B" by (metis (no_types) P99 out_trivial perp_not_eq_2) have "C' Out C D" by (simp add: S4 l6_6) then show ?thesis using S6 S7 S8 S9 S10 S11 l9_4_2 by blast qed then show ?thesis using l8_18_existence by blast qed lemma outer_pasch_R1: assumes "Col P Q C" and "Bet A C P" and "Bet B Q C" shows "\ X. Bet A X B \ Bet P Q X" by (smt Bet_perm Col_def assms(1) assms(2) assms(3) between_exchange3 between_trivial outer_transitivity_between2) lemma outer_pasch_R2: assumes "\ Col P Q C" and "Bet A C P" and "Bet B Q C" shows "\ X. Bet A X B \ Bet P Q X" proof cases assume "B = Q" then show ?thesis using between_trivial by blast next assume P1: "B \ Q" have P2: "A \ P" using assms(1) assms(2) between_identity col_trivial_3 by blast have P3: "P \ Q" using assms(1) col_trivial_1 by blast have P4: "P \ B" using assms(1) assms(3) bet_col by blast have P5: "P Q TS C B" proof - have Q1: "\ Col C P Q" using Col_cases assms(1) by blast have Q2: "\ Col B P Q" by (metis Col_cases P1 Tarski_neutral_dimensionless.colx Tarski_neutral_dimensionless_axioms assms(1) assms(3) bet_col col_trivial_2) have "\ T. Col T P Q \ Bet C T B" using Col_cases assms(3) between_symmetry col_trivial_2 by blast then show ?thesis by (simp add: Q1 Q2 TS_def) qed have P6: "P Q TS A B" by (metis P5 assms(1) assms(2) bet_out_1 l9_5 not_col_distincts) obtain X where P7: "Col X P Q \ Bet A X B" using P6 TS_def by blast have "Bet P Q X" proof - obtain T where P8: "Bet X T P \ Bet C T B" using P7 assms(2) between_symmetry inner_pasch by blast have P9: "B \ C" using P1 assms(3) bet_neq12__neq by blast have P10: "T = Q" proof - have f1: "\p pa pb. Col pb pa p \ \ Bet pb pa p" by (meson bet_col1 between_trivial) then have f2: "Col Q C B" using NCol_cases assms(3) by blast have "Col T C B" using f1 NCol_cases P8 by blast then show ?thesis using f2 f1 by (metis (no_types) NCol_cases P7 P8 assms(1) between_trivial l6_16_1 l6_2 not_bet_and_out) qed then show ?thesis using P8 between_symmetry by blast qed then show ?thesis using P7 by blast qed lemma outer_pasch: assumes "Bet A C P" and "Bet B Q C" shows "\ X. Bet A X B \ Bet P Q X" using assms(1) assms(2) outer_pasch_R1 outer_pasch_R2 by blast lemma os_distincts: assumes "A B OS X Y" shows "A \ B \ A \ X \ A \ Y \ B \ X \ B \ Y" using OS_def assms ts_distincts by blast lemma invert_one_side: assumes "A B OS P Q" shows "B A OS P Q" proof - obtain T where "A B TS P T \ A B TS Q T" using OS_def assms by blast then have "B A TS P T \ B A TS Q T" using invert_two_sides by blast thus ?thesis using OS_def by blast qed lemma l9_8_1: assumes "P Q TS A C" and "P Q TS B C" shows "P Q OS A B" proof - have "\ R::'p. (P Q TS A R \ P Q TS B R)" using assms(1) assms(2) by blast then show ?thesis using OS_def by blast qed lemma not_two_sides_id: shows "\ P Q TS A A" using ts_distincts by blast lemma l9_8_2: assumes "P Q TS A C" and "P Q OS A B" shows "P Q TS B C" proof - obtain D where P1: "P Q TS A D \ P Q TS B D" using assms(2) OS_def by blast then have "P \ Q" using ts_distincts by blast obtain T where P2: "Col T P Q \ Bet A T C" using TS_def assms(1) by blast obtain X where P3: "Col X P Q \ Bet A X D" using TS_def P1 by blast obtain Y where P4: "Col Y P Q \ Bet B Y D" using TS_def P1 by blast then obtain M where P5: "Bet Y M A \ Bet X M B" using P3 inner_pasch by blast have P6: "A \ D" using P1 ts_distincts by blast have P7: "B \ D" using P1 not_two_sides_id by blast { assume Q0: "Col A B D" have "P Q TS B C" proof cases assume Q1: "M = Y" have "X = Y" proof - have S1: "\ Col P Q A" using TS_def assms(1) not_col_permutation_1 by blast have S3: "Col P Q X" using Col_perm P3 by blast have S4: "Col P Q Y" using Col_perm P4 by blast have S5: "Col A D X" by (simp add: P3 bet_col col_permutation_5) have "Col A D Y" by (metis Col_def P5 Q1 S5 Q0 between_equality between_trivial l6_16_1) then show ?thesis using S1 S3 S4 S5 P6 l6_21 by blast qed then have "X Out A B" by (metis P1 P3 P4 TS_def l6_2) then show ?thesis using assms(1) P3 l9_5 by blast next assume Z1: "\ M = Y" have "X = Y" proof - have S1: "\ Col P Q A" using TS_def assms(1) not_col_permutation_1 by blast have S3: "Col P Q X" using Col_perm P3 by blast have S4: "Col P Q Y" using Col_perm P4 by blast have S5: "Col A D X" by (simp add: P3 bet_col col_permutation_5) have "Col A D Y" by (metis Col_def P4 Q0 P7 l6_16_1) then show ?thesis using S1 S3 S4 S5 P6 l6_21 by blast qed then have Z3: "M \ X" using Z1 by blast have Z4: "P Q TS M C" by (meson Out_cases P4 P5 Tarski_neutral_dimensionless.l9_5 Tarski_neutral_dimensionless_axioms Z1 assms(1) bet_out) have "X Out M B" using P5 Z3 bet_out by auto then show ?thesis using Z4 P3 l9_5 by blast qed } then have Z99: "Col A B D \ P Q TS B C" by blast { assume Q0: "\ Col A B D" have Q1: "P Q TS M C" proof - have S3: "Y Out A M" proof - have T1: "A \ Y" using Col_def P4 Q0 col_permutation_4 by blast have T2: "M \ Y" proof - { assume T3: "M = Y" have "Col B D X" proof - have U1: "B \ M" using P1 P4 T3 TS_def by blast have U2: "Col B M D" by (simp add: P4 T3 bet_col) have "Col B M X" by (simp add: P5 bet_col between_symmetry) then show ?thesis using U1 U2 using col_transitivity_1 by blast qed have "False" by (metis NCol_cases P1 P3 TS_def \Col B D X\ Q0 bet_col col_trivial_2 l6_21) } then show ?thesis by blast qed have "Bet Y A M \ Bet Y M A" using P5 by blast then show ?thesis using T1 T2 by (simp add: Out_def) qed then have "X Out M B" by (metis P1 P3 P4 P5 TS_def bet_out l9_5) then show ?thesis using assms(1) S3 l9_5 P3 P4 by blast qed have "X Out M B" by (metis P3 P5 Q1 TS_def bet_out) then have "P Q TS B C" using Q1 P3 l9_5 by blast } then have "\ Col A B D \ P Q TS B C" by blast then show ?thesis using Z99 by blast qed lemma l9_9: assumes "P Q TS A B" shows "\ P Q OS A B" using assms l9_8_2 not_two_sides_id by blast lemma l9_9_bis: assumes "P Q OS A B" shows "\ P Q TS A B" using assms l9_9 by blast lemma one_side_chara: assumes "P Q OS A B" shows "\ X. Col X P Q \ \ Bet A X B" proof - have "\ Col A P Q \ \ Col B P Q" using OS_def TS_def assms by auto then show ?thesis using l9_9_bis TS_def assms by blast qed lemma l9_10: assumes "\ Col A P Q" shows "\ C. P Q TS A C" by (meson Col_perm assms mid_two_sides midpoint_existence symmetric_point_construction) lemma one_side_reflexivity: assumes "\ Col A P Q" shows "P Q OS A A" using assms l9_10 l9_8_1 by blast lemma one_side_symmetry: assumes "P Q OS A B" shows "P Q OS B A" by (meson Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless_axioms assms invert_two_sides) lemma one_side_transitivity: assumes "P Q OS A B" and "P Q OS B C" shows "P Q OS A C" by (meson Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless.l9_8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2)) lemma l9_17: assumes "P Q OS A C" and "Bet A B C" shows "P Q OS A B" proof cases assume "A = C" then show ?thesis using assms(1) assms(2) between_identity by blast next assume P1: "\ A = C" obtain D where P2: "P Q TS A D \ P Q TS C D" using OS_def assms(1) by blast then have P3: "P \ Q" using ts_distincts by blast obtain X where P4: "Col X P Q \ Bet A X D" using P2 TS_def by blast obtain Y where P5: "Col Y P Q \ Bet C Y D" using P2 TS_def by blast obtain T where P6: "Bet B T D \ Bet X T Y" using P4 P5 assms(2) l3_17 by blast have P7: "P Q TS A D" by (simp add: P2) have "P Q TS B D" proof - have Q1: "\ Col B P Q" using assms(1) assms(2) one_side_chara by blast have Q2: "\ Col D P Q" using P2 TS_def by blast obtain T0 where "Col T0 P Q \ Bet B T0 D" proof - assume a1: "\T0. Col T0 P Q \ Bet B T0 D \ thesis" obtain pp :: 'p where f2: "Bet B pp D \ Bet X pp Y" using \\thesis. (\T. Bet B T D \ Bet X T Y \ thesis) \ thesis\ by blast have "Col P Q Y" using Col_def P5 by blast then have "Y = X \ Col P Q pp" using f2 Col_def P4 colx by blast then show ?thesis using f2 a1 by (metis BetSEq BetS_def Col_def P4) qed then show ?thesis using Q1 Q2 using TS_def by blast qed then show ?thesis using P7 using OS_def by blast qed lemma l9_18_R1: assumes "Col X Y P" and "Col A B P" and "X Y TS A B" shows "Bet A P B \ \ Col X Y A \ \ Col X Y B" by (meson TS_def assms(1) assms(2) assms(3) col_permutation_5 l9_5 not_col_permutation_1 not_out_bet not_two_sides_id) lemma l9_18_R2: assumes "Col X Y P" and "Col A B P" and "Bet A P B" and "\ Col X Y A" and "\ Col X Y B" shows "X Y TS A B" using Col_perm TS_def assms(1) assms(3) assms(4) assms(5) by blast lemma l9_18: assumes "Col X Y P" and "Col A B P" shows "X Y TS A B \ (Bet A P B \ \ Col X Y A \ \ Col X Y B)" using l9_18_R1 l9_18_R2 assms(1) assms(2) by blast lemma l9_19_R1: assumes "Col X Y P" and "Col A B P" and "X Y OS A B" shows "P Out A B \ \ Col X Y A" by (meson OS_def TS_def assms(1) assms(2) assms(3) col_permutation_5 not_col_permutation_1 not_out_bet one_side_chara) lemma l9_19_R2: assumes "Col X Y P" and (* "Col A B P" and *) "P Out A B" and "\ Col X Y A" shows "X Y OS A B" proof - obtain D where "X Y TS A D" using Col_perm assms(3) l9_10 by blast then show ?thesis using OS_def assms(1) assms(2) l9_5 not_col_permutation_1 by blast qed lemma l9_19: assumes "Col X Y P" and "Col A B P" shows "X Y OS A B \ (P Out A B \ \ Col X Y A)" using l9_19_R1 l9_19_R2 assms(1) assms(2) by blast lemma one_side_not_col123: assumes "A B OS X Y" shows "\ Col A B X" using assms col_trivial_3 l9_19 by blast lemma one_side_not_col124: assumes "A B OS X Y" shows "\ Col A B Y" using assms one_side_not_col123 one_side_symmetry by blast lemma col_two_sides: assumes "Col A B C" and "A \ C" and "A B TS P Q" shows "A C TS P Q" using assms(1) assms(2) assms(3) col_preserves_two_sides col_trivial_3 by blast lemma col_one_side: assumes "Col A B C" and "A \ C" and "A B OS P Q" shows "A C OS P Q" proof - obtain T where "A B TS P T \ A B TS Q T" using assms(1) assms(2) assms(3) OS_def by blast then show ?thesis using col_two_sides OS_def assms(1) assms(2) by blast qed lemma out_out_one_side: assumes "A B OS X Y" and "A Out Y Z" shows "A B OS X Z" by (meson Col_cases Tarski_neutral_dimensionless.OS_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_trivial_3 l9_5) lemma out_one_side: assumes "\ Col A B X \ \ Col A B Y" and "A Out X Y" shows "A B OS X Y" using assms(1) assms(2) l6_6 not_col_permutation_2 one_side_reflexivity one_side_symmetry out_out_one_side by blast lemma bet__ts: assumes "A \ Y" and "\ Col A B X" and "Bet X A Y" shows "A B TS X Y" proof - have "\ Col Y A B" using NCol_cases assms(1) assms(2) assms(3) bet_col col2__eq by blast then show ?thesis by (meson TS_def assms(2) assms(3) col_permutation_3 col_permutation_5 col_trivial_3) qed lemma bet_ts__ts: assumes "A B TS X Y" and "Bet X Y Z" shows "A B TS X Z" proof - have "\ Col Z A B" using assms(1) assms(2) bet_col between_equality_2 col_permutation_1 l9_18 by blast then show ?thesis using TS_def assms(1) assms(2) between_exchange4 by blast qed lemma bet_ts__os: assumes "A B TS X Y" and "Bet X Y Z" shows "A B OS Y Z" using OS_def assms(1) assms(2) bet_ts__ts l9_2 by blast lemma l9_31 : assumes "A X OS Y Z" and "A Z OS Y X" shows "A Y TS X Z" proof - have P1: "A \ X \ A \ Z \ \ Col Y A X \ \ Col Z A X \ \ Col Y A Z" using assms(1) assms(2) col_permutation_1 one_side_not_col123 one_side_not_col124 os_distincts by blast obtain Z' where P2: "Bet Z A Z' \ Cong A Z' Z A" using segment_construction by blast have P3: "Z' \ A" using P1 P2 cong_diff_4 by blast have P4: "A X TS Y Z'" by (metis (no_types) P2 P3 assms(1) bet__ts l9_8_2 one_side_not_col124 one_side_symmetry) have P5: "\ Col Y A X" using P1 by blast obtain T where P6: "Col A T X \ Bet Y T Z'" using P4 TS_def not_col_permutation_4 by blast then have P7: "T \ A" proof - have "\ Col A Z Y" by (simp add: P1 not_col_permutation_1) then have f1: "\ A Out Z Y" using out_col by blast have "A \ Z'" using P1 P2 cong_diff_4 by blast then show ?thesis using f1 by (metis (no_types) P1 P2 P6 l6_2) qed have P8: "Y A OS Z' T" by (smt P1 P2 P3 P6 Tarski_neutral_dimensionless.l6_6 Tarski_neutral_dimensionless_axioms bet_col bet_out col_trivial_2 l6_21 not_col_permutation_1 out_one_side) have P9: "A Y TS Z' Z" using Col_perm P1 P2 P8 bet__ts between_symmetry one_side_not_col123 by blast { assume Q0: "Bet T A X" have Q1: "Z' Z OS Y T" by (metis BetSEq BetS_def P1 P2 P4 P6 TS_def Tarski_neutral_dimensionless.l6_6 Tarski_neutral_dimensionless_axioms bet_col bet_out_1 col_trivial_3 colx not_col_permutation_3 not_col_permutation_4 out_one_side) then have Q2: "Z' Out T Y" by (metis P6 bet_out_1 os_distincts) then have Q3: "A Z OS Y T" by (meson Out_cases P1 P2 P6 bet_col col_permutation_3 invert_one_side l9_19_R2) have "A Z TS X T" proof - have R1: "\ Col X A Z" using P1 col_permutation_3 by blast have R2: "\ Col T A Z" using Q3 between_trivial one_side_chara by blast have "\ T0. Col T0 A Z \ Bet X T0 T" proof - have S1: "Col A A Z" by (simp add: col_trivial_1) have "Bet X A T" by (simp add: Q0 between_symmetry) then show ?thesis using S1 by blast qed then show ?thesis using R1 R2 using TS_def by auto qed have "A Y TS X Z" by (meson Q3 Tarski_neutral_dimensionless.l9_8_2 Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms \A Z TS X T\ assms(2) l9_9_bis) } then have P10: "Bet T A X \ A Y TS X Z" by blast { assume R1: "Bet A X T" then have R3: "A Y OS Z' X" by (meson Bet_cases P1 P6 P8 R1 between_equality invert_one_side not_col_permutation_4 not_out_bet out_out_one_side) have "A Y TS X Z" using R3 P9 l9_8_2 by blast } then have P11: "Bet A X T \ A Y TS X Z" by blast { assume R1: "Bet X T A" then have R3: "A Y OS T X" by (simp add: P5 P7 R1 bet_out_1 not_col_permutation_4 out_one_side) then have "A Y TS X Z" using P8 P9 invert_two_sides l9_8_2 by blast } then have "Bet X T A \ A Y TS X Z" by blast then show ?thesis using P10 P11 using P6 between_symmetry third_point by blast qed lemma col123__nos: assumes "Col P Q A" shows "\ P Q OS A B" using assms one_side_not_col123 by blast lemma col124__nos: assumes "Col P Q B" shows "\ P Q OS A B" using assms one_side_not_col124 by blast lemma col2_os__os: assumes "C \ D" and "Col A B C" and "Col A B D" and "A B OS X Y" shows "C D OS X Y" by (metis assms(1) assms(2) assms(3) assms(4) col3 col_one_side col_trivial_3 invert_one_side os_distincts) lemma os_out_os: assumes "Col A B P" and "A B OS C D" and "P Out C C'" shows "A B OS C' D" using OS_def assms(1) assms(2) assms(3) l9_5 not_col_permutation_1 by blast lemma ts_ts_os: assumes "A B TS C D" and "C D TS A B" shows "A C OS B D" proof - obtain T1 where P1: "Col T1 A B \ Bet C T1 D" using TS_def assms(1) by blast obtain T where P2: "Col T C D \ Bet A T B" using TS_def assms(2) by blast have P3: "T1 = T" proof - have "A \ B" using assms(2) ts_distincts by blast then show ?thesis proof - have "Col T1 D C" using Col_def P1 by blast then have f1: "\p. (C = T1 \ Col C p T1) \ \ Col C T1 p" by (metis assms(1) col_transitivity_1 l6_16_1 ts_distincts) have f2: "\ Col C A B" using TS_def assms(1) by presburger have f3: "(Bet B T1 A \ Bet T1 A B) \ Bet A B T1" using Col_def P1 by blast { assume "T1 \ B" then have "C \ T1 \ \ Col C T1 B \ (\p. \ Col p T1 B \ Col p T1 T) \ T \ A \ T \ B" using f3 f2 by (metis (no_types) Col_def col_transitivity_1 l6_16_1) then have "T \ A \ T \ B \ C \ T1 \ \ Col C T1 T \ T1 = T" using f3 by (meson Col_def l6_16_1) } moreover { assume "T \ A \ T \ B" then have "C \ T1 \ \ Col C T1 T \ T1 = T" using f2 by (metis (no_types) Col_def P1 P2 \A \ B\ col_transitivity_1 l6_16_1) } ultimately have "C \ T1 \ \ Col C T1 T \ T1 = T" using f2 f1 assms(1) ts_distincts by blast then show ?thesis by (metis (no_types) Col_def P1 P2 assms(1) l6_16_1 ts_distincts) qed qed have P4: "A C OS T B" by (metis Col_cases P2 TS_def assms(1) assms(2) bet_out out_one_side) then have "C A OS T D" by (metis Col_cases P1 TS_def P3 assms(2) bet_out os_distincts out_one_side) then show ?thesis by (meson P4 Tarski_neutral_dimensionless.invert_one_side Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms one_side_transitivity) qed lemma col_one_side_out: assumes "Col A X Y" and "A B OS X Y" shows "A Out X Y" by (meson assms(1) assms(2) l6_4_2 not_col_distincts not_col_permutation_4 one_side_chara) lemma col_two_sides_bet: assumes "Col A X Y" and "A B TS X Y" shows "Bet X A Y" using Col_cases assms(1) assms(2) l9_8_1 l9_9 or_bet_out out_out_one_side by blast lemma os_ts1324__os: assumes "A X OS Y Z" and "A Y TS X Z" shows "A Z OS X Y" proof - obtain P where P1: "Col P A Y \ Bet X P Z" using TS_def assms(2) by blast have P2: "A Z OS X P" by (metis Col_cases P1 TS_def assms(1) assms(2) bet_col bet_out_1 col124__nos col_trivial_2 l6_6 l9_19) have "A Z OS P Y" proof - have "\ Col A Z P \ \ Col A Z Y" using P2 col124__nos by blast moreover have "A Out P Y" proof - have "X A OS P Z" by (metis Col_cases P1 P2 assms(1) bet_out col123__nos out_one_side) then have "A X OS P Y" by (meson Tarski_neutral_dimensionless.invert_one_side Tarski_neutral_dimensionless.one_side_symmetry Tarski_neutral_dimensionless_axioms assms(1) one_side_transitivity) then show ?thesis using P1 col_one_side_out not_col_permutation_4 by blast qed ultimately show ?thesis by (simp add: out_one_side) qed then show ?thesis using P2 one_side_transitivity by blast qed lemma ts2__ex_bet2: assumes "A C TS B D" and "B D TS A C" shows "\ X. Bet A X C \ Bet B X D" by (metis TS_def assms(1) assms(2) bet_col col_permutation_5 l9_18_R1 not_col_permutation_2) lemma out_one_side_1: assumes "\ Col A B C" and "Col A B X" and "X Out C D" shows "A B OS C D" using assms(1) assms(2) assms(3) not_col_permutation_2 one_side_reflexivity one_side_symmetry os_out_os by blast lemma out_two_sides_two_sides: assumes (*"A \ PX" and *) "Col A B PX" and "PX Out X P" and "A B TS P Y" shows "A B TS X Y" using assms(1) assms(2) assms(3) l6_6 l9_5 not_col_permutation_1 by blast lemma l8_21_bis: assumes "X \ Y" and "\ Col C A B" shows "\ P. Cong A P X Y \ A B Perp P A \ A B TS C P" proof - have P1: "A \ B" using assms(2) not_col_distincts by blast then have "\ P T. A B Perp P A \ Col A B T \ Bet C T P" using l8_21 by auto then obtain P T where P2: "A B Perp P A \ Col A B T \ Bet C T P" by blast have P3: "A B TS C P" proof - have "\ Col P A B" using P2 col_permutation_1 perp_not_col by blast then show ?thesis using P2 TS_def assms(2) not_col_permutation_1 by blast qed have P4: "P \ A" using P3 ts_distincts by blast obtain P' where P5: "(Bet A P P' \ Bet A P' P) \ Cong A P' X Y" using segment_construction_2 P4 by blast have P6: "A B Perp P' A" by (smt P2 P5 Perp_perm assms(1) bet_col cong_identity cong_symmetry not_bet_distincts not_col_permutation_2 perp_col2) have P7: "\ Col P' A B" using NCol_perm P6 col_trivial_3 l8_16_1 by blast then have P8: "A B OS P P'" by (metis Out_def P4 P5 P6 col_permutation_2 out_one_side perp_not_eq_2) then have P9: "A B TS C P'" using P3 l9_2 l9_8_2 by blast then show ?thesis using P5 P6 by blast qed lemma ts__ncol: assumes "A B TS X Y" shows "\ Col A X Y \ \ Col B X Y" by (metis TS_def assms col_permutation_1 col_transitivity_2 ts_distincts) lemma one_or_two_sides_aux: assumes "\ Col C A B" and "\ Col D A B" and "Col A C X" and "Col B D X" shows "A B TS C D \ A B OS C D" proof - have P1: "A \ X" using assms(2) assms(4) col_permutation_2 by blast have P2: "B \ X" using assms(1) assms(3) col_permutation_4 by blast have P3: "\ Col X A B" using P1 assms(1) assms(3) col_permutation_5 col_transitivity_1 not_col_permutation_4 by blast { assume Q0: "Bet A C X \ Bet B D X" then have Q1: "A B OS C X" using assms(1) bet_out not_col_distincts not_col_permutation_1 out_one_side by blast then have "A B OS X D" by (metis Q0 assms(2) assms(4) bet_out_1 col_permutation_2 col_permutation_3 invert_one_side l6_4_2 not_bet_and_out not_col_distincts out_one_side) then have "A B OS C D" using Q1 one_side_transitivity by blast } then have P4: "Bet A C X \ Bet B D X \ A B OS C D" by blast { assume "Bet A C X \ Bet D X B" then have "A B OS C D" by (smt P2 assms(1) assms(4) bet_out between_equality_2 l9_10 l9_5 l9_8_1 not_bet_and_out not_col_distincts not_col_permutation_4 out_to_bet out_two_sides_two_sides) } then have P5: "Bet A C X \ Bet D X B \ A B OS C D " by blast { assume Q0: "Bet A C X \ Bet X B D" have Q1: "A B TS X D" using P3 Q0 TS_def assms(2) col_trivial_3 by blast have "A B OS X C" using Q0 assms(1) bet_out not_col_distincts one_side_reflexivity one_side_symmetry out_out_one_side by blast then have "A B TS C D" using Q1 l9_8_2 by blast } then have P6: "Bet A C X \ Bet X B D \ A B TS C D" by blast { assume Q1: "Bet C X A \ Bet B D X" then have Q2: "A B OS C X" using P1 assms(1) assms(3) between_equality_2 l6_4_2 not_col_permutation_1 not_col_permutation_4 out_one_side by blast have "A B OS X D" using Q1 assms(2) bet_out not_col_distincts one_side_reflexivity os_out_os by blast then have "A B OS C D" using Q2 using one_side_transitivity by blast } then have P7: "Bet C X A \ Bet B D X \ A B OS C D" by blast { assume "Bet C X A \ Bet D X B" then have "A B OS C D" by (smt \Bet A C X \ Bet D X B \ A B OS C D\ \Bet C X A \ Bet B D X \ A B OS C D\ assms(1) assms(2) assms(3) assms(4) between_symmetry l6_21 l9_18_R2 not_col_distincts ts_ts_os) } then have P8: "Bet C X A \ Bet D X B \ A B OS C D" by blast { assume Q1: "Bet C X A \ Bet X B D" have Q2: "A B TS X D" by (metis P3 Q1 assms(2) bet__ts invert_two_sides not_col_distincts not_col_permutation_3) have Q3: "A B OS X C" using P1 Q1 assms(1) bet_out_1 not_col_permutation_1 out_one_side by auto then have "A B TS C D" using Q2 l9_8_2 by blast } then have P9: "Bet C X A \ Bet X B D \ A B TS C D" by blast { assume Q0: "Bet X A C \ Bet B D X" have Q1: "A B TS X C" by (metis P3 Q0 assms(1) bet__ts col_permutation_2 not_col_distincts) have "A B OS X D" by (metis NCol_cases Q0 Tarski_neutral_dimensionless.out_one_side Tarski_neutral_dimensionless_axioms assms(2) assms(4) bet_out_1 invert_one_side l6_4_1 not_col_distincts not_out_bet) then have "A B TS C D" using Q1 l9_2 l9_8_2 by blast } then have P10: "Bet X A C \ Bet B D X \ A B TS C D" by blast { assume Q0: "Bet X A C \ Bet D X B" have Q1: "A B TS X C" by (metis NCol_cases P3 Q0 assms(1) bet__ts not_col_distincts) have "A B OS X D" by (metis P2 P3 Q0 bet_out_1 col_permutation_3 invert_one_side out_one_side) then have "A B TS C D" using Q1 l9_2 l9_8_2 by blast } then have P11: "Bet X A C \ Bet D X B \ A B TS C D" by blast { assume Q0: "Bet X A C \ Bet X B D" then have Q1: "A B TS C X" by (simp add: P1 Q0 assms(1) bet__ts between_symmetry not_col_permutation_1) have "A B TS D X" by (simp add: P2 Q0 assms(2) bet__ts between_symmetry invert_two_sides not_col_permutation_3) then have "A B OS C D" using Q1 l9_8_1 by blast } then have P12: "Bet X A C \ Bet X B D \ A B OS C D" by blast then show ?thesis using P4 P5 P6 P7 P8 P9 P10 P11 using Col_def assms(3) assms(4) by auto qed lemma cop__one_or_two_sides: assumes "Coplanar A B C D" and "\ Col C A B" and "\ Col D A B" shows "A B TS C D \ A B OS C D" proof - obtain X where P1: "Col A B X \ Col C D X \ Col A C X \ Col B D X \ Col A D X \ Col B C X" using Coplanar_def assms(1) by auto have P2: "Col A B X \ Col C D X \ A B TS C D \ A B OS C D" by (metis TS_def Tarski_neutral_dimensionless.l9_19_R2 Tarski_neutral_dimensionless_axioms assms(2) assms(3) not_col_permutation_3 not_col_permutation_5 not_out_bet) have P3: "Col A C X \ Col B D X \ A B TS C D \ A B OS C D" using assms(2) assms(3) one_or_two_sides_aux by blast have "Col A D X \ Col B C X \ A B TS C D \ A B OS C D" using assms(2) assms(3) l9_2 one_or_two_sides_aux one_side_symmetry by blast then show ?thesis using P1 P2 P3 by blast qed lemma os__coplanar: assumes "A B OS C D" shows "Coplanar A B C D" proof - have P1: "\ Col A B C" using assms one_side_not_col123 by blast obtain C' where P2: "Bet C B C' \ Cong B C' B C" using segment_construction by presburger have P3: "A B TS D C'" by (metis (no_types) Cong_perm OS_def P2 TS_def assms bet__ts bet_cong_eq invert_one_side l9_10 l9_8_2 one_side_not_col123 ts_distincts) obtain T where P4: "Col T A B \ Bet D T C'" using P3 TS_def by blast have P5: "C' \ T" using P3 P4 TS_def by blast have P6: "Col T B C \ Coplanar A B C D" by (metis Col_def Coplanar_def P2 P4 P5 col_trivial_2 l6_16_1) { assume Q0: "\ Col T B C" { assume R0: "Bet T B A" have S1: "B C TS T A" by (metis P1 Q0 R0 bet__ts col_permutation_2 not_col_distincts) have "C' Out T D" using P4 P5 bet_out_1 by auto then have "B C OS T D" using P2 Q0 bet_col invert_one_side not_col_permutation_3 out_one_side_1 by blast then have R1: "B C TS D A" using S1 l9_8_2 by blast then have "Coplanar A B C D" using ncoplanar_perm_9 ts__coplanar by blast } then have Q1: "Bet T B A \ Coplanar A B C D" by blast { assume R0: "\ Bet T B A" { have R2: "B C OS D T" proof - have S1: "\ Col B C D" by (metis Col_perm P2 P3 P4 Q0 bet_col colx ts_distincts) have S2: "Col B C C'" by (simp add: P2 bet_col col_permutation_4) have S3: "C' Out D T" using P4 P5 bet_out_1 l6_6 by auto then show ?thesis using S1 S2 out_one_side_1 by blast qed have R3: "B C OS T A" using P4 Q0 R0 col_permutation_2 col_permutation_5 not_bet_out out_one_side by blast } then have R1: "B C OS D A" by (metis P2 P4 Q0 bet_col bet_out_1 col_permutation_2 col_permutation_5 os_out_os) then have "Coplanar A B C D" by (simp add: R1 assms coplanar_perm_19 invert_one_side l9_31 one_side_symmetry ts__coplanar) } then have "\ Bet T B A \ Coplanar A B C D" by blast then have "Coplanar A B C D" using Q1 by blast } then have "\ Col T B C \ Coplanar A B C D" by blast then show ?thesis using P6 by blast qed lemma coplanar_trans_1: assumes "\ Col P Q R" and "Coplanar P Q R A" and "Coplanar P Q R B" shows "Coplanar Q R A B" proof - have P1: "Col Q R A \ Coplanar Q R A B" by (simp add: col__coplanar) { assume T1: "\ Col Q R A" { assume T2: "\ Col Q R B" { have "Col Q A B \ Coplanar Q R A B" using ncop__ncols by blast { assume S1: "\ Col Q A B" have U1: "Q R TS P A \ Q R OS P A" by (simp add: T1 assms(1) assms(2) cop__one_or_two_sides coplanar_perm_8 not_col_permutation_2) have U2: "Q R TS P B \ Q R OS P B" using T2 assms(1) assms(3) col_permutation_1 cop__one_or_two_sides coplanar_perm_8 by blast have W1: "Q R TS P A \ Q R OS P A \ Q R TS A B \ Q R OS A B" using l9_9 by blast have W2: "Q R TS P A \ Q R OS P B \ Q R TS A B \ Q R OS A B" using l9_2 l9_8_2 by blast have W3: "Q R TS P B \ Q R OS P A \ Q R TS A B \ Q R OS A B" using l9_8_2 by blast have "Q R TS P B \ Q R OS P B \ Q R TS A B \ Q R OS A B" using l9_9 by blast then have S2: "Q R TS A B \ Q R OS A B" using U1 U2 W1 W2 W3 using OS_def l9_2 one_side_transitivity by blast have "Coplanar Q R A B" using S2 os__coplanar ts__coplanar by blast } then have "\ Col Q A B \ Coplanar Q R A B" by blast } then have "Coplanar Q R A B" using ncop__ncols by blast } then have "\ Col Q R B \ Coplanar Q R A B" by blast } then have "\ Col Q R A \ Coplanar Q R A B" using ncop__ncols by blast then show ?thesis using P1 by blast qed lemma col_cop__cop: assumes "Coplanar A B C D" and "C \ D" and "Col C D E" shows "Coplanar A B C E" proof - have "Col D A C \ Coplanar A B C E" by (meson assms(2) assms(3) col_permutation_1 l6_16_1 ncop__ncols) moreover { assume "\ Col D A C" then have "Coplanar A C B E" by (meson assms(1) assms(3) col__coplanar coplanar_trans_1 ncoplanar_perm_11 ncoplanar_perm_13) then have "Coplanar A B C E" using ncoplanar_perm_2 by blast } ultimately show ?thesis by blast qed lemma bet_cop__cop: assumes "Coplanar A B C E" and "Bet C D E" shows "Coplanar A B C D" by (metis NCol_perm Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_col bet_neq12__neq) lemma col2_cop__cop: assumes "Coplanar A B C D" and "C \ D" and "Col C D E" and "Col C D F" shows "Coplanar A B E F" proof cases assume "C = E" then show ?thesis using assms(1) assms(2) assms(4) col_cop__cop by blast next assume "C \ E" then show ?thesis by (metis assms(1) assms(2) assms(3) assms(4) col_cop__cop col_transitivity_1 ncoplanar_perm_1 not_col_permutation_4) qed lemma col_cop2__cop: assumes "U \ V" and "Coplanar A B C U" and "Coplanar A B C V" and "Col U V P" shows "Coplanar A B C P" proof cases assume "Col A B C" then show ?thesis using ncop__ncol by blast next assume "\ Col A B C" then show ?thesis by (smt Col_perm assms(1) assms(2) assms(3) assms(4) col_cop__cop coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_14 ncoplanar_perm_15 ncoplanar_perm_23) qed lemma bet_cop2__cop: assumes "Coplanar A B C U" and "Coplanar A B C W" and "Bet U V W" shows "Coplanar A B C V" proof - have "Col U V W" using assms(3) bet_col by blast then have "Col U W V" by (meson not_col_permutation_5) then show ?thesis using assms(1) assms(2) assms(3) bet_neq23__neq col_cop2__cop by blast qed lemma coplanar_pseudo_trans: assumes "\ Col P Q R" and "Coplanar P Q R A" and "Coplanar P Q R B" and "Coplanar P Q R C" and "Coplanar P Q R D" shows "Coplanar A B C D" proof cases have LEM1: "(\ Col P Q R \ Coplanar P Q R A \ Coplanar P Q R B \ Coplanar P Q R C) \ Coplanar A B C R" by (smt col_transitivity_2 coplanar_trans_1 ncop__ncols ncoplanar_perm_19 ncoplanar_perm_21) assume P2: "Col P Q D" have P3: "P \ Q" using assms(1) col_trivial_1 by blast have P4: "Coplanar A B C Q" by (smt assms(1) assms(2) assms(3) assms(4) col2_cop__cop coplanar_trans_1 ncoplanar_perm_9 not_col_distincts) have P5: "\ Col Q R P" using Col_cases assms(1) by blast have P6: "Coplanar Q R P A" using assms(2) ncoplanar_perm_12 by blast have P7: "Coplanar Q R P B" using assms(3) ncoplanar_perm_12 by blast have P8: "Coplanar Q R P C" using assms(4) ncoplanar_perm_12 by blast then have "Coplanar A B C P" using LEM1 P5 P6 P7 by (smt col_transitivity_2 coplanar_trans_1 ncop__ncols ncoplanar_perm_19) then show ?thesis using LEM1 P2 P3 P4 col_cop2__cop by blast next assume P9: "\ Col P Q D" have P10: "Coplanar P Q D A" using NCol_cases assms(1) assms(2) assms(5) coplanar_trans_1 ncoplanar_perm_8 by blast have P11: "Coplanar P Q D B" using assms(1) assms(3) assms(5) col_permutation_1 coplanar_perm_12 coplanar_trans_1 by blast have "Coplanar P Q D C" by (meson assms(1) assms(4) assms(5) coplanar_perm_7 coplanar_trans_1 ncoplanar_perm_14 not_col_permutation_3) then show ?thesis using P9 P10 P11 by (smt P10 P11 P9 col3 coplanar_trans_1 ncop__ncol ncoplanar_perm_20 ncoplanar_perm_23 not_col_distincts) qed lemma l9_30: assumes "\ Coplanar A B C P" and "\ Col D E F" and "Coplanar D E F P" and "Coplanar A B C X" and "Coplanar A B C Y" and "Coplanar A B C Z" and "Coplanar D E F X" and "Coplanar D E F Y" and "Coplanar D E F Z" shows "Col X Y Z" proof - { assume P1: "\ Col X Y Z" have P2: "\ Col A B C" using assms(1) col__coplanar by blast have "Coplanar A B C P" proof - have Q2: "Coplanar X Y Z A" by (smt P2 assms(4) assms(5) assms(6) col2_cop__cop coplanar_trans_1 ncoplanar_perm_18 not_col_distincts) have Q3: "Coplanar X Y Z B" using P2 assms(4) assms(5) assms(6) col_trivial_3 coplanar_pseudo_trans ncop__ncols by blast have Q4: "Coplanar X Y Z C" using P2 assms(4) assms(5) assms(6) col_trivial_2 coplanar_pseudo_trans ncop__ncols by blast have "Coplanar X Y Z P" using assms(2) assms(3) assms(7) assms(8) assms(9) coplanar_pseudo_trans by blast then show ?thesis using P1 Q2 Q3 Q4 using assms(2) assms(3) assms(7) assms(8) assms(9) coplanar_pseudo_trans by blast qed then have "False" using assms(1) by blast } then show ?thesis by blast qed lemma cop_per2__col: assumes "Coplanar A X Y Z" and "A \ Z" and "Per X Z A" and "Per Y Z A" shows "Col X Y Z" proof cases assume "X = Y \ X = Z \ Y = Z" then show ?thesis using not_col_distincts by blast next assume H1:"\ (X = Y \ X = Z \ Y = Z)" obtain B where P1: "Cong X A X B \ Z Midpoint A B \ Cong Y A Y B" using Per_def assms(3) assms(4) per_double_cong by blast have P2: "X \ Y" using H1 by blast have P3: "X \ Z" using H1 by blast have P4: "Y \ Z" using H1 by blast obtain I where P5: " Col A X I \ Col Y Z I \ Col A Y I \ Col X Z I \ Col A Z I \ Col X Y I" using Coplanar_def assms(1) by auto have P6: "Col A X I \ Col Y Z I \ Col X Y Z" by (smt P1 P4 assms(2) l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2 not_col_permutation_1) have P7: "Col A Y I \ Col X Z I \ Col X Y Z" by (smt P1 P3 assms(2) col_permutation_1 col_permutation_5 l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2) have "Col A Z I \ Col X Y I \ Col X Y Z" by (metis P1 P2 assms(2) col_permutation_1 l4_17 l4_18 l7_13 l7_2 l7_3_2 midpoint_distinct_2) then show ?thesis using P5 P6 P7 by blast qed lemma cop_perp2__col: assumes "Coplanar A B Y Z" and "X Y Perp A B" and "X Z Perp A B" shows "Col X Y Z" proof cases assume P1: "Col A B X" { assume Q0: "X = A" then have Q1: "X \ B" using assms(3) perp_not_eq_2 by blast have Q2: "Coplanar B Y Z X" by (simp add: Q0 assms(1) coplanar_perm_9) have Q3: "Per Y X B" using Q0 assms(2) perp_per_2 by blast have "Per Z X B" using Q0 assms(3) perp_per_2 by blast then have "Col X Y Z" using Q1 Q2 Q3 cop_per2__col not_col_permutation_1 by blast } then have P2: "X = A \ Col X Y Z" by blast { assume Q0: "X \ A" have Q1: "A X Perp X Y" by (metis P1 Perp_perm Q0 assms(2) perp_col1) have Q2: "A X Perp X Z" by (metis P1 Perp_perm Q0 assms(3) perp_col1) have Q3: "Coplanar A Y Z X" by (smt P1 assms(1) assms(2) col2_cop__cop col_trivial_3 coplanar_perm_12 coplanar_perm_16 perp_distinct) have Q4: "Per Y X A" using Perp_perm Q1 perp_per_2 by blast have "Per Z X A" using P1 Q0 assms(3) perp_col1 perp_per_1 by auto then have "Col X Y Z" using Q0 Q3 Q4 cop_per2__col not_col_permutation_1 by blast } then have "X \ A \ Col X Y Z" by blast then show ?thesis using P2 by blast next assume P1: "\ Col A B X" obtain Y0 where P2: "Y0 PerpAt X Y A B" using Perp_def assms(2) by blast obtain Z0 where P3: "Z0 PerpAt X Z A B" using Perp_def assms(3) by auto have P4: "X Y0 Perp A B" by (metis P1 P2 assms(2) perp_col perp_in_col) have P5: "X Z0 Perp A B" by (metis P1 P3 assms(3) perp_col perp_in_col) have P6: "Y0 = Z0" by (meson P1 P2 P3 P4 P5 Perp_perm l8_18_uniqueness perp_in_col) have P7: "X \ Y0" using P4 perp_not_eq_1 by blast have P8: "Col X Y0 Y" using P2 col_permutation_5 perp_in_col by blast have "Col X Y0 Z" using P3 P6 col_permutation_5 perp_in_col by blast then show ?thesis using P7 P8 col_transitivity_1 by blast qed lemma two_sides_dec: shows "A B TS C D \ \ A B TS C D" by simp lemma cop_nts__os: assumes "Coplanar A B C D" and "\ Col C A B" and "\ Col D A B" and "\ A B TS C D" shows "A B OS C D" using assms(1) assms(2) assms(3) assms(4) cop__one_or_two_sides by blast lemma cop_nos__ts: assumes "Coplanar A B C D" and "\ Col C A B" and "\ Col D A B" and "\ A B OS C D" shows "A B TS C D" using assms(1) assms(2) assms(3) assms(4) cop_nts__os by blast lemma one_side_dec: "A B OS C D \ \ A B OS C D" by simp lemma cop_dec: "Coplanar A B C D \ \ Coplanar A B C D" by simp lemma ex_diff_cop: "\ E. Coplanar A B C E \ D \ E" by (metis col_trivial_2 diff_col_ex ncop__ncols) lemma ex_ncol_cop: assumes "D \ E" shows "\ F. Coplanar A B C F \ \ Col D E F" proof cases assume "Col A B C" then show ?thesis using assms ncop__ncols not_col_exists by blast next assume P1: "\ Col A B C" then show ?thesis proof - have P2: "(Col D E A \ Col D E B) \ (\ F. Coplanar A B C F \ \ Col D E F)" by (meson P1 assms col3 col_trivial_2 ncop__ncols) have P3: "(\Col D E A \ Col D E B) \ (\ F. Coplanar A B C F \ \ Col D E F)" using col_trivial_3 ncop__ncols by blast have P4: "(Col D E A \ \Col D E B) \ (\ F. Coplanar A B C F \ \ Col D E F)" using col_trivial_2 ncop__ncols by blast have "(\Col D E A \ \Col D E B) \ (\ F. Coplanar A B C F \ \ Col D E F)" using col_trivial_3 ncop__ncols by blast then show ?thesis using P2 P3 P4 by blast qed qed lemma ex_ncol_cop2: "\ E F. (Coplanar A B C E \ Coplanar A B C F \ \ Col D E F)" proof - have f1: "\p pa pb. Coplanar pb pa p pb" by (meson col_trivial_3 ncop__ncols) have f2: "\p pa pb. Coplanar pb pa p p" by (meson Col_perm col_trivial_3 ncop__ncols) obtain pp :: "'p \ 'p \ 'p" where f3: "\p pa. p = pa \ \ Col p pa (pp p pa)" using not_col_exists by moura have f4: "\p pa pb. Coplanar pb pb pa p" by (meson Col_perm col_trivial_3 ncop__ncols) have "\p. A \ p" by (meson col_trivial_3 diff_col_ex3) moreover { assume "B \ A" then have "D = B \ (\p. \ Col D p A \ Coplanar A B C p)" using f3 f2 by (metis (no_types) Col_perm ncop__ncols) then have "D = B \ (\p pa. Coplanar A B C p \ Coplanar A B C pa \ \ Col D p pa)" using f1 by blast } moreover { assume "D \ B" moreover { assume "\p. D \ B \ \ Coplanar A B C p" then have "D \ B \ \ Col A B C" using ncop__ncols by blast then have "\p. \ Col D p B \ Coplanar A B C p" using f2 f1 by (metis (no_types) Col_perm col_transitivity_1) } ultimately have ?thesis using f3 by (metis (no_types) col_trivial_3 ncop__ncols) } ultimately show ?thesis using f4 f3 by blast qed lemma col2_cop2__eq: assumes "\ Coplanar A B C U" and "U \ V" and "Coplanar A B C P" and "Coplanar A B C Q" and "Col U V P" and "Col U V Q" shows "P = Q" proof - have "Col U Q P" by (meson assms(2) assms(5) assms(6) col_transitivity_1) then have "Col P Q U" using not_col_permutation_3 by blast then show ?thesis using assms(1) assms(3) assms(4) col_cop2__cop by blast qed lemma cong3_cop2__col: assumes "Coplanar A B C P" and "Coplanar A B C Q" and "P \ Q" and "Cong A P A Q" and "Cong B P B Q" and "Cong C P C Q" shows "Col A B C" proof cases assume "Col A B C" then show ?thesis by blast next assume P1: "\ Col A B C" obtain M where P2: "M Midpoint P Q" using assms(6) l7_25 by blast have P3: "Per A M P" using P2 Per_def assms(4) by blast have P4: "Per B M P" using P2 Per_def assms(5) by blast have P5: "Per C M P" using P2 Per_def assms(6) by blast have "False" proof cases assume Q1: "A = M" have Q2: "Coplanar P B C A" using assms(1) ncoplanar_perm_21 by blast have Q3: "P \ A" by (metis assms(3) assms(4) cong_diff_4) have Q4: "Per B A P" by (simp add: P4 Q1) have Q5: "Per C A P" by (simp add: P5 Q1) then show ?thesis using Q1 Q2 Q3 Q4 cop_per2__col using P1 not_col_permutation_1 by blast next assume Q0: "A \ M" have Q1: "Col A B M" proof - have R1: "Coplanar A B P Q" using P1 assms(1) assms(2) coplanar_trans_1 ncoplanar_perm_8 not_col_permutation_2 by blast then have R2: "Coplanar P A B M" using P2 bet_cop__cop coplanar_perm_14 midpoint_bet ncoplanar_perm_6 by blast have R3: "P \ M" using P2 assms(3) l7_3_2 l7_9_bis by blast have R4: "Per A M P" by (simp add: P3) have R5: "Per B M P" by (simp add: P4) then show ?thesis using R2 R3 R4 cop_per2__col by blast qed have "Col A C M" proof - have R1: "Coplanar P A C M" using P1 Q1 assms(1) col2_cop__cop coplanar_perm_22 ncoplanar_perm_3 not_col_distincts by blast have R2: "P \ M" using P2 assms(3) l7_3_2 symmetric_point_uniqueness by blast have R3: "Per A M P" by (simp add: P3) have "Per C M P" by (simp add: P5) then show ?thesis using R1 R2 R3 cop_per2__col by blast qed then show ?thesis using NCol_perm P1 Q0 Q1 col_trivial_3 colx by blast qed then show ?thesis by blast qed lemma l9_38: assumes "A B C TSP P Q" shows "A B C TSP Q P" using Bet_perm TSP_def assms by blast lemma l9_39: assumes "A B C TSP P R" and "Coplanar A B C D" and "D Out P Q" shows "A B C TSP Q R" proof - have P1: "\ Col A B C" using TSP_def assms(1) ncop__ncol by blast have P2: "\ Coplanar A B C Q" by (metis TSP_def assms(1) assms(2) assms(3) col_cop2__cop l6_6 out_col out_diff2) have P3: "\ Coplanar A B C R" using TSP_def assms(1) by blast obtain T where P3A: "Coplanar A B C T \ Bet P T R" using TSP_def assms(1) by blast have W1: "D = T \ A B C TSP Q R" using P2 P3 P3A TSP_def assms(3) bet_out__bet by blast { assume V1: "D \ T" have V1A: "\ Col P D T" using P3A col_cop2__cop by (metis TSP_def V1 assms(1) assms(2) col2_cop2__eq col_trivial_2) have V1B: "D T TS P R" by (metis P3 P3A V1A bet__ts invert_two_sides not_col_permutation_3) have "D T OS P Q" using V1A assms(3) not_col_permutation_1 out_one_side by blast then have V2: "D T TS Q R" using V1B l9_8_2 by blast then obtain T' where V3: "Col T' D T \ Bet Q T' R" using TS_def by blast have V4: "Coplanar A B C T'" using Col_cases P3A V1 V3 assms(2) col_cop2__cop by blast then have "A B C TSP Q R" using P2 P3 TSP_def V3 by blast } then have "D \ T \ A B C TSP Q R" by blast then show ?thesis using W1 by blast qed lemma l9_41_1: assumes "A B C TSP P R" and "A B C TSP Q R" shows "A B C OSP P Q" using OSP_def assms(1) assms(2) by blast lemma l9_41_2: assumes "A B C TSP P R" and "A B C OSP P Q" shows "A B C TSP Q R" proof - have P1: "\ Coplanar A B C P" using TSP_def assms(1) by blast obtain S where P2: " A B C TSP P S \ A B C TSP Q S" using OSP_def assms(2) by blast obtain X where P3: "Coplanar A B C X \ Bet P X S" using P2 TSP_def by blast have P4: "\ Coplanar A B C P \ \ Coplanar A B C S" using P2 TSP_def by blast obtain Y where P5: "Coplanar A B C Y \ Bet Q Y S" using P2 TSP_def by blast have P6: "\ Coplanar A B C Q \ \ Coplanar A B C S" using P2 TSP_def by blast have P7: "X \ P \ S \ X \ Q \ Y \ S \ Y" using P3 P4 P5 P6 by blast { assume Q1: "Col P Q S" have Q2: "X = Y" proof - have R2: "Q \ S" using P5 P6 bet_neq12__neq by blast have R5: "Col Q S X" by (smt Col_def P3 Q1 between_inner_transitivity between_symmetry col_transitivity_2) have "Col Q S Y" by (simp add: P5 bet_col col_permutation_5) then show ?thesis using P2 P3 P5 R2 R5 TSP_def col2_cop2__eq by blast qed then have "X Out P Q" by (metis P3 P5 P7 l6_2) then have "A B C TSP Q R" using P3 assms(1) l9_39 by blast } then have P7: "Col P Q S \ A B C TSP Q R" by blast { assume Q1: "\ Col P Q S" obtain Z where Q2: "Bet X Z Q \ Bet Y Z P" using P3 P5 inner_pasch by blast { assume "X = Z" then have "False" by (metis P2 P3 P5 Q1 Q2 TSP_def bet_col col_cop2__cop l6_16_1 not_col_permutation_5) } then have Q3: "X \ Z" by blast have "Y \ Z" proof - have "X \ Z" by (meson \X = Z \ False\) then have "Z \ Y" by (metis (no_types) P2 P3 P5 Q2 TSP_def bet_col col_cop2__cop) then show ?thesis by meson qed then have "Y Out P Z" using Q2 bet_out l6_6 by auto then have Q4: "A B C TSP Z R" using assms(1) P5 l9_39 by blast have "X Out Z Q" using Q2 Q3 bet_out by auto then have "A B C TSP Q R" using Q4 P3 l9_39 by blast } then have "\ Col P Q S \ A B C TSP Q R" by blast then show ?thesis using P7 by blast qed lemma tsp_exists: assumes "\ Coplanar A B C P" shows "\ Q. A B C TSP P Q" proof - obtain Q where P1: "Bet P A Q \ Cong A Q A P" using segment_construction by blast have P2: "Coplanar A B C A" using coplanar_trivial ncoplanar_perm_5 by blast have P3: "\ Coplanar A B C P" by (simp add: assms) have P4: "\ Coplanar A B C Q" by (metis P1 P2 Tarski_neutral_dimensionless.col_cop2__cop Tarski_neutral_dimensionless_axioms assms bet_col cong_diff_4 not_col_permutation_2) then show ?thesis using P1 P2 TSP_def assms by blast qed lemma osp_reflexivity: assumes "\ Coplanar A B C P" shows "A B C OSP P P" by (meson assms l9_41_1 tsp_exists) lemma osp_symmetry: assumes "A B C OSP P Q" shows "A B C OSP Q P" using OSP_def assms by auto lemma osp_transitivity: assumes "A B C OSP P Q" and "A B C OSP Q R" shows "A B C OSP P R" using OSP_def assms(1) assms(2) l9_41_2 by blast lemma cop3_tsp__tsp: assumes "\ Col D E F" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar A B C F" and "A B C TSP P Q" shows "D E F TSP P Q" proof - obtain T where P1: "Coplanar A B C T \ Bet P T Q" using TSP_def assms(5) by blast have P2: "\ Col A B C" using TSP_def assms(5) ncop__ncols by blast have P3: "Coplanar D E F A \ Coplanar D E F B \ Coplanar D E F C \ Coplanar D E F T" proof - have P3A: "Coplanar D E F A" using P2 assms(2) assms(3) assms(4) col_trivial_3 coplanar_pseudo_trans ncop__ncols by blast have P3B: "Coplanar D E F B" using P2 assms(2) assms(3) assms(4) col_trivial_2 coplanar_pseudo_trans ncop__ncols by blast have P3C: "Coplanar D E F C" by (meson P2 assms(2) assms(3) assms(4) coplanar_perm_16 coplanar_pseudo_trans coplanar_trivial) have "Coplanar D E F T" using P1 P2 assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast then show ?thesis using P3A P3B P3C by simp qed have P4: "\ Coplanar D E F P" using P3 TSP_def assms(1) assms(5) coplanar_pseudo_trans by auto have P5: "\ Coplanar D E F Q" by (metis P1 P3 P4 TSP_def assms(5) bet_col bet_col1 col2_cop2__eq) have P6: "Coplanar D E F T" by (simp add: P3) have "Bet P T Q" by (simp add: P1) then show ?thesis using P4 P5 P6 TSP_def by blast qed lemma cop3_osp__osp: assumes "\ Col D E F" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar A B C F" and "A B C OSP P Q" shows "D E F OSP P Q" proof - obtain R where P1: "A B C TSP P R \ A B C TSP Q R" using OSP_def assms(5) by blast then show ?thesis using OSP_def assms(1) assms(2) assms(3) assms(4) cop3_tsp__tsp by blast qed lemma ncop_distincts: assumes "\ Coplanar A B C D" shows "A \ B \ A \ C \ A \ D \ B \ C \ B \ D \ C \ D" using Coplanar_def assms col_trivial_1 col_trivial_2 by blast lemma tsp_distincts: assumes "A B C TSP P Q" shows "A \ B \ A \ C \ B \ C \ A \ P \ B \ P \ C \ P \ A \ Q \ B \ Q \ C \ Q \ P \ Q" proof - obtain pp :: "'p \ 'p \ 'p \ 'p \ 'p \ 'p" where "\x0 x1 x2 x3 x4. (\v5. Coplanar x4 x3 x2 v5 \ Bet x1 v5 x0) = (Coplanar x4 x3 x2 (pp x0 x1 x2 x3 x4) \ Bet x1 (pp x0 x1 x2 x3 x4) x0)" by moura then have f1: "\ Coplanar A B C P \ \ Coplanar A B C Q \ Coplanar A B C (pp Q P C B A) \ Bet P (pp Q P C B A) Q" using TSP_def assms by presburger then have "Q \ pp Q P C B A" by force then show ?thesis using f1 by (meson bet_neq32__neq ncop_distincts) qed lemma osp_distincts: assumes "A B C OSP P Q" shows "A \ B \ A \ C \ B \ C \ A \ P \ B \ P \ C \ P \ A \ Q \ B \ Q \ C \ Q" using OSP_def assms tsp_distincts by blast lemma tsp__ncop1: assumes "A B C TSP P Q" shows "\ Coplanar A B C P" using TSP_def assms by blast lemma tsp__ncop2: assumes "A B C TSP P Q" shows "\ Coplanar A B C Q" using TSP_def assms by auto lemma osp__ncop1: assumes "A B C OSP P Q" shows "\ Coplanar A B C P" using OSP_def TSP_def assms by blast lemma osp__ncop2: assumes "A B C OSP P Q" shows "\ Coplanar A B C Q" using assms osp__ncop1 osp_symmetry by blast lemma tsp__nosp: assumes "A B C TSP P Q" shows "\ A B C OSP P Q" using assms l9_41_2 tsp_distincts by blast lemma osp__ntsp: assumes "A B C OSP P Q" shows "\ A B C TSP P Q" using assms tsp__nosp by blast lemma osp_bet__osp: assumes "A B C OSP P R" and "Bet P Q R" shows "A B C OSP P Q" proof - obtain S where P1: "A B C TSP P S" using OSP_def assms(1) by blast then obtain Y where P2: "Coplanar A B C Y \ Bet R Y S" using TSP_def assms(1) l9_41_2 by blast obtain X where Q1: "Coplanar A B C X \ Bet P X S" using P1 TSP_def by blast have Q2: "P \ X \ S \ X \ R \ Y" using P1 P2 Q1 TSP_def assms(1) osp__ncop2 by auto { assume P3: "Col P R S" have P5: "A B C TSP Q S" proof - have Q3: "X = Y" proof - have R1: "\ Coplanar A B C R" using assms(1) osp__ncop2 by blast have R2: "R \ S" using P1 assms(1) osp__ntsp by blast have R5: "Col R S X" by (smt Col_def P3 Q1 bet_col1 between_exchange4 between_symmetry) have "Col R S Y" using P2 bet_col col_permutation_5 by blast then show ?thesis using R1 R2 Q1 P2 R5 col2_cop2__eq by blast qed then have "Y Out P Q" by (smt P2 P3 Q1 Q2 assms(2) bet_col1 between_exchange4 between_symmetry l6_3_2 l6_4_2 not_bet_and_out third_point) then show ?thesis using P1 P2 l9_39 by blast qed then have "A B C OSP P Q" using OSP_def P1 P2 l9_39 by blast } then have H1: "Col P R S \ A B C OSP P Q" by blast { assume T1: "\ Col P R S" have T2: "X Y OS P R" proof - have T3: "P \ X \ S \ X \ R \ Y \ S \ Y" using P1 P2 Q2 TSP_def by auto have T4: "\ Col S X Y" by (metis P2 Q1 T1 T3 bet_out_1 col_out2_col col_permutation_5 not_col_permutation_4) have T5: "X Y TS P S" by (metis Col_perm Q1 Q2 T4 bet__ts bet_col col_transitivity_2) have T6: "X Y TS R S" by (metis P2 Q1 T4 assms(1) bet__ts col_cop2__cop invert_two_sides not_col_distincts osp__ncop2) then show ?thesis using T5 l9_8_1 by auto qed then have T7: "X Y OS P Q" using assms(2) l9_17 by blast then obtain S' where T7A: "X Y TS P S' \ X Y TS Q S'" using OS_def by blast have T7B: "\ Col P X Y \ \ Col S' X Y \ (\ T::'p. Col T X Y \ Bet P T S')" using T7A TS_def by auto have T7C: "\ Col Q X Y \ \ Col S' X Y \ (\ T::'p. Col T X Y \ Bet Q T S')" using T7A TS_def by blast obtain X' where T9: "Col X' X Y \ Bet P X' S' \ X Y TS Q S'" using T7A T7B by blast obtain Y' where T10: "Col Y' X Y \ Bet Q Y' S'" using T7C by blast have T11: "Coplanar A B C X'" using Col_cases P2 Q1 T9 col_cop2__cop ts_distincts by blast have T12: "Coplanar A B C Y'" using Col_cases P2 Q1 T10 T9 col_cop2__cop ts_distincts by blast have T13: "\ Coplanar A B C S'" using T11 T7C T9 assms(1) bet_col bet_col1 col2_cop2__eq osp__ncop1 by fastforce then have "A B C OSP P Q" proof - have R1: "A B C TSP P S'" using P1 T11 T13 T9 TSP_def by blast have "A B C TSP Q S'" by (metis T10 T12 T13 T7C TSP_def bet_col col_cop2__cop) then show ?thesis using R1 by (smt l9_41_1) qed } then show ?thesis using H1 by blast qed lemma l9_18_3: assumes "Coplanar A B C P" and "Col X Y P" shows "A B C TSP X Y \ (Bet X P Y \ \ Coplanar A B C X \ \ Coplanar A B C Y)" by (meson TSP_def assms(1) assms(2) l9_39 not_bet_out not_col_permutation_5 tsp_distincts) lemma bet_cop__tsp: assumes "\ Coplanar A B C X" and "P \ Y" and "Coplanar A B C P" and "Bet X P Y" shows "A B C TSP X Y" using TSP_def assms(1) assms(2) assms(3) assms(4) bet_col bet_col1 col2_cop2__eq by fastforce lemma cop_out__osp: assumes "\ Coplanar A B C X" and "Coplanar A B C P" and "P Out X Y" shows "A B C OSP X Y" by (meson OSP_def assms(1) assms(2) assms(3) l9_39 tsp_exists) lemma l9_19_3: assumes "Coplanar A B C P" and "Col X Y P" shows "A B C OSP X Y \ (P Out X Y \ \ Coplanar A B C X)" by (meson assms(1) assms(2) cop_out__osp l6_4_2 l9_18_3 not_col_permutation_5 osp__ncop1 osp__ncop2 tsp__nosp) lemma cop2_ts__tsp: assumes "\ Coplanar A B C X" and "Coplanar A B C D" and "Coplanar A B C E" and "D E TS X Y" shows "A B C TSP X Y" proof - obtain T where P1: "Col T D E \ Bet X T Y" using TS_def assms(4) by blast have P2: "Coplanar A B C T" using P1 assms(2) assms(3) assms(4) col_cop2__cop not_col_permutation_2 ts_distincts by blast then show ?thesis by (metis P1 TS_def assms(1) assms(4) bet_cop__tsp) qed lemma cop2_os__osp: assumes "\ Coplanar A B C X" and "Coplanar A B C D" and "Coplanar A B C E" and "D E OS X Y" shows "A B C OSP X Y" proof - obtain Z where P1: "D E TS X Z \ D E TS Y Z" using OS_def assms(4) by blast then have P2: "A B C TSP X Z" using assms(1) assms(2) assms(3) cop2_ts__tsp by blast then have P3: "A B C TSP Y Z" by (meson P1 assms(2) assms(3) cop2_ts__tsp l9_2 tsp__ncop2) then show ?thesis using P2 l9_41_1 by blast qed lemma cop3_tsp__ts: assumes "D \ E" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar D E X Y" and "A B C TSP X Y" shows "D E TS X Y" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) col_cop2__cop cop2_os__osp cop_nts__os not_col_permutation_2 tsp__ncop1 tsp__ncop2 tsp__nosp) lemma cop3_osp__os: assumes "D \ E" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar D E X Y" and "A B C OSP X Y" shows "D E OS X Y" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) col_cop2__cop cop2_ts__tsp cop_nts__os not_col_permutation_2 osp__ncop1 osp__ncop2 tsp__nosp) lemma cop_tsp__ex_cop2: assumes (*"Coplanar A B C P" and*) "A B C TSP D E" shows "\ Q. (Coplanar A B C Q \ Coplanar D E P Q \ P \ Q)" proof cases assume "Col D E P" then show ?thesis by (meson ex_diff_cop ncop__ncols) next assume "\ Col D E P" then obtain Q where "Coplanar A B C Q \ Bet D Q E \ \ Col D E P" using TSP_def assms(1) by blast then show ?thesis using Col_perm bet_col ncop__ncols by blast qed lemma cop_osp__ex_cop2: assumes "Coplanar A B C P" and "A B C OSP D E" shows "\ Q. Coplanar A B C Q \ Coplanar D E P Q \ P \ Q" proof cases assume "Col D E P" then show ?thesis by (metis col_trivial_3 diff_col_ex ncop__ncols) next assume P1: "\ Col D E P" obtain E' where P2: "Bet E P E' \ Cong P E' P E" using segment_construction by blast have P3: "\ Col D E' P" by (metis P1 P2 bet_col bet_cong_eq between_symmetry col_permutation_5 l5_2 l6_16_1) have P4: "A B C TSP D E'" by (metis P2 P3 assms(1) assms(2) bet_cop__tsp l9_41_2 not_col_distincts osp__ncop2 osp_symmetry) then have "\ Coplanar A B C D \ \ Coplanar A B C E' \ (\ T. Coplanar A B C T \ Bet D T E')" by (simp add: TSP_def) then obtain Q where P7: "Coplanar A B C Q \ Bet D Q E'" by blast then have "Coplanar D E' P Q" using bet_col ncop__ncols ncoplanar_perm_5 by blast then have "Coplanar D E P Q" using Col_perm P2 P3 bet_col col_cop__cop ncoplanar_perm_5 not_col_distincts by blast then show ?thesis using P3 P7 bet_col col_permutation_5 by blast qed lemma sac__coplanar: assumes "Saccheri A B C D" shows "Coplanar A B C D" using Saccheri_def assms ncoplanar_perm_4 os__coplanar by blast subsection "Line reflexivity" subsubsection "Dimensionless" lemma Ch10_Goal1: assumes "\ Coplanar D C B A" shows "\ Coplanar A B C D" by (simp add: assms ncoplanar_perm_23) lemma ex_sym: "\ Y. (A B Perp X Y \ X = Y) \ (\ M. Col A B M \ M Midpoint X Y)" proof cases assume "Col A B X" thus ?thesis using l7_3_2 by blast next assume "\ Col A B X" then obtain M0 where P1: "Col A B M0 \ A B Perp X M0" using l8_18_existence by blast obtain Z where P2: "M0 Midpoint X Z" using symmetric_point_construction by blast thus ?thesis by (metis (full_types) P1 Perp_cases bet_col midpoint_bet perp_col) qed lemma is_image_is_image_spec: assumes "A \ B" shows "P' P Reflect A B \ P' P ReflectL A B" by (simp add: Reflect_def assms) lemma ex_sym1: assumes "A \ B" shows "\ Y. (A B Perp X Y \ X = Y) \ (\ M. Col A B M \ M Midpoint X Y \ X Y Reflect A B)" proof cases assume "Col A B X" thus ?thesis by (meson ReflectL_def Reflect_def assms l7_3_2) next assume P0: "\ Col A B X" then obtain M0 where P1: "Col A B M0 \ A B Perp X M0" using l8_18_existence by blast obtain Z where P2: "M0 Midpoint X Z" using symmetric_point_construction by blast have P3: "A B Perp X Z" proof cases assume "X = Z" thus ?thesis using P1 P2 P0 midpoint_distinct by blast next assume "X \ Z" then have P2: "X Z Perp A B" using P1 P2 Perp_cases bet_col midpoint_bet perp_col by blast show ?thesis by (simp add: Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms P2) qed have P10: "(A B Perp X Z \ X = Z)" by (simp add: P3) have "\ M. Col A B M \ M Midpoint X Z \ X Z Reflect A B" using P1 P2 P3 ReflectL_def assms is_image_is_image_spec l7_2 perp_right_comm by blast thus ?thesis using P3 by blast qed lemma l10_2_uniqueness: assumes "P1 P Reflect A B" and "P2 P Reflect A B" shows "P1 = P2" proof cases assume "A = B" thus ?thesis using Reflect_def assms(1) assms(2) symmetric_point_uniqueness by auto next assume P1: "A \ B" have P1A: "P1 P ReflectL A B" using P1 assms(1) is_image_is_image_spec by auto then have P1B: "A B Perp P P1 \ P = P1" using ReflectL_def by blast have P2A: "P2 P ReflectL A B" using P1 assms(2) is_image_is_image_spec by auto then have P2B: "A B Perp P P2 \ P = P2" using ReflectL_def by blast obtain X where R1: "X Midpoint P P1 \ Col A B X" by (metis ReflectL_def assms(1) col_trivial_1 is_image_is_image_spec midpoint_existence) obtain Y where R2: "Y Midpoint P P2 \ Col A B Y" by (metis ReflectL_def assms(2) col_trivial_1 is_image_is_image_spec midpoint_existence) { assume Q1:"(A B Perp P P1 \ A B Perp P P2)" have S1: "P \ X" proof - { assume "P = X" then have "P = P1" using R1 is_midpoint_id by blast then have "A B Perp P P" using Q1 by blast then have "False" using perp_distinct by blast } thus ?thesis by blast qed then have "P1 = P2" by (smt Perp_cases Q1 \\thesis. (\X. X Midpoint P P1 \ Col A B X \ thesis) \ thesis\ \\thesis. (\Y. Y Midpoint P P2 \ Col A B Y \ thesis) \ thesis\ col_permutation_1 l7_2 l7_9 l8_18_uniqueness midpoint_col perp_col perp_not_col2) } then have T1: "(A B Perp P P1 \ A B Perp P P2) \ P1 = P2" by blast have T2: "(P = P1 \ A B Perp P P2) \ P1 = P2" by (metis R1 R2 col3 col_trivial_2 col_trivial_3 midpoint_col midpoint_distinct_1 midpoint_distinct_2 perp_not_col2) have T3: "(P = P2 \ A B Perp P P1) \ P1 = P2" by (metis R1 R2 col_trivial_2 midpoint_col midpoint_distinct_3 perp_col2 perp_not_col2) thus ?thesis using T1 T2 T3 P1B P2B by blast qed lemma l10_2_uniqueness_spec: assumes "P1 P ReflectL A B" and "P2 P ReflectL A B" shows "P1 = P2" proof - have "A B Perp P P1 \ P = P1" using ReflectL_def assms(1) by blast moreover obtain X1 where "X1 Midpoint P P1 \ Col A B X1" using ReflectL_def assms(1) by blast moreover have "A B Perp P P2 \ P = P2" using ReflectL_def assms(2) by blast moreover obtain X2 where "X2 Midpoint P P2 \ Col A B X2" using ReflectL_def assms(2) by blast ultimately show ?thesis by (smt col_permutation_1 l8_16_1 l8_18_uniqueness midpoint_col midpoint_distinct_3 perp_col1 symmetric_point_uniqueness) qed lemma l10_2_existence_spec: "\ P'. P' P ReflectL A B" proof cases assume "Col A B P" thus ?thesis using ReflectL_def l7_3_2 by blast next assume "\ Col A B P" then obtain X where "Col A B X \ A B Perp P X" using l8_18_existence by blast moreover obtain P' where "X Midpoint P P'" using symmetric_point_construction by blast ultimately show ?thesis using ReflectL_def bet_col midpoint_bet perp_col1 by blast qed lemma l10_2_existence: "\ P'. P' P Reflect A B" by (metis Reflect_def l10_2_existence_spec symmetric_point_construction) lemma l10_4_spec: assumes "P P' ReflectL A B" shows "P' P ReflectL A B" proof - obtain X where "X Midpoint P P' \ Col A B X" using ReflectL_def assms l7_2 by blast thus ?thesis using Perp_cases ReflectL_def assms by auto qed lemma l10_4: assumes "P P' Reflect A B" shows "P' P Reflect A B" using Reflect_def Tarski_neutral_dimensionless.l7_2 Tarski_neutral_dimensionless_axioms assms l10_4_spec by fastforce lemma l10_5: assumes "P' P Reflect A B" and "P'' P' Reflect A B" shows "P = P''" by (meson assms(1) assms(2) l10_2_uniqueness l10_4) lemma l10_6_uniqueness: assumes "P P1 Reflect A B" and "P P2 Reflect A B" shows "P1 = P2" using assms(1) assms(2) l10_4 l10_5 by blast lemma l10_6_uniqueness_spec: assumes "P P1 ReflectL A B" and "P P2 ReflectL A B" shows "P1 = P2" using assms(1) assms(2) l10_2_uniqueness_spec l10_4_spec by blast lemma l10_6_existence_spec: assumes "A \ B" shows "\ P. P' P ReflectL A B" using l10_2_existence_spec l10_4_spec by blast lemma l10_6_existence: "\ P. P' P Reflect A B" using l10_2_existence l10_4 by blast lemma l10_7: assumes "P' P Reflect A B" and "Q' Q Reflect A B" and "P' = Q'" shows "P = Q" using assms(1) assms(2) assms(3) l10_6_uniqueness by blast lemma l10_8: assumes "P P Reflect A B" shows "Col P A B" by (metis Col_perm assms col_trivial_2 ex_sym1 l10_6_uniqueness l7_3) lemma col__refl: assumes "Col P A B" shows "P P ReflectL A B" using ReflectL_def assms col_permutation_1 l7_3_2 by blast lemma is_image_col_cong: assumes "A \ B" and "P P' Reflect A B" and "Col A B X" shows "Cong P X P' X" proof - have P1: "P P' ReflectL A B" using assms(1) assms(2) is_image_is_image_spec by blast obtain M0 where P2: "M0 Midpoint P' P \ Col A B M0" using P1 ReflectL_def by blast have "A B Perp P' P \ P' = P" using P1 ReflectL_def by auto moreover { assume S1: "A B Perp P' P" then have "A \ B \ P' \ P" using perp_distinct by blast have S2: "M0 = X \ Cong P X P' X" using P2 cong_4312 midpoint_cong by blast { assume "M0 \ X" then have "M0 X Perp P' P" using P2 S1 assms(3) perp_col2 by blast then have "\ Col A B P \ Per P M0 X" by (metis Col_perm P2 S1 colx l8_2 midpoint_col midpoint_distinct_1 per_col perp_col1 perp_not_col2 perp_per_1) then have "Cong P X P' X" using P2 cong_commutativity l7_2 l8_2 per_double_cong by blast } then have "Cong P X P' X" using S2 by blast } then have "A B Perp P' P \ Cong P X P' X" by blast moreover { assume "P = P'" then have "Cong P X P' X" by (simp add: cong_reflexivity) } ultimately show ?thesis by blast qed lemma is_image_spec_col_cong: assumes "P P' ReflectL A B" and "Col A B X" shows "Cong P X P' X" by (metis Col_def Reflect_def assms(1) assms(2) between_trivial col__refl cong_reflexivity is_image_col_cong l10_6_uniqueness_spec) lemma image_id: assumes "A \ B" and "Col A B T" and "T T' Reflect A B" shows "T = T'" using assms(1) assms(2) assms(3) cong_diff_4 is_image_col_cong by blast lemma osym_not_col: assumes "P P' Reflect A B" and "\ Col A B P" shows "\ Col A B P'" using assms(1) assms(2) l10_4 local.image_id not_col_distincts by blast lemma midpoint_preserves_image: assumes "A \ B" and "Col A B M" and "P P' Reflect A B" and "M Midpoint P Q" and "M Midpoint P' Q'" shows "Q Q' Reflect A B" proof - obtain X where P1: "X Midpoint P' P \ Col A B X" using ReflectL_def assms(1) assms(3) is_image_is_image_spec by blast { assume S1: "A B Perp P' P" obtain Y where S2: "M Midpoint X Y" using symmetric_point_construction by blast have S3: "Y Midpoint Q Q'" proof - have R4: "X Midpoint P P'" by (simp add: P1 l7_2) thus ?thesis using assms(4) assms(5) S2 symmetry_preserves_midpoint by blast qed have S4: "P \ P'" using S1 perp_not_eq_2 by blast then have S5: "Q \ Q'" using Tarski_neutral_dimensionless.l7_9 Tarski_neutral_dimensionless_axioms assms(4) assms(5) by fastforce have S6: "Y Midpoint Q' Q \ Col A B Y" by (metis P1 S2 S3 assms(2) colx l7_2 midpoint_col midpoint_distinct_1) have S7: "A B Perp Q' Q \ Q = Q'" proof - have R3: "Per M Y Q" proof - have S1: "Y Midpoint Q Q'" using S3 by auto have "Cong M Q M Q'" using assms(1) assms(2) assms(3) assms(4) assms(5) cong_commutativity is_image_col_cong l7_16 l7_3_2 by blast thus ?thesis using Per_def S1 by blast qed { have "X = Y \ (A B Perp Q' Q \ Q = Q')" by (metis P1 Perp_cases S1 S2 S6 assms(5) l7_3 l7_9_bis) { assume "X \ Y" then have "Y PerpAt M Y Y Q" using R3 S2 S3 S5 midpoint_distinct_1 per_perp_in by blast then have V1: "Y Y Perp Y Q \ M Y Perp Y Q" by (simp add: perp_in_perp_bis) { have "Y Y Perp Y Q \ A B Perp Q' Q \ Q = Q'" using perp_not_eq_1 by blast { assume T1: "M Y Perp Y Q" have T2: "Y Q Perp A B" proof cases assume "A = M" thus ?thesis using Perp_cases S6 T1 assms(1) col_permutation_5 perp_col by blast next assume "A \ M" thus ?thesis by (smt S6 T1 assms(1) assms(2) col2__eq col_transitivity_2 perp_col0 perp_not_eq_1) qed have "A B Perp Q' Q \ Q = Q'" by (metis S3 T2 midpoint_col not_col_distincts perp_col0) } then have "M Y Perp Y Q \ A B Perp Q' Q \ Q = Q'" by blast } then have "A B Perp Q' Q \ Q = Q'" using V1 perp_distinct by blast } then have "X \ Y \ (A B Perp Q' Q \ Q = Q')" by blast } thus ?thesis by (metis P1 Perp_cases S1 S2 S6 assms(5) l7_3 l7_9_bis) qed then have "Q Q' ReflectL A B" using ReflectL_def S6 by blast } then have "A B Perp P' P \ Q Q' ReflectL A B" by blast moreover { assume "P = P'" then have "Q Q' ReflectL A B" by (metis P1 assms(2) assms(4) assms(5) col__refl col_permutation_2 colx midpoint_col midpoint_distinct_3 symmetric_point_uniqueness) } ultimately show ?thesis using ReflectL_def assms(1) assms(3) is_image_is_image_spec by auto qed lemma image_in_is_image_spec: assumes "M ReflectLAt P P' A B" shows "P P' ReflectL A B" proof - have P1: "M Midpoint P' P" using ReflectLAt_def assms by blast have P2: "Col A B M" using ReflectLAt_def assms by blast have "A B Perp P' P \ P' = P" using ReflectLAt_def assms by blast thus ?thesis using P1 P2 using ReflectL_def by blast qed lemma image_in_gen_is_image: assumes "M ReflectAt P P' A B" shows "P P' Reflect A B" using ReflectAt_def Reflect_def assms image_in_is_image_spec by auto lemma image_image_in: assumes "P \ P'" and "P P' ReflectL A B" and "Col A B M" and "Col P M P'" shows "M ReflectLAt P P' A B" proof - obtain M' where P1: "M' Midpoint P' P \ Col A B M'" using ReflectL_def assms(2) by blast have Q1: "P M' Perp A B" by (metis Col_cases P1 Perp_perm ReflectL_def assms(1) assms(2) bet_col cong_diff_3 midpoint_bet midpoint_cong not_cong_4321 perp_col1) { assume R1: "A B Perp P' P" have R3: "P \ M'" using Q1 perp_not_eq_1 by auto have R4: "A B Perp P' P" by (simp add: R1) have R5: "Col P P' M'" using P1 midpoint_col not_col_permutation_3 by blast have R6: "M' Midpoint P' P" by (simp add: P1) have R7: "\ Col A B P" using assms(1) assms(2) col__refl col_permutation_2 l10_2_uniqueness_spec l10_4_spec by blast have R8: "P \ P'" by (simp add: assms(1)) have R9: "Col A B M'" by (simp add: P1) have R10: "Col A B M" by (simp add: assms(3)) have R11: "Col P P' M'" by (simp add: R5) have R12: "Col P P' M" using Col_perm assms(4) by blast have "M = M'" proof cases assume S1: "A = M'" have "Per P M' A" by (simp add: S1 l8_5) thus ?thesis using l6_21 R8 R9 R10 R11 R12 using R7 by blast next assume "A \ M'" thus ?thesis using R10 R12 R5 R7 R8 R9 l6_21 by blast qed then have "M Midpoint P' P" using R6 by blast } then have Q2: "A B Perp P' P \ M Midpoint P' P" by blast have Q3: "P' = P \ M Midpoint P' P" using assms(1) by auto have Q4: "A B Perp P' P \ P' = P" using ReflectL_def assms(2) by auto then have "M Midpoint P' P" using Q2 Q3 by blast thus ?thesis by (simp add: ReflectLAt_def Q4 assms(3)) qed lemma image_in_col: assumes "Y ReflectLAt P P' A B" shows "Col P P' Y" using Col_perm ReflectLAt_def assms midpoint_col by blast lemma is_image_spec_rev: assumes "P P' ReflectL A B" shows "P P' ReflectL B A" proof - obtain M0 where P1: "M0 Midpoint P' P \ Col A B M0" using ReflectL_def assms by blast have P2: "Col B A M0" using Col_cases P1 by blast have "A B Perp P' P \ P' = P" using ReflectL_def assms by blast thus ?thesis using P1 P2 Perp_cases ReflectL_def by auto qed lemma is_image_rev: assumes "P P' Reflect A B" shows "P P' Reflect B A" using Reflect_def assms is_image_spec_rev by auto lemma midpoint_preserves_per: assumes "Per A B C" and "M Midpoint A A1" and "M Midpoint B B1" and "M Midpoint C C1" shows "Per A1 B1 C1" proof - obtain C' where P1: "B Midpoint C C' \ Cong A C A C'" using Per_def assms(1) by blast obtain C1' where P2: "M Midpoint C' C1'" using symmetric_point_construction by blast thus ?thesis by (meson P1 Per_def assms(2) assms(3) assms(4) l7_16 symmetry_preserves_midpoint) qed lemma col__image_spec: assumes "Col A B X" shows "X X ReflectL A B" by (simp add: assms col__refl col_permutation_2) lemma image_triv: "A A Reflect A B" by (simp add: Reflect_def col__refl col_trivial_1 l7_3_2) lemma cong_midpoint__image: assumes "Cong A X A Y" and "B Midpoint X Y" shows "Y X Reflect A B" proof cases assume "A = B" thus ?thesis by (simp add: Reflect_def assms(2)) next assume S0: "A \ B" { assume S1: "X \ Y" then have "X Y Perp A B" proof - have T1: "B \ X" using S1 assms(2) midpoint_distinct_1 by blast have T2: "B \ Y" using S1 assms(2) midpoint_not_midpoint by blast have "Per A B X" using Per_def assms(1) assms(2) by auto thus ?thesis using S0 S1 T1 T2 assms(2) col_per_perp midpoint_col by auto qed then have "A B Perp X Y \ X = Y" using Perp_perm by blast then have "Y X Reflect A B" using ReflectL_def S0 assms(2) col_trivial_2 is_image_is_image_spec by blast } then have "X \ Y \ Y X Reflect A B" by blast thus ?thesis using assms(2) image_triv is_image_rev l7_3 by blast qed lemma col_image_spec__eq: assumes "Col A B P" and "P P' ReflectL A B" shows "P = P'" using assms(1) assms(2) col__image_spec l10_2_uniqueness_spec l10_4_spec by blast lemma image_spec_triv: "A A ReflectL B B" using col__image_spec not_col_distincts by blast lemma image_spec__eq: assumes "P P' ReflectL A A" shows "P = P'" using assms col_image_spec__eq not_col_distincts by blast lemma image__midpoint: assumes "P P' Reflect A A" shows "A Midpoint P' P" using Reflect_def assms by auto lemma is_image_spec_dec: "A B ReflectL C D \ \ A B ReflectL C D" by simp lemma l10_14: assumes "P \ P'" and "A \ B" and "P P' Reflect A B" shows "A B TS P P'" proof - have P1: "P P' ReflectL A B" using assms(2) assms(3) is_image_is_image_spec by blast then obtain M0 where "M0 Midpoint P' P \ Col A B M0" using ReflectL_def by blast then have "A B Perp P' P \ A B TS P P'" by (meson TS_def assms(1) assms(2) assms(3) between_symmetry col_permutation_2 local.image_id midpoint_bet osym_not_col) thus ?thesis using assms(1) P1 ReflectL_def by blast qed lemma l10_15: assumes "Col A B C" and "\ Col A B P" shows "\ Q. A B Perp Q C \ A B OS P Q" proof - have P1: "A \ B" using assms(2) col_trivial_1 by auto obtain X where P2: "A B TS P X" using assms(2) col_permutation_1 l9_10 by blast { assume Q1: "A = C" obtain Q where Q2: "\ T. A B Perp Q A \ Col A B T \ Bet X T Q" using P1 l8_21 by blast then obtain T where "A B Perp Q A \ Col A B T \ Bet X T Q" by blast then have "A B TS Q X" by (meson P2 TS_def between_symmetry col_permutation_2 perp_not_col) then have Q5: "A B OS P Q" using P2 l9_8_1 by blast then have "\ Q. A B Perp Q C \ A B OS P Q" using Q1 Q2 by blast } then have P3: "A = C \ (\ Q. A B Perp Q C \ A B OS P Q)" by blast { assume Q1: "A \ C" then obtain Q where Q2: "\ T. C A Perp Q C \ Col C A T \ Bet X T Q" using l8_21 by presburger then obtain T where Q3: "C A Perp Q C \ Col C A T \ Bet X T Q" by blast have Q4: "A B Perp Q C" using NCol_perm P1 Q2 assms(1) col_trivial_2 perp_col2 by blast have "A B TS Q X" proof - have R1: "\ Col Q A B" using Col_perm P1 Q2 assms(1) col_trivial_2 colx perp_not_col by blast have R2: "\ Col X A B" using P2 TS_def by auto have R3: "Col T A B" by (metis Q1 Q3 assms(1) col_trivial_2 colx not_col_permutation_1) have "Bet Q T X" using Bet_cases Q3 by blast then have "\ T. Col T A B \ Bet Q T X" using R3 by blast thus ?thesis using R1 R2 by (simp add: TS_def) qed then have "A B OS P Q" using P2 l9_8_1 by blast then have "\ Q. A B Perp Q C \ A B OS P Q" using Q4 by blast } thus ?thesis using P3 by blast qed lemma ex_per_cong: assumes "A \ B" and "X \ Y" and "Col A B C" and "\ Col A B D" shows "\ P. Per P C A \ Cong P C X Y \ A B OS P D" proof - obtain Q where P1: "A B Perp Q C \ A B OS D Q" using assms(3) assms(4) l10_15 by blast obtain P where P2: "C Out Q P \ Cong C P X Y" by (metis P1 assms(2) perp_not_eq_2 segment_construction_3) have P3: "Per P C A" using P1 P2 assms(3) col_trivial_3 l8_16_1 l8_3 out_col by blast have "A B OS P D" using P1 P2 assms(3) one_side_symmetry os_out_os by blast thus ?thesis using P2 P3 cong_left_commutativity by blast qed lemma exists_cong_per: "\ C. Per A B C \ Cong B C X Y" proof cases assume "A = B" thus ?thesis by (meson Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms l8_2 segment_construction) next assume "A \ B" thus ?thesis by (metis Perp_perm bet_col between_trivial l8_16_1 l8_21 segment_construction) qed subsubsection "Upper dim 2" lemma upper_dim_implies_per2__col: assumes "upper_dim_axiom" shows "\ A B C X. (Per A X C \ X \ C \ Per B X C) \ Col A B X" proof - { fix A B C X assume "Per A X C \ X \ C \ Per B X C" moreover then obtain C' where "X Midpoint C C' \ Cong A C A C'" using Per_def by blast ultimately have "Col A B X" by (smt Col_def assms midpoint_cong midpoint_distinct_2 not_cong_2134 per_double_cong upper_dim_axiom_def) } then show ?thesis by blast qed lemma upper_dim_implies_col_perp2__col: assumes "upper_dim_axiom" shows "\ A B X Y P. (Col A B P \ A B Perp X P \ P A Perp Y P) \ Col Y X P" proof - { fix A B X Y P assume H1: "Col A B P \ A B Perp X P \ P A Perp Y P" then have H2: "P \ A" using perp_not_eq_1 by blast have "Col Y X P" proof - have T1: "Per Y P A" using H1 l8_2 perp_per_1 by blast moreover have "Per X P A" using H1 col_trivial_3 l8_16_1 by blast then show ?thesis using T1 H2 using assms upper_dim_implies_per2__col by blast qed } then show ?thesis by blast qed lemma upper_dim_implies_perp2__col: assumes "upper_dim_axiom" shows "\ X Y Z A B. (X Y Perp A B \ X Z Perp A B) \ Col X Y Z" proof - { fix X Y Z A B assume H1: "X Y Perp A B \ X Z Perp A B" then have H1A: "X Y Perp A B" by blast have H1B: "X Z Perp A B" using H1 by blast obtain C where H2: "C PerpAt X Y A B" using H1 Perp_def by blast obtain C' where H3: "C' PerpAt X Z A B" using H1 Perp_def by blast have "Col X Y Z" proof cases assume H2: "Col A B X" { assume "X = A" then have "Col X Y Z" using upper_dim_implies_col_perp2__col by (metis H1 H2 Perp_cases assms col_permutation_1) } then have P1: "X = A \ Col X Y Z" by blast { assume P2: "X \ A" then have P3: "A B Perp X Y" using perp_sym using H1 Perp_perm by blast have "Col A B X" by (simp add: H2) then have P4: "A X Perp X Y" using perp_col using P2 P3 by auto have P5: "A X Perp X Z" by (metis H1 H2 P2 Perp_perm col_trivial_3 perp_col0) have P6: "Col Y Z X" proof - have Q1: "upper_dim_axiom" by (simp add: assms) have Q2: "Per Y X A" using P4 Perp_cases perp_per_2 by blast have "Per Z X A" by (meson P5 Tarski_neutral_dimensionless.Perp_cases Tarski_neutral_dimensionless_axioms perp_per_2) then show ?thesis using Q1 Q2 P2 using upper_dim_implies_per2__col by blast qed then have "Col X Y Z" using Col_perm by blast } then show ?thesis using P1 by blast next assume T1: "\ Col A B X" obtain Y0 where Q3: "Y0 PerpAt X Y A B" using H1 Perp_def by blast obtain Z0 where Q4: "Z0 PerpAt X Z A B" using Perp_def H1 by blast have Q5: "X Y0 Perp A B" proof - have R1: "X \ Y0" using Q3 T1 perp_in_col by blast have R2: "X Y Perp A B" by (simp add: H1A) then show ?thesis using R1 using Q3 perp_col perp_in_col by blast qed have "X Z0 Perp A B" by (metis H1B Q4 T1 perp_col perp_in_col) then have Q7: "Y0 = Z0" by (meson Q3 Q4 Q5 T1 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms l8_18_uniqueness perp_in_col) have "Col X Y Z" proof - have "X \ Y0" using Q5 perp_distinct by auto moreover have "Col X Y0 Y" using Q3 not_col_permutation_5 perp_in_col by blast moreover have "Col X Y0 Z" using Q4 Q7 col_permutation_5 perp_in_col by blast ultimately show ?thesis using col_transitivity_1 by blast qed then show ?thesis using l8_18_uniqueness by (smt H1 H2 Perp_cases T1 col_trivial_3 perp_col1 perp_in_col perp_not_col) qed } then show ?thesis by blast qed lemma upper_dim_implies_not_two_sides_one_side_aux: assumes "upper_dim_axiom" shows "\ A B X Y PX. (A \ B \ PX \ A \ A B Perp X PX \ Col A B PX \ \ Col X A B \ \ Col Y A B \ \ A B TS X Y) \ A B OS X Y" proof - { fix A B X Y PX assume H1: "A \ B \ PX \ A \ A B Perp X PX \ Col A B PX \ \ Col X A B \ \ Col Y A B \ \ A B TS X Y" have H1A: "A \ B" using H1 by simp have H1B: "PX \ A" using H1 by simp have H1C: "A B Perp X PX" using H1 by simp have H1D: "Col A B PX" using H1 by simp have H1E: "\ Col X A B" using H1 by simp have H1F: "\ Col Y A B" using H1 by simp have H1G: "\ A B TS X Y" using H1 by simp have "\ P T. PX A Perp P PX \ Col PX A T \ Bet Y T P" using H1B l8_21 by blast then obtain P T where T1: "PX A Perp P PX \ Col PX A T \ Bet Y T P" by blast have J1: "PX A Perp P PX" using T1 by blast have J2: "Col PX A T" using T1 by blast have J3: "Bet Y T P" using T1 by blast have P9: "Col P X PX" using upper_dim_implies_col_perp2__col using H1C H1D J1 assms by blast have J4: "\ Col P A B" using H1A H1D T1 col_trivial_2 colx not_col_permutation_3 perp_not_col by blast have J5: "PX A TS P Y" proof - have f1: "Col PX A B" using H1D not_col_permutation_1 by blast then have f2: "Col B PX A" using not_col_permutation_1 by blast have f3: "\p. (T = A \ Col p A PX) \ \ Col p A T" by (metis J2 l6_16_1) have f4: "Col T PX A" using J2 not_col_permutation_1 by blast have f5: "\p. Col p PX B \ \ Col p PX A" using f2 by (meson H1B l6_16_1) have f6: "\p. (B = PX \ Col p B A) \ \ Col p B PX" using H1D l6_16_1 by blast have f7: "\p pa. ((B = PX \ Col p PX pa) \ \ Col p PX B) \ \ Col pa PX A" using f5 by (metis l6_16_1) have f8: "\p. ((T = A \ B = PX) \ Col p A B) \ \ Col p A PX" using f2 by (metis H1B l6_16_1 not_col_permutation_1) have "Col B T PX" using f5 f4 not_col_permutation_1 by blast then have f9: "\p. (T = PX \ Col p T B) \ \ Col p T PX" using l6_16_1 by blast have "B = PX \ \ Col Y PX A \ \ Col P PX A" using f1 by (metis (no_types) H1B H1F J4 l6_16_1 not_col_permutation_1) then show ?thesis using f9 f8 f7 f6 f5 f4 f3 by (metis (no_types) H1B H1F J3 J4 TS_def l9_2 not_col_permutation_1) qed have J6: "X \ PX" using H1 perp_not_eq_2 by blast have J7: "P \ X" using H1A H1D H1G J5 col_preserves_two_sides col_trivial_2 not_col_permutation_1 by blast have J8: "Bet X PX P \ PX Out X P \ \ Col X PX P" using l6_4_2 by blast have J9: "PX A TS P X" by (metis H1A H1D H1G J5 J6 J8 Out_cases P9 TS_def bet__ts between_symmetry col_permutation_1 col_preserves_two_sides col_trivial_2 l9_5) then have "A B OS X Y" by (meson H1A H1D J5 col2_os__os col_trivial_2 l9_2 l9_8_1 not_col_permutation_1) } then show ?thesis by blast qed lemma upper_dim_implies_not_two_sides_one_side: assumes "upper_dim_axiom" shows "\ A B X Y. (\ Col X A B \ \ Col Y A B \ \ A B TS X Y) \ A B OS X Y" proof - { fix A B X Y assume H1: "\ Col X A B \ \ Col Y A B \ \ A B TS X Y" have H1A: "\ Col X A B" using H1 by simp have H1B: "\ Col Y A B" using H1 by simp have H1C: "\ A B TS X Y" using H1 by simp have P1: "A \ B" using H1A col_trivial_2 by blast obtain PX where P2: "Col A B PX \ A B Perp X PX" using Col_cases H1 l8_18_existence by blast have "A B OS X Y" proof cases assume H5: "PX = A" have "B A OS X Y" proof - have F1: "B A Perp X A" using P2 Perp_perm H5 by blast have F2: "Col B A A" using not_col_distincts by blast have F3: "\ Col X B A" using Col_cases H1A by blast have F4: "\ Col Y B A" using Col_cases H1B by blast have "\ B A TS X Y" using H1C invert_two_sides by blast then show ?thesis by (metis F1 F3 F4 assms col_trivial_2 upper_dim_implies_not_two_sides_one_side_aux) qed then show ?thesis by (simp add: invert_one_side) next assume "PX \ A" then show ?thesis using H1 P1 P2 assms upper_dim_implies_not_two_sides_one_side_aux by blast qed } then show ?thesis by blast qed lemma upper_dim_implies_not_one_side_two_sides: assumes "upper_dim_axiom" shows "\ A B X Y. (\ Col X A B \ \ Col Y A B \ \ A B OS X Y) \ A B TS X Y" using assms upper_dim_implies_not_two_sides_one_side by blast lemma upper_dim_implies_one_or_two_sides: assumes "upper_dim_axiom" shows "\ A B X Y. (\ Col X A B \ \ Col Y A B) \ (A B TS X Y \ A B OS X Y)" using assms upper_dim_implies_not_two_sides_one_side by blast lemma upper_dim_implies_all_coplanar: assumes "upper_dim_axiom" shows "all_coplanar_axiom" using all_coplanar_axiom_def assms upper_dim_axiom_def by auto lemma all_coplanar_implies_upper_dim: assumes "all_coplanar_axiom" shows "upper_dim_axiom" using all_coplanar_axiom_def assms upper_dim_axiom_def by auto lemma all_coplanar_upper_dim: shows "all_coplanar_axiom \ upper_dim_axiom" using all_coplanar_implies_upper_dim upper_dim_implies_all_coplanar by auto lemma upper_dim_stab: shows "\ \ upper_dim_axiom \ upper_dim_axiom" by blast lemma cop__cong_on_bissect: assumes "Coplanar A B X P" and "M Midpoint A B" and "M PerpAt A B P M" and "Cong X A X B" shows "Col M P X" proof - have P1: "X = M \ \ Col A B X \ M PerpAt X M A B" using assms(2) assms(3) assms(4) cong_commutativity cong_perp_or_mid perp_in_distinct by blast { assume H1: "\ Col A B X \ M PerpAt X M A B" then have Q1: "X M Perp A B" using perp_in_perp by blast have Q2: "A B Perp P M" using assms(3) perp_in_perp by blast have P2: "Col M A B" by (simp add: assms(2) midpoint_col) then have "Col M P X" using cop_perp2__col by (meson Perp_perm Q1 Q2 assms(1) coplanar_perm_1) } then show ?thesis using P1 not_col_distincts by blast qed lemma cong_cop_mid_perp__col: assumes "Coplanar A B X P" and "Cong A X B X" and "M Midpoint A B" and "A B Perp P M" shows "Col M P X" proof - have "M PerpAt A B P M" using Col_perm assms(3) assms(4) bet_col l8_15_1 midpoint_bet by blast then show ?thesis using assms(1) assms(2) assms(3) cop__cong_on_bissect not_cong_2143 by blast qed lemma cop_image_in2__col: assumes "Coplanar A B P Q" and "M ReflectLAt P P' A B" and "M ReflectLAt Q Q' A B" shows "Col M P Q" proof - have P1: "P P' ReflectL A B" using assms(2) image_in_is_image_spec by auto then have P2: "A B Perp P' P \ P' = P" using ReflectL_def by auto have P3: "Q Q' ReflectL A B" using assms(3) image_in_is_image_spec by blast then have P4: "A B Perp Q' Q \ Q' = Q" using ReflectL_def by auto { assume S1: "A B Perp P' P \ A B Perp Q' Q" { assume T1: "A = M" have T2: "Per B A P" by (metis P1 Perp_perm S1 T1 assms(2) image_in_col is_image_is_image_spec l10_14 perp_col1 perp_distinct perp_per_1 ts_distincts) have T3: "Per B A Q" by (metis S1 T1 assms(3) image_in_col l8_5 perp_col1 perp_per_1 perp_right_comm) have T4: "Coplanar B P Q A" using assms(1) ncoplanar_perm_18 by blast have T5: "B \ A" using S1 perp_distinct by blast have T6: "Per P A B" by (simp add: T2 l8_2) have T7: "Per Q A B" using Per_cases T3 by blast then have "Col P Q A" using T4 T5 T6 using cop_per2__col by blast then have "Col A P Q" using not_col_permutation_1 by blast then have "Col M P Q" using T1 by blast } then have S2: "A = M \ Col M P Q" by blast { assume D0: "A \ M" have D1: "Per A M P" proof - have E1: "M Midpoint P P'" using ReflectLAt_def assms(2) l7_2 by blast have "Cong P A P' A" using P1 col_trivial_3 is_image_spec_col_cong by blast then have "Cong A P A P'" using Cong_perm by blast then show ?thesis using E1 Per_def by blast qed have D2: "Per A M Q" proof - have E2: "M Midpoint Q Q'" using ReflectLAt_def assms(3) l7_2 by blast have "Cong A Q A Q'" using P3 col_trivial_3 cong_commutativity is_image_spec_col_cong by blast then show ?thesis using E2 Per_def by blast qed have "Col P Q M" proof - have W1: "Coplanar P Q A B" using assms(1) ncoplanar_perm_16 by blast have W2: "A \ B" using S1 perp_distinct by blast have "Col A B M" using ReflectLAt_def assms(2) by blast then have "Coplanar P Q A M" using W1 W2 col2_cop__cop col_trivial_3 by blast then have V1: "Coplanar A P Q M" using ncoplanar_perm_8 by blast have V3: "Per P M A" by (simp add: D1 l8_2) have "Per Q M A" using D2 Per_perm by blast then show ?thesis using V1 D0 V3 cop_per2__col by blast qed then have "Col M P Q" using Col_perm by blast } then have "A \ M \ Col M P Q" by blast then have "Col M P Q" using S2 by blast } then have P5: "(A B Perp P' P \ A B Perp Q' Q) \ Col M P Q" by blast have P6: "(A B Perp P' P \ (Q' = Q)) \ Col M P Q" using ReflectLAt_def assms(3) l7_3 not_col_distincts by blast have P7: "(P' = P \ A B Perp Q' Q) \ Col M P Q" using ReflectLAt_def assms(2) l7_3 not_col_distincts by blast have "(P' = P \ Q' = Q) \ Col M P Q" using ReflectLAt_def assms(3) col_trivial_3 l7_3 by blast then show ?thesis using P2 P4 P5 P6 P7 by blast qed lemma l10_10_spec: assumes "P' P ReflectL A B" and "Q' Q ReflectL A B" shows "Cong P Q P' Q'" proof cases assume "A = B" then show ?thesis using assms(1) assms(2) cong_reflexivity image_spec__eq by blast next assume H1: "A \ B" obtain X where P1: "X Midpoint P P' \ Col A B X" using ReflectL_def assms(1) by blast obtain Y where P2: "Y Midpoint Q Q' \ Col A B Y" using ReflectL_def assms(2) by blast obtain Z where P3: "Z Midpoint X Y" using midpoint_existence by blast have P4: "Col A B Z" proof cases assume "X = Y" then show ?thesis by (metis P2 P3 midpoint_distinct_3) next assume "X \ Y" then show ?thesis by (metis P1 P2 P3 l6_21 midpoint_col not_col_distincts) qed obtain R where P5: "Z Midpoint P R" using symmetric_point_construction by blast obtain R' where P6: "Z Midpoint P' R'" using symmetric_point_construction by blast have P7: "A B Perp P P' \ P = P'" using ReflectL_def assms(1) by auto have P8: "A B Perp Q Q' \ Q = Q'" using ReflectL_def assms(2) by blast { assume Q1: "A B Perp P P' \ A B Perp Q Q'" have Q2: "R R' ReflectL A B" proof - have "P P' Reflect A B" by (simp add: H1 assms(1) is_image_is_image_spec l10_4_spec) then have "R R' Reflect A B" using H1 P4 P5 P6 midpoint_preserves_image by blast then show ?thesis using H1 is_image_is_image_spec by blast qed have Q3: "R \ R'" using P5 P6 Q1 l7_9 perp_not_eq_2 by blast have Q4: "Y Midpoint R R'" using P1 P3 P5 P6 symmetry_preserves_midpoint by blast have Q5: "Cong Q' R' Q R" using P2 Q4 l7_13 by blast have Q6: "Cong P' Z P Z" using P4 assms(1) is_image_spec_col_cong by auto have Q7: "Cong Q' Z Q Z" using P4 assms(2) is_image_spec_col_cong by blast then have "Cong P Q P' Q'" proof - have S1: "Cong R Z R' Z" using P5 P6 Q6 cong_symmetry l7_16 l7_3_2 by blast have "R \ Z" using Q3 S1 cong_reverse_identity by blast then show ?thesis by (meson P5 P6 Q5 Q6 Q7 S1 between_symmetry five_segment midpoint_bet not_cong_2143 not_cong_3412) qed } then have P9: "(A B Perp P P' \ A B Perp Q Q') \ Cong P Q P' Q'" by blast have P10: "(A B Perp P P' \ Q = Q') \ Cong P Q P' Q'" using P2 Tarski_neutral_dimensionless.l7_3 Tarski_neutral_dimensionless_axioms assms(1) cong_symmetry is_image_spec_col_cong by fastforce have P11: "(P = P' \ A B Perp Q Q') \ Cong P Q P' Q'" using P1 Tarski_neutral_dimensionless.l7_3 Tarski_neutral_dimensionless.not_cong_4321 Tarski_neutral_dimensionless_axioms assms(2) is_image_spec_col_cong by fastforce have "(P = P' \ Q = Q') \ Cong P Q P' Q'" using cong_reflexivity by blast then show ?thesis using P10 P11 P7 P8 P9 by blast qed lemma l10_10: assumes "P' P Reflect A B" and "Q' Q Reflect A B" shows "Cong P Q P' Q'" using Reflect_def assms(1) assms(2) cong_4321 l10_10_spec l7_13 by auto lemma image_preserves_bet: assumes "A A' ReflectL X Y" and "B B' ReflectL X Y" and "C C' ReflectL X Y" and "Bet A B C" shows "Bet A' B' C'" proof - have P3: "A B C Cong3 A' B' C'" using Cong3_def assms(1) assms(2) assms(3) l10_10_spec l10_4_spec by blast then show ?thesis using assms(4) l4_6 by blast qed lemma image_gen_preserves_bet: assumes "A A' Reflect X Y" and "B B' Reflect X Y" and "C C' Reflect X Y" and "Bet A B C" shows "Bet A' B' C'" proof cases assume "X = Y" then show ?thesis by (metis (full_types) assms(1) assms(2) assms(3) assms(4) image__midpoint l7_15 l7_2) next assume P1: "X \ Y" then have P2: "A A' ReflectL X Y" using assms(1) is_image_is_image_spec by blast have P3: "B B' ReflectL X Y" using P1 assms(2) is_image_is_image_spec by auto have "C C' ReflectL X Y" using P1 assms(3) is_image_is_image_spec by blast then show ?thesis using image_preserves_bet using assms(4) P2 P3 by blast qed lemma image_preserves_col: assumes "A A' ReflectL X Y" and "B B' ReflectL X Y" and "C C' ReflectL X Y" and "Col A B C" shows "Col A' B' C'" using image_preserves_bet using Col_def assms(1) assms(2) assms(3) assms(4) by auto lemma image_gen_preserves_col: assumes "A A' Reflect X Y" and "B B' Reflect X Y" and "C C' Reflect X Y" and "Col A B C" shows "Col A' B' C'" using Col_def assms(1) assms(2) assms(3) assms(4) image_gen_preserves_bet by auto lemma image_gen_preserves_ncol: assumes "A A' Reflect X Y" and "B B' Reflect X Y" and "C C' Reflect X Y" and "\ Col A B C" shows "\ Col A' B' C'" using assms(1) assms(2) assms(3) assms(4)image_gen_preserves_col l10_4 by blast lemma image_gen_preserves_inter: assumes "A A' Reflect X Y" and "B B' Reflect X Y" and "C C' Reflect X Y" and "D D' Reflect X Y" and "\ Col A B C" and "C \ D" and "Col A B I" and "Col C D I" and "Col A' B' I'" and "Col C' D' I'" shows "I I' Reflect X Y" proof - obtain I0 where P1: "I I0 Reflect X Y" using l10_6_existence by blast then show ?thesis by (smt Tarski_neutral_dimensionless.image_gen_preserves_col Tarski_neutral_dimensionless_axioms assms(1) assms(10) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9) l10_4 l10_7 l6_21) qed lemma intersection_with_image_gen: assumes "A A' Reflect X Y" and "B B' Reflect X Y" and "\ Col A B A'" and "Col A B C" and "Col A' B' C" shows "Col C X Y" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) image_gen_preserves_inter l10_2_uniqueness l10_4 l10_8 not_col_distincts) lemma image_preserves_midpoint : assumes "A A' ReflectL X Y" and "B B' ReflectL X Y" and "C C' ReflectL X Y" and "A Midpoint B C" shows "A' Midpoint B' C'" proof - have P1: "Bet B' A' C'" using image_preserves_bet using assms(1) assms(2) assms(3) assms(4) midpoint_bet by auto have "Cong B' A' A' C'" by (metis Cong_perm Tarski_neutral_dimensionless.l10_10_spec Tarski_neutral_dimensionless.l7_13 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong_transitivity l7_3_2) then show ?thesis by (simp add: Midpoint_def P1) qed lemma image_spec_preserves_per: assumes "A A' ReflectL X Y" and "B B' ReflectL X Y" and "C C' ReflectL X Y" and "Per A B C" shows "Per A' B' C'" proof cases assume "X = Y" then show ?thesis using assms(1) assms(2) assms(3) assms(4) image_spec__eq by blast next assume P1: "X \ Y" obtain C1 where P2: "B Midpoint C C1" using symmetric_point_construction by blast obtain C1' where P3: "C1 C1' ReflectL X Y" by (meson P1 l10_6_existence_spec) then have P4: "B' Midpoint C' C1'" using P2 assms(2) assms(3) image_preserves_midpoint by blast have "Cong A' C' A' C1'" proof - have Q1: "Cong A' C' A C" using assms(1) assms(3) l10_10_spec by auto have "Cong A C A' C1'" by (metis P2 P3 Tarski_neutral_dimensionless.l10_10_spec Tarski_neutral_dimensionless_axioms assms(1) assms(4) cong_inner_transitivity cong_symmetry per_double_cong) then show ?thesis using Q1 cong_transitivity by blast qed then show ?thesis using P4 Per_def by blast qed lemma image_preserves_per: assumes "A A' Reflect X Y" and "B B' Reflect X Y"and "C C' Reflect X Y" and "Per A B C" shows "Per A' B' C'" proof cases assume "X = Y" then show ?thesis using midpoint_preserves_per using assms(1) assms(2) assms(3) assms(4) image__midpoint l7_2 by blast next assume P1: "X \ Y" have P2: "X \ Y \ A A' ReflectL X Y" using P1 assms(1) is_image_is_image_spec by blast have P3: "X \ Y \ B B' ReflectL X Y" using P1 assms(2) is_image_is_image_spec by blast have P4: "X \ Y \ C C' ReflectL X Y" using P1 assms(3) is_image_is_image_spec by blast then show ?thesis using image_spec_preserves_per using P2 P3 assms(4) by blast qed lemma l10_12: assumes "Per A B C" and "Per A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "Cong A C A' C'" proof cases assume P1: "B = C" then have "B' = C'" using assms(4) cong_reverse_identity by blast then show ?thesis using P1 assms(3) by blast next assume P2: "B \ C" have "Cong A C A' C'" proof cases assume "A = B" then show ?thesis using assms(3) assms(4) cong_diff_3 by force next assume P3: "A \ B" obtain X where P4: "X Midpoint B B'" using midpoint_existence by blast obtain A1 where P5: "X Midpoint A' A1" using Mid_perm symmetric_point_construction by blast obtain C1 where P6: "X Midpoint C' C1" using Mid_perm symmetric_point_construction by blast have Q1: "A' B' C' Cong3 A1 B C1" using Cong3_def P4 P5 P6 l7_13 l7_2 by blast have Q2: "Per A1 B C1" using assms(2)Q1 l8_10 by blast have Q3: "Cong A B A1 B" by (metis Cong3_def Q1 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms assms(3) cong_inner_transitivity) have Q4: "Cong B C B C1" by (metis Cong3_def Q1 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms assms(4) cong_inner_transitivity) obtain Y where P7: "Y Midpoint C C1" using midpoint_existence by auto then have R1: "C1 C Reflect B Y" using cong_midpoint__image using Q4 by blast obtain A2 where R2: "A1 A2 Reflect B Y" using l10_6_existence by blast have R3: "Cong C A2 C1 A1" using R1 R2 l10_10 by blast have R5: "B B Reflect B Y" using image_triv by blast have R6: "Per A2 B C" using image_preserves_per using Q2 R1 R2 image_triv by blast have R7: "Cong A B A2 B" using l10_10 Cong_perm Q3 R2 cong_transitivity image_triv by blast obtain Z where R7A: "Z Midpoint A A2" using midpoint_existence by blast have "Cong B A B A2" using Cong_perm R7 by blast then have T1: "A2 A Reflect B Z" using R7A cong_midpoint__image by blast obtain C0 where T2: "B Midpoint C C0" using symmetric_point_construction by blast have T3: "Cong A C A C0" using T2 assms(1) per_double_cong by blast have T4: "Cong A2 C A2 C0" using R6 T2 per_double_cong by blast have T5: "C0 C Reflect B Z" proof - have "C0 C Reflect Z B" proof cases assume "A = A2" then show ?thesis by (metis R7A T2 T3 cong_midpoint__image midpoint_distinct_3) next assume "A \ A2" then show ?thesis using l4_17 cong_midpoint__image by (metis R7A T2 T3 T4 midpoint_col not_col_permutation_3) qed then show ?thesis using is_image_rev by blast qed have T6: "Cong A C A2 C0" using T1 T5 l10_10 by auto have R4: "Cong A C A2 C" by (metis T4 T6 Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms cong_inner_transitivity) then have Q5: "Cong A C A1 C1" by (meson R3 cong_inner_transitivity not_cong_3421) then show ?thesis using Cong3_def Q1 Q5 cong_symmetry cong_transitivity by blast qed then show ?thesis by blast qed lemma l10_16: assumes "\ Col A B C" and "\ Col A' B' P" and "Cong A B A' B'" shows "\ C'. A B C Cong3 A' B' C' \ A' B' OS P C'" proof cases assume "A = B" then show ?thesis using assms(1) not_col_distincts by auto next assume P1: "A \ B" obtain X where P2: "Col A B X \ A B Perp C X" using assms(1) l8_18_existence by blast obtain X' where P3: "A B X Cong3 A' B' X'" using P2 assms(3) l4_14 by blast obtain Q where P4: "A' B' Perp Q X' \ A' B' OS P Q" using P2 P3 assms(2) l10_15 l4_13 by blast obtain C' where P5: "X' Out C' Q \ Cong X' C' X C" by (metis P2 P4 l6_11_existence perp_distinct) have P6: "Cong A C A' C'" proof cases assume "A = X" then show ?thesis by (metis Cong3_def P3 P5 cong_4321 cong_commutativity cong_diff_3) next assume "A \ X" have P7: "Per A X C" using P2 col_trivial_3 l8_16_1 l8_2 by blast have P8: "Per A' X' C'" proof - have "X' PerpAt A' X' X' C'" proof - have Z1: "A' X' Perp X' C'" proof - have W1: "X' \ C'" using P5 out_distinct by blast have W2: "X' Q Perp A' B'" using P4 Perp_perm by blast then have "X' C' Perp A' B'" by (metis P5 Perp_perm W1 col_trivial_3 not_col_permutation_5 out_col perp_col2_bis) then show ?thesis by (metis Cong3_def P2 P3 Perp_perm \A \ X\ col_trivial_3 cong_identity l4_13 perp_col2_bis) qed have Z2: "Col X' A' X'" using not_col_distincts by blast have "Col X' X' C'" by (simp add: col_trivial_1) then show ?thesis by (simp add: Z1 Z2 l8_14_2_1b_bis) qed then show ?thesis by (simp add: perp_in_per) qed have P9: "Cong A X A' X'" using Cong3_def P3 by auto have "Cong X C X' C'" using Cong_perm P5 by blast then show ?thesis using l10_12 using P7 P8 P9 by blast qed have P10: "Cong B C B' C'" proof cases assume "B = X" then show ?thesis by (metis Cong3_def P3 P5 cong_4321 cong_commutativity cong_diff_3) next assume "B \ X" have Q1: "Per B X C" using P2 col_trivial_2 l8_16_1 l8_2 by blast have "X' PerpAt B' X' X' C'" proof - have Q2: "B' X' Perp X' C'" proof - have R1: "B' \ X'" using Cong3_def P3 \B \ X\ cong_identity by blast have "X' C' Perp B' A'" proof - have S1: "X' \ C'" using Out_def P5 by blast have S2: "X' Q Perp B' A'" using P4 Perp_perm by blast have "Col X' Q C'" using Col_perm P5 out_col by blast then show ?thesis using S1 S2 perp_col by blast qed then have R2: "B' A' Perp X' C'" using Perp_perm by blast have R3: "Col B' A' X'" using Col_perm P2 P3 l4_13 by blast then show ?thesis using R1 R2 perp_col by blast qed have Q3: "Col X' B' X'" by (simp add: col_trivial_3) have "Col X' X' C'" by (simp add: col_trivial_1) then show ?thesis using l8_14_2_1b_bis using Q2 Q3 by blast qed then have Q2: "Per B' X' C'" by (simp add: perp_in_per) have Q3: "Cong B X B' X'" using Cong3_def P3 by blast have Q4: "Cong X C X' C'" using P5 not_cong_3412 by blast then show ?thesis using Q1 Q2 Q3 l10_12 by blast qed have P12: "A' B' OS C' Q \ X' Out C' Q \ \ Col A' B' C'" using l9_19 l4_13 by (meson P2 P3 P5 one_side_not_col123 out_one_side_1) then have P13: "A' B' OS C' Q" using l4_13 by (meson P2 P3 P4 P5 l6_6 one_side_not_col124 out_one_side_1) then show ?thesis using Cong3_def P10 P4 P6 assms(3) one_side_symmetry one_side_transitivity by blast qed lemma cong_cop_image__col: assumes "P \ P'" and "P P' Reflect A B" and "Cong P X P' X" and "Coplanar A B P X" shows "Col A B X" proof - have P1: "(A \ B \ P P' ReflectL A B) \ (A = B \ A Midpoint P' P)" by (metis assms(2) image__midpoint is_image_is_image_spec) { assume Q1: "A \ B \ P P' ReflectL A B" then obtain M where Q2: "M Midpoint P' P \ Col A B M" using ReflectL_def by blast have "Col A B X" proof cases assume R1: "A = M" have R2: "P A Perp A B" proof - have S1: "P \ A" using Q2 R1 assms(1) midpoint_distinct_2 by blast have S2: "P P' Perp A B" using Perp_perm Q1 ReflectL_def assms(1) by blast have "Col P P' A" using Q2 R1 midpoint_col not_col_permutation_3 by blast then show ?thesis using perp_col using S1 S2 by blast qed have R3: "Per P A B" by (simp add: R2 perp_comm perp_per_1) then have R3A: "Per B A P" using l8_2 by blast have "A Midpoint P P' \ Cong X P X P'" using Cong_cases Q2 R1 assms(3) l7_2 by blast then have R4: "Per X A P" using Per_def by blast have R5: "Coplanar P B X A" using assms(4) ncoplanar_perm_20 by blast have "P \ A" using R2 perp_not_eq_1 by auto then show ?thesis using R4 R5 R3A using cop_per2__col not_col_permutation_1 by blast next assume T1: "A \ M" have T3: "P \ M" using Q2 assms(1) l7_3_2 sym_preserve_diff by blast have T2: "P M Perp M A" proof - have T4: "P P' Perp M A" using Perp_perm Q1 Q2 ReflectL_def T1 assms(1) col_trivial_3 perp_col0 by blast have "Col P P' M" by (simp add: Col_perm Q2 midpoint_col) then show ?thesis using T3 T4 perp_col by blast qed then have "M P Perp A M" using perp_comm by blast then have "M PerpAt M P A M" using perp_perp_in by blast then have "M PerpAt P M M A" by (simp add: perp_in_comm) then have U1: "Per P M A" by (simp add: perp_in_per) have U2: "Per X M P" using l7_2 cong_commutativity using Per_def Q2 assms(3) by blast have "Col A X M" proof - have W2: "Coplanar P A X M" by (meson Q1 Q2 assms(4) col_cop2__cop coplanar_perm_13 ncop_distincts) have "Per A M P" by (simp add: U1 l8_2) then show ?thesis using cop_per2__col using U2 T3 W2 by blast qed then show ?thesis using Q2 T1 col2__eq not_col_permutation_4 by blast qed } then have P2: "(A \ B \ P P' ReflectL A B) \ Col A B X" by blast have "(A = B \ A Midpoint P' P) \ Col A B X" using col_trivial_1 by blast then show ?thesis using P1 P2 by blast qed lemma cong_cop_per2_1: assumes "A \ B" and "Per A B X" and "Per A B Y" and "Cong B X B Y" and "Coplanar A B X Y" shows "X = Y \ B Midpoint X Y" by (meson Per_cases assms(1) assms(2) assms(3) assms(4) assms(5) cop_per2__col coplanar_perm_3 l7_20_bis not_col_permutation_5) lemma cong_cop_per2: assumes "A \ B" and "Per A B X" and "Per A B Y" and "Cong B X B Y" and "Coplanar A B X Y" shows "X = Y \ X Y ReflectL A B" proof - have "X = Y \ B Midpoint X Y" using assms(1) assms(2) assms(3) assms(4) assms(5) cong_cop_per2_1 by blast then show ?thesis by (metis Mid_perm Per_def Reflect_def assms(1) assms(3) cong_midpoint__image symmetric_point_uniqueness) qed lemma cong_cop_per2_gen: assumes "A \ B" and "Per A B X" and "Per A B Y" and "Cong B X B Y" and "Coplanar A B X Y" shows "X = Y \ X Y Reflect A B" proof - have "X = Y \ B Midpoint X Y" using assms(1) assms(2) assms(3) assms(4) assms(5) cong_cop_per2_1 by blast then show ?thesis using assms(2) cong_midpoint__image l10_4 per_double_cong by blast qed lemma ex_perp_cop: assumes "A \ B" shows "\ Q. A B Perp Q C \ Coplanar A B P Q" proof - { assume "Col A B C \ Col A B P" then have "\ Q. A B Perp Q C \ Coplanar A B P Q" using assms ex_ncol_cop l10_15 ncop__ncols by blast } then have T1: "(Col A B C \ Col A B P) \ (\ Q. A B Perp Q C \ Coplanar A B P Q)" by blast { assume "\Col A B C \ Col A B P" then have "\ Q. A B Perp Q C \ Coplanar A B P Q" by (metis Perp_cases ncop__ncols not_col_distincts perp_exists) } then have T2: "(\Col A B C \ Col A B P) \ (\ Q. A B Perp Q C \ Coplanar A B P Q)" by blast { assume "Col A B C \ \Col A B P" then have "\ Q. A B Perp Q C \ Coplanar A B P Q" using l10_15 os__coplanar by blast } then have T3: "(Col A B C \ \Col A B P) \ (\ Q. A B Perp Q C \ Coplanar A B P Q)" by blast { assume "\Col A B C \ \Col A B P" then have "\ Q. A B Perp Q C \ Coplanar A B P Q" using l8_18_existence ncop__ncols perp_right_comm by blast } then have "(\Col A B C \ \Col A B P) \ (\ Q. A B Perp Q C \ Coplanar A B P Q)" by blast then show ?thesis using T1 T2 T3 by blast qed lemma hilbert_s_version_of_pasch_aux: assumes "Coplanar A B C P" and "\ Col A I P" and "\ Col B C P" and "Bet B I C" and "B \ I" and "I \ C" and "B \ C" shows "\ X. Col I P X \ ((Bet A X B \ A \ X \ X \ B \ A \ B) \ (Bet A X C \ A \ X \ X \ C \ A \ C))" proof - have T1: "I P TS B C" using Col_perm assms(3) assms(4) assms(5) assms(6) bet__ts bet_col col_transitivity_1 by blast have T2: "Coplanar A P B I" using assms(1) assms(4) bet_cop__cop coplanar_perm_6 ncoplanar_perm_9 by blast have T3: "I P TS A B \ I P TS A C" by (meson T1 T2 TS_def assms(2) cop_nos__ts coplanar_perm_21 l9_2 l9_8_2) have T4: "I P TS A B \ (\ X. Col I P X \ ((Bet A X B \ A \ X \ X \ B \ A \ B) \ (Bet A X C \ A \ X \ X \ C \ A \ C)))" by (metis TS_def not_col_permutation_2 ts_distincts) have "I P TS A C \ (\ X. Col I P X \ ((Bet A X B \ A \ X \ X \ B \ A \ B) \ (Bet A X C \ A \ X \ X \ C \ A \ C)))" by (metis TS_def not_col_permutation_2 ts_distincts) then show ?thesis using T3 T4 by blast qed lemma hilbert_s_version_of_pasch: assumes "Coplanar A B C P" and "\ Col C Q P" and "\ Col A B P" and "BetS A Q B" shows "\ X. Col P Q X \ (BetS A X C \ BetS B X C)" proof - obtain X where "Col Q P X \ (Bet C X A \ C \ X \ X \ A \ C \ A \ Bet C X B \ C \ X \ X \ B \ C \ B)" using BetSEq assms(1) assms(2) assms(3) assms(4) coplanar_perm_12 hilbert_s_version_of_pasch_aux by fastforce then show ?thesis by (metis BetS_def Bet_cases Col_perm) qed lemma two_sides_cases: assumes "\ Col PO A B" and "PO P OS A B" shows "PO A TS P B \ PO B TS P A" by (meson assms(1) assms(2) cop_nts__os l9_31 ncoplanar_perm_3 not_col_permutation_4 one_side_not_col124 one_side_symmetry os__coplanar) lemma not_par_two_sides: assumes "C \ D" and "Col A B I" and "Col C D I" and "\ Col A B C" shows "\ X Y. Col C D X \ Col C D Y \ A B TS X Y" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. Bet p pa (pp p pa) \ pa \ (pp p pa)" by (meson point_construction_different) then have f2: "\p pa pb pc. (Col pc pb p \ \ Col pc pb (pp p pa)) \ \ Col pc pb pa" by (meson Col_def colx) have f3: "\p pa. Col pa p pa" by (meson Col_def between_trivial) have f4: "\p pa. Col pa p p" by (meson Col_def between_trivial) have f5: "Col I D C" by (meson Col_perm assms(3)) have f6: "\p pa. Col (pp pa p) p pa" using f4 f3 f2 by blast then have f7: "\p pa. Col pa (pp pa p) p" by (meson Col_perm) then have f8: "\p pa pb pc. (pc pb TS p (pp p pa) \ Col pc pb p) \ \ Col pc pb pa" using f2 f1 by (meson l9_18) have "I = D \ Col D (pp D I) C" using f7 f5 f3 colx by blast then have "I = D \ Col C D (pp D I)" using Col_perm by blast then show ?thesis using f8 f6 f3 by (metis Col_perm assms(2) assms(4)) qed lemma cop_not_par_other_side: assumes "C \ D" and "Col A B I" and "Col C D I" and "\ Col A B C" and "\ Col A B P" and "Coplanar A B C P" shows "\ Q. Col C D Q \ A B TS P Q" proof - obtain X Y where P1:"Col C D X \ Col C D Y \ A B TS X Y" using assms(1) assms(2) assms(3) assms(4) not_par_two_sides by blast then have "Coplanar C A B X" using Coplanar_def assms(1) assms(2) assms(3) col_transitivity_1 by blast then have "Coplanar A B P X" using assms(4) assms(6) col_permutation_3 coplanar_trans_1 ncoplanar_perm_2 ncoplanar_perm_6 by blast then show ?thesis by (meson P1 l9_8_2 TS_def assms(5) cop_nts__os not_col_permutation_2 one_side_symmetry) qed lemma cop_not_par_same_side: assumes "C \ D" and "Col A B I" and "Col C D I" and "\ Col A B C" and "\ Col A B P" and "Coplanar A B C P" shows "\ Q. Col C D Q \ A B OS P Q" proof - obtain X Y where P1: "Col C D X \ Col C D Y \ A B TS X Y" using assms(1) assms(2) assms(3) assms(4) not_par_two_sides by blast then have "Coplanar C A B X" using Coplanar_def assms(1) assms(2) assms(3) col_transitivity_1 by blast then have "Coplanar A B P X" using assms(4) assms(6) col_permutation_1 coplanar_perm_2 coplanar_trans_1 ncoplanar_perm_14 by blast then show ?thesis by (meson P1 TS_def assms(5) cop_nts__os l9_2 l9_8_1 not_col_permutation_2) qed end subsubsection "Line reflexivity: 2D" context Tarski_2D begin lemma all_coplanar: "Coplanar A B C D" proof - have "\ A B C P Q. P \ Q \ Cong A P A Q \ Cong B P B Q\ Cong C P C Q \ (Bet A B C \ Bet B C A \ Bet C A B)" using upper_dim by blast then show ?thesis using upper_dim_implies_all_coplanar by (smt Tarski_neutral_dimensionless.not_col_permutation_2 Tarski_neutral_dimensionless_axioms all_coplanar_axiom_def all_coplanar_implies_upper_dim coplanar_perm_9 ncop__ncol os__coplanar ts__coplanar upper_dim_implies_not_one_side_two_sides) qed lemma per2__col: assumes "Per A X C" and "X \ C" and "Per B X C" shows "Col A B X" using all_coplanar_axiom_def all_coplanar_upper_dim assms(1) assms(2) assms(3) upper_dim upper_dim_implies_per2__col by blast lemma perp2__col: assumes "X Y Perp A B" and "X Z Perp A B" shows "Col X Y Z" by (meson Tarski_neutral_dimensionless.cop_perp2__col Tarski_neutral_dimensionless_axioms all_coplanar assms(1) assms(2)) end subsection "Angles" subsubsection "Some generalites" context Tarski_neutral_dimensionless begin lemma l11_3: assumes "A B C CongA D E F" shows "\ A' C' D' F'. B Out A' A \ B Out C C' \ E Out D' D \ E Out F F' \ A' B C' Cong3 D' E F'" proof - obtain A' C' D' F' where P1: "Bet B A A' \ Cong A A' E D \ Bet B C C' \ Cong C C' E F \ Bet E D D' \ Cong D D' B A \ Bet E F F' \ Cong F F' B C \ Cong A' C' D' F'" using CongA_def using assms by auto then have "A' B C' Cong3 D' E F'" by (meson Cong3_def between_symmetry l2_11_b not_cong_1243 not_cong_4312) thus ?thesis by (metis CongA_def P1 assms bet_out l6_6) qed lemma l11_aux: assumes "B Out A A'" and "E Out D D'" and "Cong B A' E D'" and "Bet B A A0" and "Bet E D D0" and "Cong A A0 E D" and "Cong D D0 B A" shows "Cong B A0 E D0 \ Cong A' A0 D' D0" proof - have P2: "Cong B A0 E D0" by (meson Bet_cases assms(4) assms(5) assms(6) assms(7) l2_11_b not_cong_1243 not_cong_4312) have P3: "Bet B A A' \ Bet B A' A" using Out_def assms(1) by auto have P4: "Bet E D D' \ Bet E D' D" using Out_def assms(2) by auto have P5: "Bet B A A' \ Cong A' A0 D' D0" by (smt P2 assms(1) assms(2) assms(3) assms(4) assms(5) bet_out l6_6 l6_7 out_cong_cong out_diff1) have P6: "Bet B A' A \ Cong A' A0 D' D0" proof - have "E Out D D0" using assms(2) assms(5) bet_out out_diff1 by blast thus ?thesis by (meson P2 assms(2) assms(3) assms(4) between_exchange4 cong_preserves_bet l4_3_1 l6_6 l6_7) qed have P7: "Bet E D D' \ Cong A' A0 D' D0" using P3 P5 P6 by blast have "Bet E D' D \ Cong A' A0 D' D0" using P3 P5 P6 by blast thus ?thesis using P2 P3 P4 P5 P6 P7 by blast qed lemma l11_3_bis: assumes "\ A' C' D' F'. (B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ A' B C' Cong3 D' E F')" shows "A B C CongA D E F" proof - obtain A' C' D' F' where P1: "B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ A' B C' Cong3 D' E F'" using assms by blast obtain A0 where P2: "Bet B A A0 \ Cong A A0 E D" using segment_construction by presburger obtain C0 where P3: "Bet B C C0 \ Cong C C0 E F" using segment_construction by presburger obtain D0 where P4: "Bet E D D0 \ Cong D D0 B A" using segment_construction by presburger obtain F0 where P5: "Bet E F F0 \ Cong F F0 B C" using segment_construction by presburger have P6: "A \ B \ C \ B \ D \ E \ F \ E" using P1 out_diff2 by blast have "Cong A0 C0 D0 F0" proof - have Q1: "Cong B A0 E D0 \ Cong A' A0 D' D0" proof - have R1: "B Out A A'" by (simp add: P1 l6_6) have R2: "E Out D D'" by (simp add: P1 l6_6) have "Cong B A' E D'" using Cong3_def P1 cong_commutativity by blast thus ?thesis using l11_aux using P2 P4 R1 R2 by blast qed have Q2: "Cong B C0 E F0 \ Cong C' C0 F' F0" by (smt Cong3_def Out_cases P1 P3 P5 Tarski_neutral_dimensionless.l11_aux Tarski_neutral_dimensionless_axioms) have Q3: "B A' A0 Cong3 E D' D0" by (meson Cong3_def P1 Q1 cong_3_swap) have Q4: "B C' C0 Cong3 E F' F0" using Cong3_def P1 Q2 by blast have "Cong C0 A' F0 D'" proof - have R1: "B C' C0 A' FSC E F' F0 D'" proof - have S1: "Col B C' C0" by (metis (no_types) Col_perm P1 P3 P6 bet_col col_transitivity_1 out_col) have S3: "Cong B A' E D'" using Cong3_def Q3 by blast have "Cong C' A' F' D'" using Cong3_def P1 cong_commutativity by blast thus ?thesis by (simp add: FSC_def S1 Q4 S3) qed have "B \ C'" using P1 out_distinct by blast thus ?thesis using R1 l4_16 by blast qed then have Q6: "B A' A0 C0 FSC E D' D0 F0" by (meson FSC_def P1 P2 P6 Q2 Q3 bet_out l6_7 not_cong_2143 out_col) have "B \ A'" using Out_def P1 by blast thus ?thesis using Q6 l4_16 by blast qed thus ?thesis using P6 P2 P3 P4 P5 CongA_def by auto qed lemma l11_4_1: assumes "A B C CongA D E F" and (*"A \ B" and "C \ B" and "D \ E" and "F \ E" and*) "B Out A' A" and "B Out C' C" and "E Out D' D" and "E Out F' F" and "Cong B A' E D'" and "Cong B C' E F'" shows "Cong A' C' D' F'" proof - obtain A0 C0 D0 F0 where P1: "B Out A0 A \ B Out C C0 \ E Out D0 D \ E Out F F0 \ A0 B C0 Cong3 D0 E F0" using assms(1) l11_3 by blast have P2: "B Out A' A0" using P1 assms(2) l6_6 l6_7 by blast have P3: "E Out D' D0" by (meson P1 assms(4) l6_6 l6_7) have P4: "Cong A' A0 D' D0" proof - have "Cong B A0 E D0" using Cong3_def P1 cong_3_swap by blast thus ?thesis using P2 P3 using assms(6) out_cong_cong by blast qed have P5: "Cong A' C0 D' F0" proof - have P6: "B A0 A' C0 FSC E D0 D' F0" by (meson Cong3_def Cong_perm FSC_def P1 P2 P4 assms(6) not_col_permutation_5 out_col) thus ?thesis using P2 Tarski_neutral_dimensionless.l4_16 Tarski_neutral_dimensionless_axioms out_diff2 by fastforce qed have P6: "B Out C' C0" using P1 assms(3) l6_7 by blast have "E Out F' F0" using P1 assms(5) l6_7 by blast then have "Cong C' C0 F' F0" using Cong3_def P1 P6 assms(7) out_cong_cong by auto then have P9: "B C0 C' A' FSC E F0 F' D'" by (smt Cong3_def Cong_perm FSC_def P1 P5 P6 assms(6) assms(7) not_col_permutation_5 out_col) then have "Cong C' A' F' D'" using P6 Tarski_neutral_dimensionless.l4_16 Tarski_neutral_dimensionless_axioms out_diff2 by fastforce thus ?thesis using Tarski_neutral_dimensionless.not_cong_2143 Tarski_neutral_dimensionless_axioms by fastforce qed lemma l11_4_2: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" and "\ A' C' D' F'. (B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ Cong B A' E D' \ Cong B C' E F' \ Cong A' C' D' F')" shows "A B C CongA D E F" proof - obtain A' where P1: "Bet B A A' \ Cong A A' E D" using segment_construction by fastforce obtain C' where P2: "Bet B C C' \ Cong C C' E F" using segment_construction by fastforce obtain D' where P3: "Bet E D D' \ Cong D D' B A" using segment_construction by fastforce obtain F' where P4: "Bet E F F' \ Cong F F' B C" using segment_construction by fastforce have P5: "Cong A' B D' E" by (meson Bet_cases P1 P3 l2_11_b not_cong_1243 not_cong_4312) have P6: "Cong B C' E F'" by (meson P2 P4 between_symmetry cong_3421 cong_right_commutativity l2_11_b) have "B Out A' A \ B Out C' C \ E Out D' D \ E Out F' F \ A' B C' Cong3 D' E F'" by (metis (no_types, lifting) Cong3_def P1 P2 P3 P4 P5 P6 Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) bet_neq12__neq cong_commutativity) thus ?thesis using l11_3_bis by blast qed lemma conga_refl: assumes "A \ B" and "C \ B" shows "A B C CongA A B C" by (meson CongA_def assms(1) assms(2) cong_reflexivity segment_construction) lemma conga_sym: assumes "A B C CongA A' B' C'" shows "A' B' C' CongA A B C" proof - obtain A0 C0 D0 F0 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' D0 \ Cong A' D0 B A \ Bet B' C' F0 \ Cong C' F0 B C \ Cong A0 C0 D0 F0" using CongA_def assms by auto thus ?thesis proof - have "\p pa pb pc. Bet B' A' p \ Cong A' p B A \ Bet B' C' pa \ Cong C' pa B C \Bet B A pb \ Cong A pb B' A' \Bet B C pc \ Cong C pc B' C' \ Cong p pa pb pc" by (metis (no_types) Tarski_neutral_dimensionless.cong_symmetry Tarski_neutral_dimensionless_axioms P1) thus ?thesis using CongA_def assms by auto qed qed lemma l11_10: assumes "A B C CongA D E F" and "B Out A' A" and "B Out C' C" and "E Out D' D" and "E Out F' F" shows "A' B C' CongA D' E F'" proof - have P1: "A' \ B" using assms(2) out_distinct by auto have P2: "C' \ B" using Out_def assms(3) by force have P3: "D' \ E" using Out_def assms(4) by blast have P4: "F' \ E" using assms(5) out_diff1 by auto have P5: "\ A'0 C'0 D'0 F'0. (B Out A'0 A' \ B Out C'0 C' \ E Out D'0 D' \ E Out F'0 F' \ Cong B A'0 E D'0 \ Cong B C'0 E F'0) \ Cong A'0 C'0 D'0 F'0" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) l11_4_1 l6_7) thus ?thesis using P1 P2 P3 P4 P5 l11_4_2 by blast qed lemma out2__conga: assumes "B Out A' A" and "B Out C' C" shows "A B C CongA A' B C'" by (smt assms(1) assms(2) between_trivial2 conga_refl l11_10 out2_bet_out out_distinct) lemma cong3_diff: assumes "A \ B" and "A B C Cong3 A' B' C'" shows "A' \ B'" using Cong3_def assms(1) assms(2) cong_diff by blast lemma cong3_diff2: assumes "B \ C" and "A B C Cong3 A' B' C'" shows "B' \ C'" using Cong3_def assms(1) assms(2) cong_diff by blast lemma cong3_conga: assumes "A \ B" and "C \ B" and "A B C Cong3 A' B' C'" shows "A B C CongA A' B' C'" by (metis assms(1) assms(2) assms(3) cong3_diff cong3_diff2 l11_3_bis out_trivial) lemma cong3_conga2: assumes "A B C Cong3 A' B' C'" and "A B C CongA A'' B'' C''" shows "A' B' C' CongA A'' B'' C''" proof - obtain A0 C0 A2 C2 where P1: "Bet B A A0 \ Cong A A0 B'' A'' \ Bet B C C0 \ Cong C C0 B'' C''\ Bet B'' A'' A2 \ Cong A'' A2 B A \ Bet B'' C'' C2 \ Cong C'' C2 B C \ Cong A0 C0 A2 C2" using CongA_def assms(2) by auto obtain A1 where P5: "Bet B' A' A1 \ Cong A' A1 B'' A''" using segment_construction by blast obtain C1 where P6: "Bet B' C' C1 \ Cong C' C1 B'' C''" using segment_construction by blast have P7: "Cong A A0 A' A1" proof - have "Cong B'' A'' A' A1" using P5 using Cong_perm by blast thus ?thesis using Cong_perm P1 cong_inner_transitivity by blast qed have P8: "Cong B A0 B' A1" using Cong3_def P1 P5 P7 assms(1) cong_commutativity l2_11_b by blast have P9: "Cong C C0 C' C1" using P1 P6 cong_inner_transitivity cong_symmetry by blast have P10: "Cong B C0 B' C1" using Cong3_def P1 P6 P9 assms(1) l2_11_b by blast have "B A A0 C FSC B' A' A1 C'" using FSC_def P1 P5 P7 P8 Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless_axioms assms(1) bet_col l4_3 by fastforce then have P12: "Cong A0 C A1 C'" using CongA_def assms(2) l4_16 by auto then have "B C C0 A0 FSC B' C' C1 A1" using Cong3_def FSC_def P1 P10 P8 P9 assms(1) bet_col cong_commutativity by auto then have P13: "Cong C0 A0 C1 A1" using l4_16 CongA_def assms(2) by blast have Q2: "Cong A' A1 B'' A''" using P1 P7 cong_inner_transitivity by blast have Q5: "Bet B'' A'' A2" using P1 by blast have Q6: "Cong A'' A2 B' A'" proof - have "Cong B A B' A'" using P1 P7 P8 P5 l4_3 by blast thus ?thesis using P1 cong_transitivity by blast qed have Q7: "Bet B'' C'' C2" using P1 by blast have Q8: "Cong C'' C2 B' C'" proof - have "Cong B C B' C'" using Cong3_def assms(1) by blast thus ?thesis using P1 cong_transitivity by blast qed have R2: "Cong C0 A0 C2 A2" using Cong_cases P1 by blast have "Cong C1 A1 A0 C0" using Cong_cases P13 by blast then have Q9: "Cong A1 C1 A2 C2" using R2 P13 cong_inner_transitivity not_cong_4321 by blast thus ?thesis using CongA_def P5 Q2 P6 Q5 Q6 Q7 Q8 by (metis assms(1) assms(2) cong3_diff cong3_diff2) qed lemma conga_diff1: assumes "A B C CongA A' B' C'" shows "A \ B" using CongA_def assms by blast lemma conga_diff2: assumes "A B C CongA A' B' C'" shows "C \ B" using CongA_def assms by blast lemma conga_diff45: assumes "A B C CongA A' B' C'" shows "A' \ B'" using CongA_def assms by blast lemma conga_diff56: assumes "A B C CongA A' B' C'" shows "C' \ B'" using CongA_def assms by blast lemma conga_trans: assumes "A B C CongA A' B' C'" and "A' B' C' CongA A'' B'' C''" shows "A B C CongA A'' B'' C''" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C'\ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(1) by auto have P2: "A''\ B'' \ C'' \ B''" using CongA_def assms(2) by auto have P3: "A1 B' C1 CongA A'' B'' C''" proof - have L2: "B' Out A1 A'" using P1 by (metis Out_def assms(2) bet_neq12__neq conga_diff1) have L3: "B' Out C1 C'" using P1 by (metis Out_def assms(1) bet_neq12__neq conga_diff56) have L4: "B'' Out A'' A''" using P2 out_trivial by auto have "B'' Out C'' C''" by (simp add: P2 out_trivial) thus ?thesis using assms(2) L2 L3 L4 l11_10 by blast qed have L6: "A0 B C0 CongA A' B' C'" by (smt Out_cases P1 Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless.conga_diff45 Tarski_neutral_dimensionless_axioms assms(1) bet_out conga_diff56 l11_10 l5_3) have L7: "Cong B A0 B' A1" by (meson P1 between_symmetry cong_3421 l2_11_b not_cong_1243) have L8: "Cong B C0 B' C1" using P1 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast have L10: "A0 B C0 Cong3 A1 B' C1" by (simp add: Cong3_def L7 L8 P1 cong_commutativity) then have L11: "A0 B C0 CongA A'' B'' C''" by (meson Tarski_neutral_dimensionless.cong3_conga2 Tarski_neutral_dimensionless_axioms P3 cong_3_sym) thus ?thesis using l11_10 proof - have D2: "B Out A A0" using P1 using CongA_def assms(1) bet_out by auto have D3: "B Out C C0" using P1 using CongA_def assms(1) bet_out by auto have D4: "B'' Out A'' A''" using P2 out_trivial by blast have "B'' Out C'' C''" using P2 out_trivial by auto thus ?thesis using l11_10 L11 D2 D3 D4 by blast qed qed lemma conga_pseudo_refl: assumes "A \ B" and "C \ B" shows "A B C CongA C B A" by (meson CongA_def assms(1) assms(2) between_trivial cong_pseudo_reflexivity segment_construction) lemma conga_trivial_1: assumes "A \ B" and "C \ D" shows "A B A CongA C D C" by (meson CongA_def assms(1) assms(2) cong_trivial_identity segment_construction) lemma l11_13: assumes "A B C CongA D E F" and "Bet A B A'" and "A'\ B" and "Bet D E D'" and "D'\ E" shows "A' B C CongA D' E F" proof - obtain A'' C'' D'' F'' where P1: "Bet B A A'' \ Cong A A'' E D \ Bet B C C'' \ Cong C C'' E F \ Bet E D D'' \ Cong D D'' B A \ Bet E F F'' \ Cong F F'' B C \ Cong A'' C'' D'' F''" using CongA_def assms(1) by auto obtain A0 where P2:"Bet B A' A0 \ Cong A' A0 E D'" using segment_construction by blast obtain D0 where P3: "Bet E D' D0 \ Cong D' D0 B A'" using segment_construction by blast have "Cong A0 C'' D0 F''" proof - have L1: "A'' B A0 C'' OFSC D'' E D0 F''" proof - have L2: "Bet A'' B A0" proof - have M1: "Bet A'' A B" using Bet_perm P1 by blast have M2: "Bet A B A0" using P2 assms(2) assms(3) outer_transitivity_between by blast have "A \ B" using CongA_def assms(1) by blast thus ?thesis using M1 M2 outer_transitivity_between2 by blast qed have L3: "Bet D'' E D0" using Bet_perm P1 P2 outer_transitivity_between CongA_def by (metis P3 assms(1) assms(4) assms(5)) have L4: "Cong A'' B D'' E" by (meson P1 between_symmetry cong_3421 cong_left_commutativity l2_11_b) have L5: "Cong B A0 E D0" by (meson P2 P3 between_symmetry cong_3421 cong_right_commutativity l2_11_b) have "Cong B C'' E F''" by (meson P1 between_symmetry cong_3421 cong_right_commutativity l2_11_b) thus ?thesis using P1 L2 L3 L4 L5 by (simp add: OFSC_def) qed have "A'' \ B" using CongA_def P1 assms(1) bet_neq12__neq by fastforce thus ?thesis using L1 five_segment_with_def by auto qed thus ?thesis using CongA_def P1 P2 P3 assms(1) assms(3) assms(5) by auto qed lemma conga_right_comm: assumes "A B C CongA D E F" shows "A B C CongA F E D" by (metis Tarski_neutral_dimensionless.conga_diff45 Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms assms conga_diff56 conga_pseudo_refl) lemma conga_left_comm: assumes "A B C CongA D E F" shows "C B A CongA D E F" by (meson assms conga_right_comm conga_sym) lemma conga_comm: assumes "A B C CongA D E F" shows "C B A CongA F E D" by (meson Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms assms) lemma conga_line: assumes "A \ B" and "B \ C" and "A' \ B'" and "B' \ C'" and "Bet A B C" and "Bet A' B' C'" shows "A B C CongA A' B' C'" by (metis Bet_cases assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) conga_trivial_1 l11_13) lemma l11_14: assumes "Bet A B A'" and "A \ B" and "A' \ B" and "Bet C B C'" and "B \ C" and "B \ C'" shows "A B C CongA A' B C'" by (metis Bet_perm assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) conga_pseudo_refl conga_right_comm l11_13) lemma l11_16: assumes "Per A B C" and "A \ B" and "C \ B" and "Per A' B' C'" and "A'\ B'" and "C'\ B'" shows "A B C CongA A' B' C'" proof - obtain C0 where P1: "Bet B C C0 \ Cong C C0 B' C'" using segment_construction by blast obtain C1 where P2: "Bet B' C' C1 \ Cong C' C1 B C" using segment_construction by blast obtain A0 where P3: "Bet B A A0 \ Cong A A0 B' A'" using segment_construction by blast obtain A1 where P4: "Bet B' A' A1 \ Cong A' A1 B A" using segment_construction by blast have "Cong A0 C0 A1 C1" proof - have Q1: "Per A0 B C0" by (metis P1 P3 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) bet_col per_col) have Q2: "Per A1 B' C1" by (metis P2 P4 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(4) assms(5) assms(6) bet_col per_col) have Q3: "Cong A0 B A1 B'" by (meson P3 P4 between_symmetry cong_3421 cong_left_commutativity l2_11_b) have "Cong B C0 B' C1" using P1 P2 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast thus ?thesis using Q1 Q2 Q3 l10_12 by blast qed thus ?thesis using CongA_def P1 P2 P3 P4 assms(2) assms(3) assms(5) assms(6) by auto qed lemma l11_17: assumes "Per A B C" and "A B C CongA A' B' C'" shows "Per A' B' C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "Per A0 B C0" proof - have Q1: "B \ C" using assms(2) conga_diff2 by blast have Q2: "Per A0 B C" by (metis P1 Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_col conga_diff1 per_col) have "Col B C C0" using P1 bet_col by blast thus ?thesis using Q1 Q2 per_col by blast qed have P3: "Per A1 B' C1" proof - have "A0 B C0 Cong3 A1 B' C1" by (meson Bet_cases Cong3_def P1 l2_11_b not_cong_2134 not_cong_3421) thus ?thesis using P2 l8_10 by blast qed have P4: "B' \ C1" using P1 assms(2) between_identity conga_diff56 by blast have P5: "Per A' B' C1" proof - have P6: "B' \ A1" using P1 assms(2) between_identity conga_diff45 by blast have P7: "Per C1 B' A1" by (simp add: P3 l8_2) have "Col B' A1 A'" using P1 NCol_cases bet_col by blast thus ?thesis using P3 P6 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms by fastforce qed have "Col B' C1 C'" using P1 bet_col col_permutation_5 by blast thus ?thesis using P4 P5 per_col by blast qed lemma l11_18_1: assumes "Bet C B D" and "B \ C" and "B \ D" and "A \ B" and "Per A B C" shows "A B C CongA A B D" by (smt Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) bet_col col_per2__per l11_16) lemma l11_18_2: assumes "Bet C B D" and "A B C CongA A B D" shows "Per A B C" proof - obtain A0 C0 A1 D0 where P1: "Bet B A A0 \ Cong A A0 B A \ Bet B C C0 \ Cong C C0 B D \ Bet B A A1 \ Cong A A1 B A \ Bet B D D0 \ Cong D D0 B C \ Cong A0 C0 A1 D0" using CongA_def assms(2) by auto have P2: "A0 = A1" by (metis P1 assms(2) conga_diff45 construction_uniqueness) have P3: "Per A0 B C0" proof - have Q1: "Bet C0 B D0" proof - have R1: "Bet C0 C B" using P1 between_symmetry by blast have R2: "Bet C B D0" proof - have S1: "Bet C B D" by (simp add: assms(1)) have S2: "Bet B D D0" using P1 by blast have "B \ D" using assms(2) conga_diff56 by blast thus ?thesis using S1 S2 outer_transitivity_between by blast qed have "C \ B" using assms(2) conga_diff2 by auto thus ?thesis using R1 R2 outer_transitivity_between2 by blast qed have Q2: "Cong C0 B B D0" by (meson P1 between_symmetry cong_3421 l2_11_b not_cong_1243) have "Cong A0 C0 A0 D0" using P1 P2 by blast thus ?thesis using Per_def Q1 Q2 midpoint_def by blast qed have P4: "B \ C0" using P1 assms(2) bet_neq12__neq conga_diff2 by blast have P5: "Per A B C0" by (metis P1 P3 Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms assms(2) bet_col bet_col1 bet_neq21__neq col_transitivity_1 conga_diff45) have "Col B C0 C" using P1 using NCol_cases bet_col by blast thus ?thesis using P4 P5 per_col by blast qed lemma cong3_preserves_out: assumes "A Out B C" and "A B C Cong3 A' B' C'" shows "A' Out B' C'" by (meson assms(1) assms(2) col_permutation_4 cong3_symmetry cong_3_swap l4_13 l4_6 not_bet_and_out or_bet_out out_col) lemma l11_21_a: assumes "B Out A C" and "A B C CongA A' B' C'" shows "B' Out A' C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B' A' \ Bet B C C0 \ Cong C C0 B' C' \ Bet B' A' A1 \ Cong A' A1 B A \ Bet B' C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "B Out A0 C0" by (metis P1 assms(1) bet_out l6_6 l6_7 out_diff1) have P3: "B' Out A1 C1" proof - have "B A0 C0 Cong3 B' A1 C1" by (meson Cong3_def P1 between_symmetry cong_right_commutativity l2_11_b not_cong_4312) thus ?thesis using P2 cong3_preserves_out by blast qed thus ?thesis by (metis P1 assms(2) bet_out conga_diff45 conga_diff56 l6_6 l6_7) qed lemma l11_21_b: assumes "B Out A C" and "B' Out A' C'" shows "A B C CongA A' B' C'" by (smt assms(1) assms(2) between_trivial2 conga_trivial_1 l11_10 out2_bet_out out_distinct) lemma conga_cop__or_out_ts: assumes "Coplanar A B C C'" and "A B C CongA A B C'" shows "B Out C C' \ A B TS C C'" proof - obtain A0 C0 A1 C1 where P1: "Bet B A A0 \ Cong A A0 B A \Bet B C C0 \ Cong C C0 B C' \Bet B A A1 \ Cong A A1 B A \Bet B C' C1 \ Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have P2: "A0 = A1" using P1 by (metis assms(2) conga_diff1 construction_uniqueness) have "B Out C C' \ A B TS C C'" proof cases assume "C0 = C1" thus ?thesis by (metis P1 assms(2) bet2__out conga_diff2 conga_diff56) next assume R1: "C0 \ C1" obtain M where R2: "M Midpoint C0 C1" using midpoint_existence by blast have R3: "Cong B C0 B C1" by (meson Bet_cases P1 l2_11_b not_cong_2134 not_cong_3421) have R3A: "Cong A0 C0 A0 C1" using P1 P2 by blast then have R4: "Per A0 M C0" using R2 using Per_def by blast have R5: "Per B M C0" using Per_def R2 R3 by auto then have R6: "Per B M C1" using R2 l8_4 by blast have R7: "B \ A0" using P1 assms(2) bet_neq12__neq conga_diff45 by blast then have "Cong A C0 A C1" by (meson Col_perm P1 R3 R3A bet_col l4_17) then have R9: "Per A M C0" using Per_def R2 by blast then have R10: "Per A M C1" by (meson R2 Tarski_neutral_dimensionless.l8_4 Tarski_neutral_dimensionless_axioms) have R11: "Col B A M" proof - have S1: "Coplanar C0 B A M" proof - have "Coplanar B A C0 M" proof - have T1: "Coplanar B A C0 C1" proof - have "Coplanar A C0 B C'" proof - have "Coplanar A C' B C0" proof - have U1: "Coplanar A C' B C" by (simp add: assms(1) coplanar_perm_4) have U2: "B \ C" using assms(2) conga_diff2 by blast have "Col B C C0" by (simp add: P1 bet_col) thus ?thesis by (meson Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms U1 U2) qed thus ?thesis using ncoplanar_perm_5 by blast qed thus ?thesis by (metis P1 Tarski_neutral_dimensionless.col_cop__cop Tarski_neutral_dimensionless_axioms assms(2) bet_col conga_diff56 coplanar_perm_12) qed have "Col C0 C1 M" using Col_perm R2 midpoint_col by blast thus ?thesis using T1 R1 col_cop__cop by blast qed thus ?thesis using ncoplanar_perm_8 by blast qed have "C0 \ M" using R1 R2 midpoint_distinct_1 by blast thus ?thesis using R5 R9 S1 cop_per2__col by blast qed have "B Out C C' \ A B TS C C'" proof cases assume Q1: "B = M" have Q2: "\ Col A B C" by (metis Col_def P1 Q1 R9 assms(2) conga_diff1 conga_diff2 l6_16_1 l8_9 not_bet_and_out out_trivial) have Q3: "\ Col A B C'" by (metis Col_def P1 Q1 R10 assms(2) conga_diff1 conga_diff56 l6_16_1 l8_9 not_bet_and_out out_trivial) have Q4: "Col B A B" by (simp add: col_trivial_3) have "Bet C B C'" proof - have S1: "Bet C1 C' B" using Bet_cases P1 by blast have "Bet C1 B C" proof - have T1: "Bet C0 C B" using Bet_cases P1 by blast have "Bet C0 B C1" by (simp add: Q1 R2 midpoint_bet) thus ?thesis using T1 between_exchange3 between_symmetry by blast qed thus ?thesis using S1 between_exchange3 between_symmetry by blast qed thus ?thesis by (metis Q2 Q3 Q4 bet__ts col_permutation_4 invert_two_sides) next assume L1: "B \ M" have L2: "B M TS C0 C1" proof - have M1: "\ Col C0 B M" by (metis (no_types) Col_perm L1 R1 R2 R5 is_midpoint_id l8_9) have M2: "\ Col C1 B M" using Col_perm L1 R1 R2 R6 l8_9 midpoint_not_midpoint by blast have M3: "Col M B M" using col_trivial_3 by auto have "Bet C0 M C1" by (simp add: R2 midpoint_bet) thus ?thesis using M1 M2 M3 TS_def by blast qed have "A B TS C C'" proof - have W2: "A B TS C C1" proof - have V1: "A B TS C0 C1" using L2 P1 R11 R7 col_two_sides cong_diff invert_two_sides not_col_permutation_5 by blast have "B Out C0 C" using L2 Out_def P1 TS_def assms(2) col_trivial_1 conga_diff2 by auto thus ?thesis using V1 col_trivial_3 l9_5 by blast qed then have W1: "A B TS C' C" proof - have Z1: "A B TS C1 C" by (simp add: W2 l9_2) have Z2: "Col B A B" using not_col_distincts by blast have "B Out C1 C'" using L2 Out_def P1 TS_def assms(2) col_trivial_1 conga_diff56 by auto thus ?thesis using Z1 Z2 l9_5 by blast qed thus ?thesis by (simp add: l9_2) qed thus ?thesis by blast qed thus ?thesis by blast qed thus ?thesis by blast qed lemma conga_os__out: assumes "A B C CongA A B C'" and "A B OS C C'" shows "B Out C C'" using assms(1) assms(2) conga_cop__or_out_ts l9_9 os__coplanar by blast lemma cong2_conga_cong: assumes "A B C CongA A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "Cong A C A' C'" by (smt assms(1) assms(2) assms(3) cong_4321 l11_3 l11_4_1 not_cong_3412 out_distinct out_trivial) lemma angle_construction_1: assumes "\ Col A B C" and "\ Col A' B' P" shows "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof - obtain C0 where P1: "Col B A C0 \ B A Perp C C0" using assms(1) col_permutation_4 l8_18_existence by blast have "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof cases assume P1A: "B = C0" obtain C' where P2: "Per C' B' A' \ Cong C' B' C B \ A' B' OS C' P" by (metis assms(1) assms(2) col_trivial_1 col_trivial_2 ex_per_cong) have P3: "A B C CongA A' B' C'" by (metis P1 P2 Tarski_neutral_dimensionless.l8_2 Tarski_neutral_dimensionless.os_distincts Tarski_neutral_dimensionless_axioms P1A assms(1) l11_16 not_col_distincts perp_per_1) thus ?thesis using P2 by blast next assume P4: "B \ C0" have "\ C'. (A B C CongA A' B' C' \ A' B' OS C' P)" proof cases assume R1: "B Out A C0" obtain C0' where R2: "B' Out A' C0' \ Cong B' C0' B C0" by (metis P4 assms(2) col_trivial_1 segment_construction_3) have "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" proof - have R4: "B' \ C0'" using Out_def R2 by auto have R5: "C \ C0" using P1 perp_distinct by blast have R6: "Col B' C0' C0'" by (simp add: col_trivial_2) have "\ Col B' C0' P" using NCol_cases R2 R4 assms(2) col_transitivity_1 out_col by blast then have "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" using R4 R5 R6 ex_per_cong by blast thus ?thesis by auto qed then obtain C' where R7: "Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" by auto then have R8: "C0 B C Cong3 C0' B' C'" by (meson Cong3_def P1 R2 col_trivial_2 l10_12 l8_16_1 not_col_permutation_2 not_cong_2143 not_cong_4321) have R9: "A B C CongA A' B' C'" proof - have S1: "C0 B C CongA C0' B' C'" by (metis P4 R8 assms(1) cong3_conga not_col_distincts) have S3: "B Out C C" using assms(1) not_col_distincts out_trivial by force have "B' \ C'" using R8 assms(1) cong3_diff2 not_col_distincts by blast then have "B' Out C' C'" using out_trivial by auto thus ?thesis using S1 R1 S3 R2 l11_10 by blast qed have "B' A' OS C' P" proof - have T1: "Col B' C0' A'" by (meson NCol_cases R2 Tarski_neutral_dimensionless.out_col Tarski_neutral_dimensionless_axioms) have "B' \ A'" using assms(2) col_trivial_1 by auto thus ?thesis using T1 R7 col_one_side by blast qed then have "A' B' OS C' P" by (simp add: invert_one_side) thus ?thesis using R9 by blast next assume U1: "\ B Out A C0" then have U2: "Bet A B C0" using NCol_perm P1 or_bet_out by blast obtain C0' where U3: "Bet A' B' C0' \ Cong B' C0' B C0" using segment_construction by blast have U4: "\ C'. Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" proof - have V2: "C \ C0" using Col_cases P1 assms(1) by blast have "B' \ C0'" using P4 U3 cong_diff_3 by blast then have "\ Col B' C0' P" using Col_def U3 assms(2) col_transitivity_1 by blast thus ?thesis using ex_per_cong using V2 not_col_distincts by blast qed then obtain C' where U5: "Per C' C0' B' \ Cong C' C0' C C0 \ B' C0' OS C' P" by blast have U98: "A B C CongA A' B' C'" proof - have X1: "C0 B C Cong3 C0' B' C'" proof - have X2: "Cong C0 B C0' B'" using Cong_cases U3 by blast have X3: "Cong C0 C C0' C'" using U5 not_cong_4321 by blast have "Cong B C B' C'" proof - have Y1: "Per C C0 B" using P1 col_trivial_3 l8_16_1 by blast have "Cong C C0 C' C0'" using U5 not_cong_3412 by blast thus ?thesis using Cong_perm Y1 U5 X2 l10_12 by blast qed thus ?thesis by (simp add: Cong3_def X2 X3) qed have X22: "Bet C0 B A" using U2 between_symmetry by blast have X24: "Bet C0' B' A'" using Bet_cases U3 by blast have "A' \ B'" using assms(2) not_col_distincts by blast thus ?thesis by (metis P4 X1 X22 X24 assms(1) cong3_conga l11_13 not_col_distincts) qed have "A' B' OS C' P" proof - have "B' A' OS C' P" proof - have W1: "Col B' C0' A'" by (simp add: Col_def U3) have "B' \ A'" using assms(2) not_col_distincts by blast thus ?thesis using W1 U5 col_one_side by blast qed thus ?thesis by (simp add: invert_one_side) qed thus ?thesis using U98 by blast qed thus ?thesis by auto qed thus ?thesis by auto qed lemma angle_construction_2: assumes "A \ B" (*and "A \ C"*) and "B \ C" (*and "A' \ B'"*) and "\ Col A' B' P" shows "\ C'. (A B C CongA A' B' C' \ (A' B' OS C' P \ Col A' B' C'))" by (metis Col_def angle_construction_1 assms(1) assms(2) assms(3) col_trivial_3 conga_line l11_21_b or_bet_out out_trivial point_construction_different) lemma ex_conga_ts: assumes "\ Col A B C" and "\ Col A' B' P" shows "\ C'. A B C CongA A' B' C' \ A' B' TS C' P" proof - obtain P' where P1: "A' Midpoint P P'" using symmetric_point_construction by blast have P2: "\ Col A' B' P'" by (metis P1 assms(2) col_transitivity_1 midpoint_col midpoint_distinct_2 not_col_distincts) obtain C' where P3: "A B C CongA A' B' C' \ A' B' OS C' P'" using P2 angle_construction_1 assms(1) by blast have "A' B' TS P' P" using P1 P2 assms(2) bet__ts l9_2 midpoint_bet not_col_distincts by auto thus ?thesis using P3 l9_8_2 one_side_symmetry by blast qed lemma l11_15: assumes "\ Col A B C" and "\ Col D E P" shows "\ F. (A B C CongA D E F \ E D OS F P) \ (\ F1 F2. ((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P)) \ E Out F1 F2)" proof - obtain F where P1: "A B C CongA D E F \ D E OS F P" using angle_construction_1 assms(1) assms(2) by blast then have P2: "A B C CongA D E F \ E D OS F P" using invert_one_side by blast have "(\ F1 F2. ((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P)) \ E Out F1 F2)" proof - { fix F1 F2 assume P3: "((A B C CongA D E F1 \ E D OS F1 P) \ (A B C CongA D E F2 \ E D OS F2 P))" then have P4: "A B C CongA D E F1" by simp have P5: "E D OS F1 P" using P3 by simp have P6: "A B C CongA D E F2" using P3 by simp have P7: "E D OS F2 P" using P3 by simp have P8: "D E F1 CongA D E F2" using P4 conga_sym P6 conga_trans by blast have "D E OS F1 F2" using P5 P7 invert_one_side one_side_symmetry one_side_transitivity by blast then have "E Out F1 F2" using P8 conga_os__out by (meson P3 conga_sym conga_trans) } thus ?thesis by auto qed thus ?thesis using P2 by blast qed lemma l11_19: assumes "Per A B P1" and "Per A B P2" and "A B OS P1 P2" shows "B Out P1 P2" proof cases assume "Col A B P1" thus ?thesis using assms(3) col123__nos by blast next assume P1: "\ Col A B P1" have "B Out P1 P2" proof cases assume "Col A B P2" thus ?thesis using assms(3) one_side_not_col124 by blast next assume P2: "\ Col A B P2" obtain x where "A B P1 CongA A B x \ B A OS x P2 \ (\ F1 F2. ((A B P1 CongA A B F1 \ B A OS F1 P2) \ (A B P1 CongA A B F2 \ B A OS F2 P2))\ B Out F1 F2)" using P1 P2 l11_15 by blast thus ?thesis by (metis P1 P2 assms(1) assms(2) assms(3) conga_os__out l11_16 not_col_distincts) qed thus ?thesis by simp qed lemma l11_22_bet: assumes "Bet A B C" and "P' B' TS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "Bet A' B' C'" proof - obtain C'' where P1: "Bet A' B' C'' \ Cong B' C'' B C" using segment_construction by blast have P2: "C B P CongA C'' B' P'" by (metis P1 assms(1) assms(3) assms(4) cong_diff_3 conga_diff2 l11_13) have P3: "C'' B' P' CongA C' B' P'" by (meson P2 Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless_axioms assms(4) conga_comm conga_trans) have P4: "B' Out C' C'' \ P' B' TS C' C''" proof - have P5: "Coplanar P' B' C' C''" by (meson P1 TS_def assms(2) bet__coplanar coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_8 ts__coplanar) have "P' B' C' CongA P' B' C''" using P3 conga_comm conga_sym by blast thus ?thesis by (simp add: P5 conga_cop__or_out_ts) qed have P6: "B' Out C' C'' \ Bet A' B' C'" proof - { assume "B' Out C' C''" then have "Bet A' B' C'" using P1 bet_out_out_bet between_exchange4 between_trivial2 col_trivial_3 l6_6 not_bet_out by blast } thus ?thesis by simp qed have "P' B' TS C' C'' \ Bet A' B' C'" proof - { assume P7: "P' B' TS C' C''" then have "Bet A' B' C'" proof cases assume "Col C' B' P'" thus ?thesis using Col_perm TS_def assms(2) by blast next assume Q1: "\ Col C' B' P'" then have Q2: "B' \ P'" using not_col_distincts by blast have Q3: "B' P' TS A' C''" by (metis Col_perm P1 TS_def P7 assms(2) col_trivial_3) have Q4: "B' P' OS C' C''" using P7 Q3 assms(2) invert_two_sides l9_8_1 l9_9 by blast thus ?thesis using P7 invert_one_side l9_9 by blast qed } thus ?thesis by simp qed thus ?thesis using P6 P4 by blast qed lemma l11_22a: assumes "B P TS A C" and "B' P' TS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" proof - have P1: "A \ B \ A' \ B' \ P \ B \ P' \ B' \ C \ B \ C' \ B'" using assms(3) assms(4) conga_diff1 conga_diff2 conga_diff45 conga_diff56 by auto have P2: "A \ C \ A' \ C'" using assms(1) assms(2) not_two_sides_id by blast obtain A'' where P3: "B' Out A' A'' \ Cong B' A'' B A" using P1 segment_construction_3 by force have P4: "\ Col A B P" using TS_def assms(1) by blast obtain T where P5: "Col T B P \ Bet A T C" using TS_def assms(1) by blast have "A B C CongA A' B' C'" proof cases assume "B = T" thus ?thesis by (metis P1 P5 assms(2) assms(3) assms(4) conga_line invert_two_sides l11_22_bet) next assume P6: "B \ T" have "A B C CongA A' B' C'" proof cases assume P7A: "Bet P B T" obtain T'' where T1: "Bet P' B' T'' \ Cong B' T'' B T" using segment_construction by blast have "\ T''. Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" proof - have T2: "Col B' P' T''" using T1 by (simp add: bet_col col_permutation_4) have "(B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" using P7A T1 not_bet_and_out by blast thus ?thesis using T2 by blast qed then obtain T'' where T3: "Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" by blast then have T4: "B' \ T''" using P6 cong_diff_3 by blast obtain C'' where T5: "Bet A'' T'' C'' \ Cong T'' C'' T C" using segment_construction by blast have T6: "A B T CongA A' B' T''" by (smt Out_cases P5 P6 T3 T4 P7A assms(3) between_symmetry col_permutation_4 conga_comm l11_13 l6_4_1 or_bet_out) then have T7: "A B T CongA A'' B' T''" by (smt P3 P4 P6 T3 Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms bet_out col_trivial_3 cong_diff_3 l5_2 l6_6 not_col_permutation_1 or_bet_out) then have T8: "Cong A T A'' T''" using P3 T3 cong2_conga_cong cong_4321 not_cong_3412 by blast have T9: "Cong A C A'' C''" using P5 T5 T8 cong_symmetry l2_11_b by blast have T10: "Cong C B C'' B'" by (smt P3 P4 P5 T3 T5 T8 cong_commutativity cong_symmetry five_segment) have "A B C Cong3 A'' B' C''" using Cong3_def P3 T10 T9 not_cong_2143 not_cong_4321 by blast then have T11: "A B C CongA A'' B' C''" by (simp add: Tarski_neutral_dimensionless.cong3_conga Tarski_neutral_dimensionless_axioms P1) have "C B T Cong3 C'' B' T''" by (simp add: Cong3_def T10 T3 T5 cong_4321 cong_symmetry) then have T12: "C B T CongA C'' B' T''" using P1 P6 cong3_conga by auto have T13: "P B C CongA P' B' C''" proof - have K4: "Bet T B P" using Bet_perm P7A by blast have "Bet T'' B' P'" using Col_perm P7A T3 l6_6 not_bet_and_out or_bet_out by blast thus ?thesis using K4 P1 T12 conga_comm l11_13 by blast qed have T14: "P' B' C' CongA P' B' C''" proof - have "P' B' C' CongA P B C" by (simp add: assms(4) conga_sym) thus ?thesis using T13 conga_trans by blast qed have T15: "B' Out C' C'' \ P' B' TS C' C''" proof - have K7: "Coplanar P' B' C' C''" proof - have K8: "Coplanar A' P' B' C'" using assms(2) coplanar_perm_14 ts__coplanar by blast have K8A: "Coplanar P' C'' B' A''" proof - have "Col P' B' T'' \ Col C'' A'' T''" using Col_def Col_perm T3 T5 by blast then have "Col P' C'' T'' \ Col B' A'' T'' \ Col P' B' T'' \ Col C'' A'' T'' \ Col P' A'' T'' \ Col C'' B' T''" by simp thus ?thesis using Coplanar_def by auto qed then have "Coplanar A' P' B' C''" proof - have K9: "B' \ A''" using P3 out_distinct by blast have "Col B' A'' A'" using Col_perm P3 out_col by blast thus ?thesis using K8A K9 col_cop__cop coplanar_perm_19 by blast qed thus ?thesis by (meson K8 TS_def assms(2) coplanar_perm_7 coplanar_trans_1 ncoplanar_perm_2) qed thus ?thesis by (simp add: T14 K7 conga_cop__or_out_ts) qed have "A B C CongA A' B' C'" proof cases assume "B' Out C' C''" thus ?thesis using P1 P3 T11 l11_10 out_trivial by blast next assume W1: "\ B' Out C' C''" then have W1A: " P' B' TS C' C''" using T15 by simp have W2: "B' P' TS A'' C'" using P3 assms(2) col_trivial_1 l9_5 by blast then have W3: "B' P' OS A'' C''" using T15 W1 invert_two_sides l9_2 l9_8_1 by blast have W4: "B' P' TS A'' C''" proof - have "\ Col A' B' P'" using TS_def assms(2) by auto thus ?thesis using Col_perm T3 T5 W3 one_side_chara by blast qed thus ?thesis using W1A W2 invert_two_sides l9_8_1 l9_9 by blast qed thus ?thesis by simp next assume R1: "\ Bet P B T" then have R2: "B Out P T" using Col_cases P5 l6_4_2 by blast have R2A: "\ T''. Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" proof - obtain T'' where R3: "B' Out P' T'' \ Cong B' T'' B T" using P1 P6 segment_construction_3 by fastforce thus ?thesis using R2 out_col by blast qed then obtain T'' where R4: "Col B' P' T'' \ (B' Out P' T'' \ B Out P T) \ Cong B' T'' B T" by auto have R5: "B' \ T''" using P6 R4 cong_diff_3 by blast obtain C'' where R6: "Bet A'' T'' C'' \ Cong T'' C'' T C" using segment_construction by blast have R7: "A B T CongA A' B' T''" using P1 R2 R4 assms(3) l11_10 l6_6 out_trivial by auto have R8: "A B T CongA A'' B' T''" using P3 P4 R2 R4 assms(3) l11_10 l6_6 not_col_distincts out_trivial by blast have R9: "Cong A T A'' T''" using Cong_cases P3 R4 R8 cong2_conga_cong by blast have R10: "Cong A C A'' C''" using P5 R6 R9 l2_11_b not_cong_3412 by blast have R11: "Cong C B C'' B'" by (smt P3 P4 P5 R4 R6 R9 cong_commutativity cong_symmetry five_segment) have "A B C Cong3 A'' B' C''" by (simp add: Cong3_def P3 R10 R11 cong_4321 cong_commutativity) then have R12: "A B C CongA A'' B' C''" using P1 by (simp add: cong3_conga) have "C B T Cong3 C'' B' T''" using Cong3_def R11 R4 R6 not_cong_3412 not_cong_4321 by blast then have R13: "C B T CongA C'' B' T''" using P1 P6 Tarski_neutral_dimensionless.cong3_conga Tarski_neutral_dimensionless_axioms by fastforce have R13A: "\ Col A' B' P'" using TS_def assms(2) by blast have R14: "B' Out C' C'' \ P' B' TS C' C''" proof - have S1: "Coplanar P' B' C' C''" proof - have T1: "Coplanar A' P' B' C'" using assms(2) ncoplanar_perm_14 ts__coplanar by blast have "Coplanar A' P' B' C''" proof - have U6: "B' \ A''" using P3 out_diff2 by blast have "Coplanar P' C'' B' A''" proof - have "Col P' B' T'' \ Col C'' A'' T''" using Col_def Col_perm R4 R6 by blast thus ?thesis using Coplanar_def by auto qed thus ?thesis by (meson Col_cases P3 U6 col_cop__cop ncoplanar_perm_21 ncoplanar_perm_6 out_col) qed thus ?thesis using NCol_cases R13A T1 coplanar_trans_1 by blast qed have "P' B' C' CongA P' B' C''" proof - have "C B P CongA C'' B' P'" using P1 R12 R13 R2 R4 conga_diff56 l11_10 out_trivial by presburger then have "C' B' P' CongA C'' B' P'" by (meson Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms assms(4) conga_comm conga_sym) thus ?thesis by (simp add: conga_comm) qed thus ?thesis by (simp add: S1 conga_cop__or_out_ts) qed have S1: "B Out A A" using P4 not_col_distincts out_trivial by blast have S2: "B Out C C" using TS_def assms(1) not_col_distincts out_trivial by auto have S3: "B' Out A' A''" using P3 by simp have "A B C CongA A' B' C'" proof cases assume "B' Out C' C''" thus ?thesis using S1 S2 S3 using R12 l11_10 by blast next assume "\ B' Out C' C''" then have Z3: "P' B' TS C' C''" using R14 by simp have Q1: "B' P' TS A'' C'" using S3 assms(2) l9_5 not_col_distincts by blast have Q2: "B' P' OS A'' C''" proof - have "B' P' TS C'' C'" proof - have "B' P' TS C' C''" using Z3 using invert_two_sides by blast thus ?thesis by (simp add: l9_2) qed thus ?thesis using Q1 l9_8_1 by blast qed have "B' P' TS A'' C''" using Col_perm Q2 R4 R6 one_side_chara by blast thus ?thesis using Q2 l9_9 by blast qed thus ?thesis using S1 S2 S3 using R12 l11_10 by blast qed thus ?thesis by simp qed thus ?thesis by simp qed lemma l11_22b: assumes "B P OS A C" and "B' P' OS A' C'" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" proof - obtain D where P1: "Bet A B D \ Cong B D A B" using segment_construction by blast obtain D' where P2: "Bet A' B' D' \ Cong B' D' A' B'" using segment_construction by blast have P3: "D B P CongA D' B' P'" proof - have Q3: "D \ B" by (metis P1 assms(1) col_trivial_3 cong_diff_3 one_side_not_col124 one_side_symmetry) have Q5: "D' \ B'" by (metis P2 assms(2) col_trivial_3 cong_diff_3 one_side_not_col124 one_side_symmetry) thus ?thesis using assms(3) P1 Q3 P2 l11_13 by blast qed have P5: "D B C CongA D' B' C'" proof - have Q1: "B P TS D C" by (metis P1 assms(1) bet__ts col_trivial_3 cong_diff_3 l9_2 l9_8_2 one_side_not_col124 one_side_symmetry) have "B' P' TS D' C'" by (metis Cong_perm P2 assms(2) bet__ts between_cong between_trivial2 l9_2 l9_8_2 one_side_not_col123 point_construction_different ts_distincts) thus ?thesis using assms(4) Q1 P3 l11_22a by blast qed have P6: "Bet D B A" using Bet_perm P1 by blast have P7: "A \ B" using assms(3) conga_diff1 by auto have P8: "Bet D' B' A'" using Bet_cases P2 by blast have "A' \ B'" using assms(3) conga_diff45 by blast thus ?thesis using P5 P6 P7 P8 l11_13 by blast qed lemma l11_22: assumes "((B P TS A C \ B' P' TS A' C')\(B P OS A C \ B' P' OS A' C'))" and "A B P CongA A' B' P'" and "P B C CongA P' B' C'" shows "A B C CongA A' B' C'" by (meson assms(1) assms(2) assms(3) l11_22a l11_22b) lemma l11_24: assumes "P InAngle A B C" shows "P InAngle C B A" proof - obtain pp :: "'p \ 'p \ 'p \ 'p \ 'p" where "\x0 x1 x2 x3. (\v4. Bet x2 v4 x0 \ (v4 = x1 \ x1 Out v4 x3)) = (Bet x2 (pp x0 x1 x2 x3) x0 \ ((pp x0 x1 x2 x3) = x1 \ x1 Out (pp x0 x1 x2 x3) x3))" by moura then have "A \ B \ C \ B \ P \ B \Bet A (pp C B A P) C \ ((pp C B A P) = B \ B Out (pp C B A P) P)" using InAngle_def assms by presburger thus ?thesis by (metis (no_types) InAngle_def between_symmetry) qed lemma col_in_angle: assumes "A \ B" and "C \ B" and "P \ B" and "B Out A P \ B Out C P" shows "P InAngle A B C" by (meson InAngle_def assms(1) assms(2) assms(3) assms(4) between_trivial between_trivial2) lemma out321__inangle: assumes "C \ B" and "B Out A P" shows "P InAngle A B C" using assms(1) assms(2) col_in_angle out_distinct by auto lemma inangle1123: assumes "A \ B" and "C \ B" shows "A InAngle A B C" by (simp add: assms(1) assms(2) out321__inangle out_trivial) lemma out341__inangle: assumes "A \ B" and "B Out C P" shows "P InAngle A B C" using assms(1) assms(2) col_in_angle out_distinct by auto lemma inangle3123: assumes "A \ B" and "C \ B" shows "C InAngle A B C" by (simp add: assms(1) assms(2) inangle1123 l11_24) lemma in_angle_two_sides: assumes "\ Col B A P" and "\ Col B C P" and "P InAngle A B C" shows "P B TS A C" by (metis InAngle_def TS_def assms(1) assms(2) assms(3) not_col_distincts not_col_permutation_1 out_col) lemma in_angle_out: assumes "B Out A C" and "P InAngle A B C" shows "B Out A P" by (metis InAngle_def assms(1) assms(2) not_bet_and_out out2_bet_out) lemma col_in_angle_out: assumes "Col B A P" and "\ Bet A B C" and "P InAngle A B C" shows "B Out A P" proof - obtain X where P1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(3) by auto have "B Out A P" proof cases assume "X = B" thus ?thesis using P1 assms(2) by blast next assume P2: "X \ B" thus ?thesis proof - have f1: "Bet B A P \ A Out B P" by (meson assms(1) l6_4_2) have f2: "B Out X P" using P1 P2 by blast have f3: "(Bet B P A \Bet B A P) \Bet P B A" using f1 by (meson Bet_perm Out_def) have f4: "Bet B P X \Bet P X B" using f2 by (meson Bet_perm Out_def) then have f5: "((Bet B P X \Bet X B A) \Bet B P A) \Bet B A P" using f3 by (meson between_exchange3) have "\p. Bet p X C \ \Bet A p X" using P1 between_exchange3 by blast then have f6: "(P = B \Bet B A P) \Bet B P A" using f5 f3 by (meson Bet_perm P2 assms(2) outer_transitivity_between2) have f7: "Bet C X A" using Bet_perm P1 by blast have "P \ B" using f2 by (simp add: Out_def) moreover { assume "Bet B B C" then have "A \ B" using assms(2) by blast } ultimately have "A \ B" using f7 f4 f1 by (meson Bet_perm Out_def P2 between_exchange3 outer_transitivity_between2) thus ?thesis using f6 f2 by (simp add: Out_def) qed qed thus ?thesis by blast qed lemma l11_25_aux: assumes "P InAngle A B C" and "\ Bet A B C" and "B Out A' A" shows "P InAngle A' B C" proof - have P1: "Bet B A' A \ Bet B A A'" using Out_def assms(3) by auto { assume P2: "Bet B A' A" obtain X where P3: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(1) by auto obtain T where P4: "Bet A' T C \ Bet X T B" using Bet_perm P2 P3 inner_pasch by blast { assume "X = B" then have "P InAngle A' B C" using P3 assms(2) by blast } { assume "B Out X P" then have "P InAngle A' B C" by (metis InAngle_def P4 assms(1) assms(3) bet_out_1 l6_7 out_diff1) } then have "P InAngle A' B C" using P3 \X = B \ P InAngle A' B C\ by blast } { assume Q0: "Bet B A A'" obtain X where Q1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(1) by auto { assume "X = B" then have "P InAngle A' B C" using Q1 assms(2) by blast } { assume Q2: "B Out X P" obtain T where Q3: "Bet A' T C \ Bet B X T" using Bet_perm Q1 Q0 outer_pasch by blast then have "P InAngle A' B C" by (metis InAngle_def Q0 Q2 assms(1) bet_out l6_6 l6_7 out_diff1) } then have "P InAngle A' B C" using \X = B \ P InAngle A' B C\ Q1 by blast } thus ?thesis using P1 \Bet B A' A \ P InAngle A' B C\ by blast qed lemma l11_25: assumes "P InAngle A B C" and "B Out A' A" and "B Out C' C" and "B Out P' P" shows "P' InAngle A' B C'" proof cases assume "Bet A B C" thus ?thesis by (metis Bet_perm InAngle_def assms(2) assms(3) assms(4) bet_out__bet l6_6 out_distinct) next assume P1: "\ Bet A B C" have P2: "P InAngle A' B C" using P1 assms(1) assms(2) l11_25_aux by blast have P3: "P InAngle A' B C'" proof - have "P InAngle C' B A'" using l11_25_aux using Bet_perm P1 P2 assms(2) assms(3) bet_out__bet l11_24 by blast thus ?thesis using l11_24 by blast qed obtain X where P4: "Bet A' X C' \ (X = B \ B Out X P)" using InAngle_def P3 by auto { assume "X = B" then have "P' InAngle A' B C'" using InAngle_def P3 P4 assms(4) out_diff1 by auto } { assume "B Out X P" then have "P' InAngle A' B C'" proof - have "\p. B Out p P' \ \ B Out p P" by (meson Out_cases assms(4) l6_7) thus ?thesis by (metis (no_types) InAngle_def P3 assms(4) out_diff1) qed } thus ?thesis using InAngle_def P4 assms(2) assms(3) assms(4) out_distinct by auto qed lemma inangle_distincts: assumes "P InAngle A B C" shows "A \ B \ C \ B \ P \ B" using InAngle_def assms by auto lemma segment_construction_0: shows "\ B'. Cong A' B' A B" using segment_construction by blast lemma angle_construction_3: assumes "A \ B" and "C \ B" and "A' \ B'" shows "\ C'. A B C CongA A' B' C'" by (metis angle_construction_2 assms(1) assms(2) assms(3) not_col_exists) lemma l11_28: assumes "A B C Cong3 A' B' C'" and "Col A C D" shows "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume P1: "A = C" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = B" thus ?thesis by (metis P1 assms(1) cong3_diff cong3_symmetry cong_3_swap_2 not_cong_3421 segment_construction_0) next assume "A \ B" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = D" thus ?thesis using Cong3_def P1 assms(1) cong_trivial_identity by blast next assume "A \ D" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "B = D" thus ?thesis using Cong3_def assms(1) cong_3_swap_2 cong_trivial_identity by blast next assume Q1: "B \ D" obtain D'' where Q2: "B A D CongA B' A' D''" by (metis \A \ B\ \A \ D\ angle_construction_3 assms(1) cong3_diff) obtain D' where Q3: "A' Out D'' D' \ Cong A' D' A D" by (metis Q2 \A \ D\ conga_diff56 segment_construction_3) have Q5: "Cong A D A' D'" using Q3 not_cong_3412 by blast have "B A D CongA B' A' D'" using Q2 Q3 \A \ B\ \A \ D\ conga_diff45 l11_10 l6_6 out_trivial by auto then have "Cong B D B' D'" using Cong3_def Cong_perm Q5 assms(1) cong2_conga_cong by blast thus ?thesis using Cong3_def P1 Q5 assms(1) cong_reverse_identity by blast qed thus ?thesis by simp qed thus ?thesis by simp qed thus ?thesis by simp next assume Z1: "A \ C" have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" proof cases assume "A = D" thus ?thesis using Cong3_def Cong_perm assms(1) cong_trivial_identity by blast next assume "A \ D" { assume "Bet A C D" obtain D' where W1: "Bet A' C' D' \ Cong C' D' C D" using segment_construction by blast have W2: "Cong A D A' D'" by (meson Cong3_def W1 \Bet A C D\ assms(1) cong_symmetry l2_11_b) have W3: "Cong B D B' D'" proof - have X1: "Cong C D C' D'" using W1 not_cong_3412 by blast have "Cong C B C' B'" using Cong3_def assms(1) cong_commutativity by presburger then have W4: "A C D B OFSC A' C' D' B'" using Cong3_def OFSC_def W1 X1 \Bet A C D\ assms(1) by blast have "Cong D B D' B'" using W4 \A \ C\ five_segment_with_def by blast thus ?thesis using Z1 not_cong_2143 by blast qed have "Cong C D C' D'" by (simp add: W1 cong_symmetry) then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W2 W3 by blast } { assume W3B: "Bet C D A" then obtain D' where W4A: "Bet A' D' C' \ A D C Cong3 A' D' C'" using Bet_perm Cong3_def assms(1) l4_5 by blast have W5: "Cong A D A' D'" using Cong3_def W4A by blast have "A D C B IFSC A' D' C' B'" by (meson Bet_perm Cong3_def Cong_perm IFSC_def W4A W3B assms(1)) then have "Cong D B D' B'" using l4_2 by blast then have W6: "Cong B D B' D'" using Cong_perm by blast then have "Cong C D C' D'" using Cong3_def W4A not_cong_2143 by blast then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W5 W6 by blast } { assume W7: "Bet D A C" obtain D' where W7A: "Bet C' A' D' \ Cong A' D' A D" using segment_construction by blast then have W8: "Cong A D A' D'" using Cong_cases by blast have "C A D B OFSC C' A' D' B'" by (meson Bet_perm Cong3_def Cong_perm OFSC_def W7 W7A assms(1)) then have "Cong D B D' B'" using Z1 five_segment_with_def by auto then have w9: "Cong B D B' D'" using Cong_perm by blast have "Cong C D C' D'" proof - have L1: "Bet C A D" using Bet_perm W7 by blast have L2: "Bet C' A' D'" using Bet_perm W7 using W7A by blast have "Cong C A C' A'" using assms(1) using Cong3_def assms(1) not_cong_2143 by blast thus ?thesis using l2_11 using L1 L2 W8 l2_11 by blast qed then have "\ D'. (Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D')" using W8 w9 by blast } thus ?thesis using Bet_cases \Bet A C D \ \D'. Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D'\ \Bet C D A \ \D'. Cong A D A' D' \ Cong B D B' D' \ Cong C D C' D'\ assms(2) third_point by blast qed thus ?thesis by blast qed lemma bet_conga__bet: assumes "Bet A B C" and "A B C CongA A' B' C'" shows "Bet A' B' C'" proof - obtain A0 C0 A1 C1 where P1:" Bet B A A0 \Cong A A0 B' A' \ Bet B C C0 \Cong C C0 B' C' \ Bet B' A' A1 \Cong A' A1 B A \ Bet B' C' C1 \Cong C' C1 B C \ Cong A0 C0 A1 C1" using CongA_def assms(2) by auto have "Bet C B A0" using P1 outer_transitivity_between by (metis assms(1) assms(2) between_symmetry conga_diff1) then have "Bet A0 B C" using Bet_cases by blast then have P2: "Bet A0 B C0" using P1 assms(2) conga_diff2 outer_transitivity_between by blast have P3: "A0 B C0 Cong3 A1 B' C1" proof - have Q1: "Cong A0 B A1 B'" by (meson Bet_cases P1 l2_11_b not_cong_1243 not_cong_4312) have Q3: "Cong B C0 B' C1" using P1 between_symmetry cong_3421 l2_11_b not_cong_1243 by blast thus ?thesis by (simp add: Cong3_def Q1 P1) qed then have P4: "Bet A1 B' C1" using P2 l4_6 by blast then have "Bet A' B' C1" using P1 Bet_cases between_exchange3 by blast thus ?thesis using between_inner_transitivity P1 by blast qed lemma in_angle_one_side: assumes "\ Col A B C" and "\ Col B A P" and "P InAngle A B C" shows "A B OS P C" proof - obtain X where P1: "Bet A X C \ (X = B \ B Out X P)" using InAngle_def assms(3) by auto { assume "X = B" then have "A B OS P C" using P1 assms(1) bet_col by blast } { assume P2: "B Out X P" obtain C' where P2A: "Bet C A C' \ Cong A C' C A" using segment_construction by blast have "A B TS X C'" proof - have Q1: "\ Col X A B" by (metis Col_def P1 assms(1) assms(2) col_transitivity_2 out_col) have Q2 :"\ Col C' A B" by (metis Col_def Cong_perm P2A assms(1) cong_diff l6_16_1) have "\ T. Col T A B \ Bet X T C'" using Bet_cases P1 P2A between_exchange3 col_trivial_1 by blast thus ?thesis by (simp add: Q1 Q2 TS_def) qed then have P3: "A B TS P C'" using P2 col_trivial_3 l9_5 by blast then have "A B TS C C'" by (smt P1 P2 bet_out bet_ts__os between_trivial col123__nos col_trivial_3 invert_two_sides l6_6 l9_2 l9_5) then have "A B OS P C" using OS_def P3 by blast } thus ?thesis using P1 \X = B \ A B OS P C\ by blast qed lemma inangle_one_side: assumes "\ Col A B C" and "\ Col A B P" and "\ Col A B Q" and "P InAngle A B C" and "Q InAngle A B C" shows "A B OS P Q" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) in_angle_one_side not_col_permutation_4 one_side_symmetry one_side_transitivity) lemma inangle_one_side2: assumes "\ Col A B C" and "\ Col A B P" and "\ Col A B Q" and "\ Col C B P" and "\ Col C B Q" and "P InAngle A B C" and "Q InAngle A B C" shows "A B OS P Q \ C B OS P Q" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) inangle_one_side l11_24 not_col_permutation_3) lemma col_conga_col: assumes "Col A B C" and "A B C CongA D E F" shows "Col D E F" proof - { assume "Bet A B C" then have "Col D E F" using Col_def assms(2) bet_conga__bet by blast } { assume "Bet B C A" then have "Col D E F" by (meson Col_perm Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \Bet A B C \ Col D E F\ assms(1) assms(2) or_bet_out out_col) } { assume "Bet C A B" then have "Col D E F" by (meson Col_perm Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \Bet A B C \ Col D E F\ assms(1) assms(2) or_bet_out out_col) } thus ?thesis using Col_def \Bet A B C \ Col D E F\ \Bet B C A \ Col D E F\ assms(1) by blast qed lemma ncol_conga_ncol: assumes "\ Col A B C" and "A B C CongA D E F" shows "\ Col D E F" using assms(1) assms(2) col_conga_col conga_sym by blast lemma angle_construction_4: assumes "A \ B" and "C \ B" and "A' \ B'" shows "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" proof cases assume "Col A' B' P" thus ?thesis using angle_construction_3 assms(1) assms(2) assms(3) ncop__ncols by blast next assume "\ Col A' B' P" { assume "Col A B C" then have "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" by (meson angle_construction_3 assms(1) assms(2) assms(3) col__coplanar col_conga_col) } { assume "\ Col A B C" then obtain C' where "A B C CongA A' B' C' \ A' B' OS C' P" using \\ Col A' B' P\ angle_construction_1 by blast then have "\C'. (A B C CongA A' B' C' \ Coplanar A' B' C' P)" using os__coplanar by blast } thus ?thesis using \Col A B C \ \C'. A B C CongA A' B' C' \ Coplanar A' B' C' P\ by blast qed lemma lea_distincts: assumes "A B C LeA D E F" shows "A\B \ C\B \ D\E \ F\E" by (metis (no_types) LeA_def Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless_axioms assms inangle_distincts) lemma l11_29_a: assumes "A B C LeA D E F" shows "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" proof - obtain P where P1: "P InAngle D E F \ A B C CongA D E P" using LeA_def assms by blast then have P2: "E \ D \ B \ A \ E \ F \ E \ P \ B \ C" using conga_diff1 conga_diff2 inangle_distincts by blast then have P3: "A \ B \ C \ B" by blast { assume Q1: "Bet A B C" then have Q2: "Bet D E P" by (meson P1 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms) have Q3: "C InAngle A B C" by (simp add: P3 inangle3123) obtain X where Q4: "Bet D X F \ (X = E \ E Out X P)" using InAngle_def P1 by auto have "A B C CongA D E F" proof - { assume R1: "X = E" have R2: "Bet E F P \ Bet E P F" proof - have R3: "D \ E" using P2 by blast have "Bet D E F" using Col_def Col_perm P1 Q2 col_in_angle_out not_bet_and_out by blast have "Bet D E P" using Q2 by blast thus ?thesis using l5_2 using R3 \Bet D E F\ by blast qed then have "A B C CongA D E F" by (smt P1 P2 bet_out l11_10 l6_6 out_trivial) } { assume S1: "E Out X P" have S2: "E Out P F" proof - { assume "Bet E X P" then have "E Out P F" proof - have "Bet E X F" by (meson Bet_perm Q2 Q4 \Bet E X P\ between_exchange3) thus ?thesis by (metis Bet_perm S1 bet2__out between_equality_2 between_trivial2 out2_bet_out out_diff1) qed } { assume "Bet E P X" then have "E Out P F" by (smt Bet_perm Q2 Q4 S1 bet_out_1 between_exchange3 not_bet_and_out outer_transitivity_between2) } thus ?thesis using Out_def S1 \Bet E X P \ E Out P F\ by blast qed then have "A B C CongA D E F" by (metis Bet_perm P2 Q1 Q2 bet_out__bet conga_line) } thus ?thesis using Q4 \X = E \ A B C CongA D E F\ by blast qed then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using conga_diff1 conga_diff2 inangle3123 by blast } { assume "B Out A C" obtain Q where "D E F CongA A B Q" by (metis P2 angle_construction_3) then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" by (metis Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms \B Out A C\ conga_diff1 conga_sym out321__inangle) } { assume ZZ: "\ Col A B C" have Z1: "D \ E" using P2 by blast have Z2: "F \ E" using P2 by blast have Z3: "Bet D E F \ E Out D F \ \ Col D E F" using not_bet_out by blast { assume "Bet D E F" obtain Q where Z4: "Bet A B Q \ Cong B Q E F" using segment_construction by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" by (metis InAngle_def P3 Z1 Z2 \Bet D E F\ conga_line point_construction_different) } { assume "E Out D F" then have Z5: "E Out D P" using P1 in_angle_out by blast have "D E P CongA A B C" by (simp add: P1 conga_sym) then have Z6: "B Out A C" using l11_21_a Z5 by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using \B Out A C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } { assume W1: "\ Col D E F" obtain Q where W2: "D E F CongA A B Q \ A B OS Q C" - using W1 ZZ angle_construction_1 by moura + using W1 ZZ angle_construction_1 by blast obtain DD where W3: "E Out D DD \ Cong E DD B A" using P3 Z1 segment_construction_3 by force obtain FF where W4: "E Out F FF \ Cong E FF B Q" by (metis P2 W2 conga_diff56 segment_construction_3) then have W5: "P InAngle DD E FF" by (smt Out_cases P1 P2 W3 l11_25 out_trivial) obtain X where W6: "Bet DD X FF \ (X = E \ E Out X P)" using InAngle_def W5 by presburger { assume W7: "X = E" have W8: "Bet D E F" proof - have W10: "E Out DD D" by (simp add: W3 l6_6) have "E Out FF F" by (simp add: W4 l6_6) thus ?thesis using W6 W7 W10 bet_out_out_bet by blast qed then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using \Bet D E F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } { assume V1: "E Out X P" have "B \ C \ E \ X" using P3 V1 out_diff1 by blast then obtain CC where V2: "B Out C CC \ Cong B CC E X" using segment_construction_3 by blast then have V3: "A B CC CongA DD E X" by (smt P1 P2 V1 W3 l11_10 l6_6 out_trivial) have V4: "Cong A CC DD X" proof - have "Cong A B DD E" using W3 not_cong_4321 by blast thus ?thesis using V2 V3 cong2_conga_cong by blast qed have V5: "A B Q CongA DD E FF" proof - have U1: "D E F CongA A B Q" by (simp add: W2) then have U1A: "A B Q CongA D E F" by (simp add: conga_sym) have U2: "B Out A A" by (simp add: P3 out_trivial) have U3: "B Out Q Q" using W2 conga_diff56 out_trivial by blast have U4: "E Out DD D" using W3 l6_6 by blast have "E Out FF F" by (simp add: W4 l6_6) thus ?thesis using l11_10 using U1A U2 U3 U4 by blast qed then have V6: "Cong A Q DD FF" using Cong_perm W3 W4 cong2_conga_cong by blast have "CC B Q CongA X E FF" proof - have U1: "B A OS CC Q" by (metis (no_types) V2 W2 col124__nos invert_one_side one_side_symmetry one_side_transitivity out_one_side) have U2: "E DD OS X FF" proof - have "\ Col DD E FF" by (metis Col_perm OS_def TS_def U1 V5 ncol_conga_ncol) then have "\ Col E DD X" by (metis Col_def V2 V4 W6 ZZ cong_identity l6_16_1 os_distincts out_one_side) then have "DD E OS X FF" by (metis Col_perm W6 bet_out not_col_distincts one_side_reflexivity out_out_one_side) thus ?thesis by (simp add: invert_one_side) qed have "CC B A CongA X E DD" by (simp add: V3 conga_comm) thus ?thesis using U1 U2 V5 l11_22b by blast qed then have V8: "Cong CC Q X FF" using V2 W4 cong2_conga_cong cong_commutativity not_cong_3412 by blast have V9: "CC InAngle A B Q" proof - have T2: "Q \ B" using W2 conga_diff56 by blast have T3: "CC \ B" using V2 out_distinct by blast have "Bet A CC Q" proof - have T4: "DD X FF Cong3 A CC Q" using Cong3_def V4 V6 V8 not_cong_3412 by blast thus ?thesis using W6 l4_6 by blast qed then have "\ X0. Bet A X0 Q \ (X0 = B \ B Out X0 CC)" using out_trivial by blast thus ?thesis by (simp add: InAngle_def P3 T2 T3) qed then have "C InAngle A B Q" using V2 inangle_distincts l11_25 out_trivial by blast then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using W2 conga_sym by blast } then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using W6 \X = E \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } then have "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" using Z3 \E Out D F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ \Bet D E F \ \Q. C InAngle A B Q \ A B Q CongA D E F\ by blast } thus ?thesis using \B Out A C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ \Bet A B C \ \Q. C InAngle A B Q \ A B Q CongA D E F\ not_bet_out by blast qed lemma in_angle_line: assumes "P \ B" and "A \ B" and "C \ B" and "Bet A B C" shows "P InAngle A B C" using InAngle_def assms(1) assms(2) assms(3) assms(4) by auto lemma l11_29_b: assumes "\ Q. (C InAngle A B Q \ A B Q CongA D E F)" shows "A B C LeA D E F" proof - obtain Q where P1: "C InAngle A B Q \ A B Q CongA D E F" using assms by blast obtain X where P2: "Bet A X Q \ (X = B \ B Out X C)" using InAngle_def P1 by auto { assume P2A: "X = B" obtain P where P3: "A B C CongA D E P" using angle_construction_3 assms conga_diff45 inangle_distincts by fastforce have "P InAngle D E F" proof - have O1: "Bet D E F" by (metis (no_types) P1 P2 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms P2A) have O2: "P \ E" using P3 conga_diff56 by auto have O3: "D \ E" using P3 conga_diff45 by auto have "F \ E" using P1 conga_diff56 by blast thus ?thesis using in_angle_line by (simp add: O1 O2 O3) qed then have "A B C LeA D E F" using LeA_def P3 by blast } { assume G1: "B Out X C" obtain DD where G2: "E Out D DD \ Cong E DD B A" by (metis assms conga_diff1 conga_diff45 segment_construction_3) have G3: "D \ E \ DD \ E" using G2 out_diff1 out_diff2 by blast obtain FF where G3G: "E Out F FF \ Cong E FF B Q" by (metis P1 conga_diff56 inangle_distincts segment_construction_3) then have G3A: "F \ E" using out_diff1 by blast have G3B: "FF \ E" using G3G out_distinct by blast have G4: "Bet A B C \ B Out A C \ \ Col A B C" using not_bet_out by blast { assume G5: "Bet A B C" have G6: "F InAngle D E F" by (simp add: G3 G3A inangle3123) have "A B C CongA D E F" by (smt Bet_perm G3 G3A G5 Out_def P1 P2 bet_conga__bet between_exchange3 conga_line inangle_distincts outer_transitivity_between2) then have "A B C LeA D E F" using G6 LeA_def by blast } { assume G8: "B Out A C" have G9: "D InAngle D E F" by (simp add: G3 G3A inangle1123) have "A B C CongA D E D" by (simp add: G3 G8 l11_21_b out_trivial) then have "A B C LeA D E F" using G9 LeA_def by blast } { assume R1: "\ Col A B C" have R2: "Bet A B Q \ B Out A Q \ \ Col A B Q" using not_bet_out by blast { assume R3: "Bet A B Q" obtain P where R4: "A B C CongA D E P" by (metis G3 LeA_def \Bet A B C \ A B C LeA D E F\ angle_construction_3 not_bet_distincts) have R5: "P InAngle D E F" proof - have R6: "P \ E" using R4 conga_diff56 by auto have "Bet D E F" by (metis (no_types) P1 R3 Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless_axioms) thus ?thesis by (simp add: R6 G3 G3A in_angle_line) qed then have "A B C LeA D E F" using R4 R5 LeA_def by blast } { assume S1: "B Out A Q" have S2: "B Out A C" using G1 P2 S1 l6_7 out_bet_out_1 by blast have S3: "Col A B C" by (simp add: Col_perm S2 out_col) then have "A B C LeA D E F" using R1 by blast } { assume S3B: "\ Col A B Q" obtain P where S4: "A B C CongA D E P \ D E OS P F" by (meson P1 R1 Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms S3B angle_construction_1) obtain PP where S4A: "E Out P PP \ Cong E PP B X" by (metis G1 S4 os_distincts out_diff1 segment_construction_3) have S5: "P InAngle D E F" proof - have "PP InAngle DD E FF" proof - have Z3: "PP \ E" using S4A l6_3_1 by blast have Z4: "Bet DD PP FF" proof - have L1: "C B Q CongA P E F" proof - have K1: "B A OS C Q" using Col_perm P1 R1 S3B in_angle_one_side invert_one_side by blast have K2: "E D OS P F" by (simp add: S4 invert_one_side) have "C B A CongA P E D" by (simp add: S4 conga_comm) thus ?thesis using K1 K2 P1 l11_22b by auto qed have L2: "Cong DD FF A Q" proof - have "DD E FF CongA A B Q" proof - have L3: "A B Q CongA D E F" by (simp add: P1) then have L3A: "D E F CongA A B Q" using conga_sym by blast have L4: "E Out DD D" using G2 Out_cases by auto have L5: "E Out FF F" using G3G Out_cases by blast have L6: "B Out A A" using S3B not_col_distincts out_trivial by auto have "B Out Q Q" by (metis S3B not_col_distincts out_trivial) thus ?thesis using L3A L4 L5 L6 l11_10 by blast qed have L2B: "Cong DD E A B" using Cong_perm G2 by blast have "Cong E FF B Q" by (simp add: G3G) thus ?thesis using L2B \DD E FF CongA A B Q\ cong2_conga_cong by auto qed have L8: "Cong A X DD PP" proof - have L9: "A B X CongA DD E PP" proof - have L9B: "B Out A A" using S3B not_col_distincts out_trivial by blast have L9D: "E Out D D " using G3 out_trivial by auto have "E Out PP P" using Out_cases S4A by blast thus ?thesis using l11_10 S4 L9B G1 L9D using G2 Out_cases by blast qed have L10: "Cong A B DD E" using G2 not_cong_4321 by blast have "Cong B X E PP" using Cong_perm S4A by blast thus ?thesis using L10 L9 cong2_conga_cong by blast qed have "A X Q Cong3 DD PP FF" proof - have L12B: "Cong A Q DD FF" using L2 not_cong_3412 by blast have "Cong X Q PP FF" proof - have L13A: "X B Q CongA PP E FF" proof - have L13AC: "B Out Q Q" by (metis S3B col_trivial_2 out_trivial) have L13AD: "E Out PP P" by (simp add: S4A l6_6) have "E Out FF F" by (simp add: G3G l6_6) thus ?thesis using L1 G1 L13AC L13AD l11_10 by blast qed have L13B: "Cong X B PP E" using S4A not_cong_4321 by blast have "Cong B Q E FF" using G3G not_cong_3412 by blast thus ?thesis using L13A L13B cong2_conga_cong by auto qed thus ?thesis by (simp add: Cong3_def L12B L8) qed thus ?thesis using P2 l4_6 by blast qed have "PP = E \ E Out PP PP" using out_trivial by auto thus ?thesis using InAngle_def G3 G3B Z3 Z4 by auto qed thus ?thesis using G2 G3G S4A l11_25 by blast qed then have "A B C LeA D E F" using S4 LeA_def by blast } then have "A B C LeA D E F" using R2 \B Out A Q \ A B C LeA D E F\ \Bet A B Q \ A B C LeA D E F\ by blast } then have "A B C LeA D E F" using G4 \B Out A C \ A B C LeA D E F\ \Bet A B C \ A B C LeA D E F\ by blast } thus ?thesis using P2 \X = B \ A B C LeA D E F\ by blast qed lemma bet_in_angle_bet: assumes "Bet A B P" and "P InAngle A B C" shows "Bet A B C" by (metis (no_types) Col_def Col_perm assms(1) assms(2) col_in_angle_out not_bet_and_out) lemma lea_line: assumes "Bet A B P" and "A B P LeA A B C" shows "Bet A B C" by (metis Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless.l11_29_a Tarski_neutral_dimensionless_axioms assms(1) assms(2) bet_in_angle_bet) lemma eq_conga_out: assumes "A B A CongA D E F" shows "E Out D F" by (metis CongA_def assms l11_21_a out_trivial) lemma out_conga_out: assumes "B Out A C" and "A B C CongA D E F" shows "E Out D F" using assms(1) assms(2) l11_21_a by blast lemma conga_ex_cong3: assumes "A B C CongA A' B' C'" shows "\ AA CC. ((B Out A AA \ B Out C CC) \ AA B CC Cong3 A' B' C')" using out_diff2 by blast lemma conga_preserves_in_angle: assumes "A B C CongA A' B' C'" and "A B I CongA A' B' I'" and "I InAngle A B C" and "A' B' OS I' C'" shows "I' InAngle A' B' C'" proof - have P1: "A \ B" using assms(1) conga_diff1 by auto have P2: "B \ C" using assms(1) conga_diff2 by blast have P3: "A' \ B'" using assms(1) conga_diff45 by auto have P4: "B' \ C'" using assms(1) conga_diff56 by blast have P5: "I \ B" using assms(2) conga_diff2 by auto have P6: "I' \ B'" using assms(2) conga_diff56 by blast have P7: "Bet A B C \ B Out A C \ \ Col A B C" using l6_4_2 by blast { assume "Bet A B C" have Q1: "Bet A' B' C'" using \Bet A B C\ assms(1) assms(4) bet_col col124__nos col_conga_col by blast then have "I' InAngle A' B' C'" using assms(4) bet_col col124__nos by auto } { assume "B Out A C" then have "I' InAngle A' B' C'" by (metis P4 assms(2) assms(3) in_angle_out l11_21_a out321__inangle) } { assume Z1: "\ Col A B C" have Q2: "Bet A B I \ B Out A I \ \ Col A B I" by (simp add: or_bet_out) { assume "Bet A B I" then have "I' InAngle A' B' C'" using \Bet A B C \ I' InAngle A' B' C'\ assms(3) bet_in_angle_bet by blast } { assume "B Out A I" then have "I' InAngle A' B' C'" using P4 assms(2) l11_21_a out321__inangle by auto } { assume "\ Col A B I" obtain AA' where Q3: "B' Out A' AA' \ Cong B' AA' B A" using P1 P3 segment_construction_3 by presburger obtain CC' where Q4: "B' Out C' CC' \ Cong B' CC' B C" using P2 P4 segment_construction_3 by presburger obtain J where Q5: "Bet A J C \ (J = B \ B Out J I)" using InAngle_def assms(3) by auto have Q6: "B \ J" using Q5 Z1 bet_col by auto have Q7: "\ Col A B J" using Q5 Q6 \\ Col A B I\ col_permutation_2 col_transitivity_1 out_col by blast have "\ Col A' B' I'" by (metis assms(4) col123__nos) then have "\ C'. (A B J CongA A' B' C' \ A' B' OS C' I')" using Q7 angle_construction_1 by blast then obtain J' where Q8: "A B J CongA A' B' J' \ A' B' OS J' I'" by blast have Q9: "B' \ J'" using Q8 conga_diff56 by blast obtain JJ' where Q10: "B' Out J' JJ' \ Cong B' JJ' B J" using segment_construction_3 Q6 Q9 by blast have Q11: "\ Col A' B' J'" using Q8 col123__nos by blast have Q12: "A' \ JJ'" by (metis Col_perm Q10 Q11 out_col) have Q13: "B' \ JJ'" using Q10 out_distinct by blast have Q14: "\ Col A' B' JJ'" using Col_perm Q10 Q11 Q13 l6_16_1 out_col by blast have Q15: "A B C CongA AA' B' CC'" proof - have T2: "C \ B" using P2 by auto have T3: "AA' \ B'" using Out_def Q3 by blast have T4: "CC' \ B'" using Q4 out_distinct by blast have T5: "\ A' C' D' F'. (B Out A' A \ B Out C' C \ B' Out D' AA' \ B' Out F' CC' \Cong B A' B' D' \ Cong B C' B' F' \ Cong A' C' D' F')" by (smt Q3 Q4 Tarski_neutral_dimensionless.l11_4_1 Tarski_neutral_dimensionless_axioms assms(1) l6_6 l6_7) thus ?thesis using P1 T2 T3 T4 l11_4_2 by blast qed have Q16: "A' B' J' CongA A' B' JJ'" proof - have P9: "B' Out A' A'" by (simp add: P3 out_trivial) have "B' Out JJ' J'" using Out_cases Q10 by auto thus ?thesis using l11_10 by (simp add: P9 out2__conga) qed have Q17: "B' Out I' JJ' \ A' B' TS I' JJ'" proof - have "Coplanar A' I' B' J'" by (metis (full_types) Q8 ncoplanar_perm_3 os__coplanar) then have "Coplanar A' I' B' JJ'" using Q10 Q9 col_cop__cop out_col by blast then have R1: "Coplanar A' B' I' JJ'" using coplanar_perm_2 by blast have "A' B' I' CongA A' B' JJ'" proof - have R2: "A' B' I' CongA A B I" by (simp add: assms(2) conga_sym) have "A B I CongA A' B' JJ'" proof - have f1: "\p pa pb. \ p Out pa pb \ \ p Out pb pa \ p Out pa pb" using Out_cases by blast then have f2: "B' Out JJ' J'" using Q10 by blast have "B Out J I" by (metis Q5 Q6) thus ?thesis using f2 f1 by (meson P3 Q8 Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms \\ Col A B I\ col_one_side_out col_trivial_2 one_side_reflexivity out_trivial) qed thus ?thesis using R2 conga_trans by blast qed thus ?thesis using R1 conga_cop__or_out_ts by blast qed { assume Z2: "B' Out I' JJ'" have Z3: "J B C CongA J' B' C'" proof - have R1: "B A OS J C" by (metis Q5 Q7 Z1 bet_out invert_one_side not_col_distincts out_one_side) have R2: "B' A' OS J' C'" by (meson Q10 Z2 assms(4) invert_one_side l6_6 one_side_symmetry out_out_one_side) have "J B A CongA J' B' A'" using Q8 conga_comm by blast thus ?thesis using assms(1) R1 R2 l11_22b by blast qed then have "I' InAngle A' B' C'" proof - have "A J C Cong3 AA' JJ' CC'" proof - have R8: "Cong A J AA' JJ'" proof - have R8A: "A B J CongA AA' B' JJ'" proof - have R8AB: "B Out A A" by (simp add: P1 out_trivial) have R8AC: "B Out J I" using Q5 Q6 by auto have R8AD: "B' Out AA' A'" using Out_cases Q3 by auto have "B' Out JJ' I'" using Out_cases Z2 by blast thus ?thesis using assms(2) R8AB R8AC R8AD l11_10 by blast qed have R8B: "Cong A B AA' B'" using Q3 not_cong_4321 by blast have R8C: "Cong B J B' JJ'" using Q10 not_cong_3412 by blast thus ?thesis using R8A R8B cong2_conga_cong by blast qed have LR8A: "Cong A C AA' CC'" using Q15 Q3 Q4 cong2_conga_cong cong_4321 cong_symmetry by blast have "Cong J C JJ' CC'" proof - have K1:"B' Out JJ' J'" using Out_cases Q10 by auto have "B' Out CC' C'" using Out_cases Q4 by auto then have "J' B' C' CongA JJ' B' CC'" using K1 by (simp add: out2__conga) then have LR9A: "J B C CongA JJ' B' CC'" using Z3 conga_trans by blast have LR9B: "Cong J B JJ' B'" using Q10 not_cong_4321 by blast have "Cong B C B' CC'" using Q4 not_cong_3412 by blast thus ?thesis using LR9A LR9B cong2_conga_cong by blast qed thus ?thesis using R8 LR8A by (simp add: Cong3_def) qed then have R10: "Bet AA' JJ' CC'" using Q5 l4_6 by blast have "JJ' InAngle AA' B' CC'" proof - have R11: "AA' \ B'" using Out_def Q3 by auto have R12: "CC' \ B'" using Out_def Q4 by blast have "Bet AA' JJ' CC' \ (JJ' = B' \ B' Out JJ' JJ')" using R10 out_trivial by auto thus ?thesis using InAngle_def Q13 R11 R12 by auto qed thus ?thesis using Z2 Q3 Q4 l11_25 by blast qed } { assume X1: "A' B' TS I' JJ'" have "A' B' OS I' J'" by (simp add: Q8 one_side_symmetry) then have X2: "B' A' OS I' JJ'" using Q10 invert_one_side out_out_one_side by blast then have "I' InAngle A' B' C'" using X1 invert_one_side l9_9 by blast } then have "I' InAngle A' B' C'" using Q17 \B' Out I' JJ' \ I' InAngle A' B' C'\ by blast } then have "I' InAngle A' B' C'" using Q2 \B Out A I \ I' InAngle A' B' C'\ \Bet A B I \ I' InAngle A' B' C'\ by blast } thus ?thesis using P7 \B Out A C \ I' InAngle A' B' C'\ \Bet A B C \ I' InAngle A' B' C'\ by blast qed lemma l11_30: assumes "A B C LeA D E F" and "A B C CongA A' B' C'" and "D E F CongA D' E' F'" shows "A' B' C' LeA D' E' F'" proof - obtain Q where P1: "C InAngle A B Q \ A B Q CongA D E F" using assms(1) l11_29_a by blast have P1A: "C InAngle A B Q" using P1 by simp have P1B: "A B Q CongA D E F" using P1 by simp have P2: "A \ B" using P1A inangle_distincts by auto have P3: "C \ B" using P1A inangle_distincts by blast have P4: "A' \ B'" using CongA_def assms(2) by blast have P5: "C' \ B'" using CongA_def assms(2) by auto have P6: "D \ E" using CongA_def P1B by blast have P7: "F \ E" using CongA_def P1B by blast have P8: "D' \ E'" using CongA_def assms(3) by blast have P9: "F' \ E'" using CongA_def assms(3) by blast have P10: "Bet A' B' C' \ B' Out A' C' \ \ Col A' B' C'" using or_bet_out by blast { assume "Bet A' B' C'" then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (metis P1 P4 P5 P8 P9 assms(2) assms(3) bet_conga__bet bet_in_angle_bet conga_line conga_sym inangle3123) } { assume R1: "B' Out A' C'" obtain Q' where R2: "D' E' F' CongA A' B' Q'" using P4 P8 P9 angle_construction_3 by blast then have "C' InAngle A' B' Q'" using col_in_angle P1 R1 conga_diff56 out321__inangle by auto then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using R2 conga_sym by blast } { assume R3: "\ Col A' B' C'" have R3A: "Bet D' E' F' \ E' Out D' F' \ \ Col D' E' F'" using or_bet_out by blast { assume "Bet D' E' F'" have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (metis P4 P5 P8 P9 \Bet D' E' F'\ conga_line in_angle_line point_construction_different) } { assume R4A: "E' Out D' F'" obtain Q' where R4: "D' E' F' CongA A' B' Q'" using P4 P8 P9 angle_construction_3 by blast then have R5: "B' Out A' Q'" using out_conga_out R4A by blast have R6: "A B Q CongA D' E' F'" using P1 assms(3) conga_trans by blast then have R7: "B Out A Q" using out_conga_out R4A R4 using conga_sym by blast have R8: "B Out A C" using P1A R7 in_angle_out by blast then have R9: "B' Out A' C'" using out_conga_out assms(2) by blast have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" by (simp add: R9 \B' Out A' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\) } { assume "\ Col D' E' F'" obtain QQ where S1: "D' E' F' CongA A' B' QQ \ A' B' OS QQ C'" using R3 \\ Col D' E' F'\ angle_construction_1 by blast have S1A: "A B Q CongA A' B' QQ" using S1 using P1 assms(3) conga_trans by blast have "A' B' OS C' QQ" using S1 by (simp add: S1 one_side_symmetry) then have S2: "C' InAngle A' B' QQ" using conga_preserves_in_angle S1A using P1A assms(2) by blast have S3: "A' B' QQ CongA D' E' F'" by (simp add: S1 conga_sym) then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using S2 by auto } then have "\ Q'. (C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F')" using R3A \E' Out D' F' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ \Bet D' E' F' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ by blast } thus ?thesis using l11_29_b using P10 \B' Out A' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ \Bet A' B' C' \ \Q'. C' InAngle A' B' Q' \ A' B' Q' CongA D' E' F'\ by blast qed lemma l11_31_1: assumes "B Out A C" and "D \ E" and "F \ E" shows "A B C LeA D E F" by (metis (full_types) LeA_def assms(1) assms(2) assms(3) l11_21_b out321__inangle segment_construction_3) lemma l11_31_2: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" and "Bet D E F" shows "A B C LeA D E F" by (metis LeA_def angle_construction_3 assms(1) assms(2) assms(3) assms(4) assms(5) conga_diff56 in_angle_line) lemma lea_refl: assumes "A \ B" and "C \ B" shows "A B C LeA A B C" by (meson assms(1) assms(2) conga_refl l11_29_b out341__inangle out_trivial) lemma conga__lea: assumes "A B C CongA D E F" shows "A B C LeA D E F" by (metis Tarski_neutral_dimensionless.conga_diff1 Tarski_neutral_dimensionless.conga_diff2 Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless_axioms assms conga_refl lea_refl) lemma conga__lea456123: assumes "A B C CongA D E F" shows "D E F LeA A B C" by (simp add: Tarski_neutral_dimensionless.conga__lea Tarski_neutral_dimensionless_axioms assms conga_sym) lemma lea_left_comm: assumes "A B C LeA D E F" shows "C B A LeA D E F" by (metis assms conga_pseudo_refl conga_refl l11_30 lea_distincts) lemma lea_right_comm: assumes "A B C LeA D E F" shows "A B C LeA F E D" by (meson assms conga_right_comm l11_29_a l11_29_b) lemma lea_comm: assumes"A B C LeA D E F" shows "C B A LeA F E D" using assms lea_left_comm lea_right_comm by blast lemma lta_left_comm: assumes "A B C LtA D E F" shows "C B A LtA D E F" by (meson LtA_def Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless_axioms assms) lemma lta_right_comm: assumes "A B C LtA D E F" shows "A B C LtA F E D" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms) lemma lta_comm: assumes "A B C LtA D E F" shows "C B A LtA F E D" using assms lta_left_comm lta_right_comm by blast lemma lea_out4__lea: assumes "A B C LeA D E F" and "B Out A A'" and "B Out C C'" and "E Out D D'" and "E Out F F'" shows "A' B C' LeA D' E F'" using assms(1) assms(2) assms(3) assms(4) assms(5) l11_30 l6_6 out2__conga by auto lemma lea121345: assumes "A \ B" and "C \ D" and "D \ E" shows "A B A LeA C D E" using assms(1) assms(2) assms(3) l11_31_1 out_trivial by auto lemma inangle__lea: assumes "P InAngle A B C" shows "A B P LeA A B C" by (metis Tarski_neutral_dimensionless.l11_29_b Tarski_neutral_dimensionless_axioms assms conga_refl inangle_distincts) lemma inangle__lea_1: assumes "P InAngle A B C" shows "P B C LeA A B C" by (simp add: Tarski_neutral_dimensionless.inangle__lea Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless_axioms assms l11_24) lemma inangle__lta: assumes "\ Col P B C" and "P InAngle A B C" shows "A B P LtA A B C" by (metis LtA_def TS_def Tarski_neutral_dimensionless.conga_cop__or_out_ts Tarski_neutral_dimensionless.conga_os__out Tarski_neutral_dimensionless.inangle__lea Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_one_side_out col_trivial_3 in_angle_one_side inangle__coplanar invert_two_sides l11_21_b ncoplanar_perm_12 not_col_permutation_3 one_side_reflexivity) lemma in_angle_trans: assumes "C InAngle A B D" and "D InAngle A B E" shows "C InAngle A B E" proof - obtain CC where P1: "Bet A CC D \ (CC = B \ B Out CC C)" using InAngle_def assms(1) by auto obtain DD where P2: "Bet A DD E \ (DD = B \ B Out DD D)" using InAngle_def assms(2) by auto then have P3: "Bet A DD E" by simp have P4: "DD = B \ B Out DD D" using P2 by simp { assume "CC = B \ DD = B" then have "C InAngle A B E" using InAngle_def P2 assms(1) assms(2) by auto } { assume "CC = B \ B Out DD D" then have "C InAngle A B E" by (metis InAngle_def P1 assms(1) assms(2) bet_in_angle_bet) } { assume "B Out CC C \ DD = B" then have "C InAngle A B E" by (metis Out_def P2 assms(2) in_angle_line inangle_distincts) } { assume P3: "B Out CC C \ B Out DD D" then have P3A: "B Out CC C" by simp have P3B: "B Out DD D" using P3 by simp have "C InAngle A B DD" using P3 assms(1) inangle_distincts l11_25 out_trivial by blast then obtain CC' where T1: "Bet A CC' DD \ (CC' = B \ B Out CC' C)" using InAngle_def by auto { assume "CC' = B" then have "C InAngle A B E" by (metis P2 P3 T1 assms(2) between_exchange4 in_angle_line inangle_distincts out_diff2) } { assume "B Out CC' C" then have "C InAngle A B E" by (metis InAngle_def P2 T1 assms(1) assms(2) between_exchange4) } then have "C InAngle A B E" using T1 \CC' = B \ C InAngle A B E\ by blast } thus ?thesis using P1 P2 \B Out CC C \ DD = B \ C InAngle A B E\ \CC = B \ B Out DD D \ C InAngle A B E\ \CC = B \ DD = B \ C InAngle A B E\ by blast qed lemma lea_trans: assumes "A B C LeA A1 B1 C1" and "A1 B1 C1 LeA A2 B2 C2" shows "A B C LeA A2 B2 C2" proof - obtain P1 where T1: "P1 InAngle A1 B1 C1 \ A B C CongA A1 B1 P1" using LeA_def assms(1) by auto obtain P2 where T2: "P2 InAngle A2 B2 C2 \ A1 B1 C1 CongA A2 B2 P2" using LeA_def assms(2) by blast have T3: "A \ B" using CongA_def T1 by auto have T4: "C \ B" using CongA_def T1 by blast have T5: "A1 \ B1" using T1 inangle_distincts by blast have T6: "C1 \ B1" using T1 inangle_distincts by blast have T7: "A2 \ B2" using T2 inangle_distincts by blast have T8: "C2 \ B2" using T2 inangle_distincts by blast have T9: "Bet A B C \ B Out A C \ \ Col A B C" using not_out_bet by auto { assume "Bet A B C" then have "A B C LeA A2 B2 C2" by (metis T1 T2 T3 T4 T7 T8 bet_conga__bet bet_in_angle_bet l11_31_2) } { assume "B Out A C" then have "A B C LeA A2 B2 C2" by (simp add: T7 T8 l11_31_1) } { assume H1: "\ Col A B C" have T10: "Bet A2 B2 C2 \ B2 Out A2 C2 \ \ Col A2 B2 C2" using not_out_bet by auto { assume "Bet A2 B2 C2" then have "A B C LeA A2 B2 C2" by (simp add: T3 T4 T7 T8 l11_31_2) } { assume T10A: "B2 Out A2 C2" have "B Out A C" proof - have "B1 Out A1 P1" proof - have "B1 Out A1 C1" using T2 conga_sym T2 T10A in_angle_out out_conga_out by blast thus ?thesis using T1 in_angle_out by blast qed thus ?thesis using T1 conga_sym l11_21_a by blast qed then have "A B C LeA A2 B2 C2" using \B Out A C \ A B C LeA A2 B2 C2\ by blast } { assume T12: "\ Col A2 B2 C2" obtain P where T13: "A B C CongA A2 B2 P \ A2 B2 OS P C2" using T12 H1 angle_construction_1 by blast have T14: "A2 B2 OS P2 C2" proof - have "\ Col B2 A2 P2" proof - have "B2 \ A2" using T7 by auto { assume H2: "P2 = A2" have "A2 B2 A2 CongA A1 B1 C1" using T2 H2 conga_sym by blast then have "B1 Out A1 C1" using eq_conga_out by blast then have "B1 Out A1 P1" using T1 in_angle_out by blast then have "B Out A C" using T1 conga_sym out_conga_out by blast then have False using Col_cases H1 out_col by blast } then have "P2 \ A2" by blast have "Bet A2 B2 P2 \ B2 Out A2 P2 \ \ Col A2 B2 P2" using not_out_bet by auto { assume H4: "Bet A2 B2 P2" then have "Bet A2 B2 C2" using T2 bet_in_angle_bet by blast then have "Col B2 A2 P2 \ False" using Col_def T12 by blast then have "\ Col B2 A2 P2" using H4 bet_col not_col_permutation_4 by blast } { assume H5: "B2 Out A2 P2" then have "B1 Out A1 C1" using T2 conga_sym out_conga_out by blast then have "B1 Out A1 P1" using T1 in_angle_out by blast then have "B Out A C" using H1 T1 ncol_conga_ncol not_col_permutation_4 out_col by blast then have "\ Col B2 A2 P2" using Col_perm H1 out_col by blast } { assume "\ Col A2 B2 P2" then have "\ Col B2 A2 P2" using Col_perm by blast } thus ?thesis using \B2 Out A2 P2 \ \ Col B2 A2 P2\ \Bet A2 B2 P2 \ \ Col B2 A2 P2\ \Bet A2 B2 P2 \ B2 Out A2 P2 \ \ Col A2 B2 P2\ by blast qed thus ?thesis by (simp add: T2 T12 in_angle_one_side) qed have S1: "A2 B2 OS P P2" using T13 T14 one_side_symmetry one_side_transitivity by blast have "A1 B1 P1 CongA A2 B2 P" using conga_trans conga_sym T1 T13 by blast then have "P InAngle A2 B2 P2" using conga_preserves_in_angle T2 T1 S1 by blast then have "P InAngle A2 B2 C2" using T2 in_angle_trans by blast then have "A B C LeA A2 B2 C2" using T13 LeA_def by blast } then have "A B C LeA A2 B2 C2" using T10 \B2 Out A2 C2 \ A B C LeA A2 B2 C2\ \Bet A2 B2 C2 \ A B C LeA A2 B2 C2\ by blast } thus ?thesis using T9 \B Out A C \ A B C LeA A2 B2 C2\ \Bet A B C \ A B C LeA A2 B2 C2\ by blast qed lemma in_angle_asym: assumes "D InAngle A B C" and "C InAngle A B D" shows "A B C CongA A B D" proof - obtain CC where P1: "Bet A CC D \ (CC = B \ B Out CC C)" using InAngle_def assms(2) by auto obtain DD where P2: "Bet A DD C \ (DD = B \ B Out DD D)" using InAngle_def assms(1) by auto { assume "(CC = B) \ (DD = B)" then have "A B C CongA A B D" by (metis P1 P2 assms(2) conga_line inangle_distincts) } { assume "(CC = B) \ (B Out DD D)" then have "A B C CongA A B D" by (metis P1 assms(1) bet_in_angle_bet conga_line inangle_distincts) } { assume "(B Out CC C) \ (DD = B)" then have "A B C CongA A B D" by (metis P2 assms(2) bet_in_angle_bet conga_line inangle_distincts) } { assume V1: "(B Out CC C) \ (B Out DD D)" obtain X where P3: "Bet CC X C \ Bet DD X D" using P1 P2 between_symmetry inner_pasch by blast then have "B Out X D" using V1 out_bet_out_2 by blast then have "B Out C D" using P3 V1 out2_bet_out by blast then have "A B C CongA A B D" using assms(2) inangle_distincts l6_6 out2__conga out_trivial by blast } thus ?thesis using P1 P2 using \B Out CC C \ DD = B \ A B C CongA A B D\ \CC = B \ B Out DD D \ A B C CongA A B D\ \CC = B \ DD = B \ A B C CongA A B D\ by blast qed lemma lea_asym: assumes "A B C LeA D E F" and "D E F LeA A B C" shows "A B C CongA D E F" proof cases assume P1: "Col A B C" { assume P1A: "Bet A B C" have P2: "D \ E" using assms(1) lea_distincts by blast have P3: "F \ E" using assms(2) lea_distincts by auto have P4: "A \ B" using assms(1) lea_distincts by auto have P5: "C \ B" using assms(2) lea_distincts by blast obtain P where P6: "P InAngle D E F \ A B C CongA D E P" using LeA_def assms(1) by blast then have "A B C CongA D E P" by simp then have "Bet D E P" using P1 P1A bet_conga__bet by blast then have "Bet D E F" using P6 bet_in_angle_bet by blast then have "A B C CongA D E F" by (metis Tarski_neutral_dimensionless.bet_conga__bet Tarski_neutral_dimensionless.conga_line Tarski_neutral_dimensionless.l11_29_a Tarski_neutral_dimensionless_axioms P2 P3 P4 P5 assms(2) bet_in_angle_bet) } { assume T1: "\ Bet A B C" then have T2: "B Out A C" using P1 not_out_bet by auto obtain P where T3: "P InAngle A B C \ D E F CongA A B P" using LeA_def assms(2) by blast then have T3A: "P InAngle A B C" by simp have T3B: "D E F CongA A B P" using T3 by simp have T4: "E Out D F" proof - have T4A: "B Out A P" using T2 T3 in_angle_out by blast have "A B P CongA D E F" by (simp add: T3 conga_sym) thus ?thesis using T4A l11_21_a by blast qed then have "A B C CongA D E F" by (simp add: T2 l11_21_b) } thus ?thesis using \Bet A B C \ A B C CongA D E F\ by blast next assume T5: "\ Col A B C" obtain Q where T6: "C InAngle A B Q \ A B Q CongA D E F" using assms(1) l11_29_a by blast then have T6A: "C InAngle A B Q" by simp have T6B: "A B Q CongA D E F" by (simp add: T6) obtain P where T7: "P InAngle A B C \ D E F CongA A B P" using LeA_def assms(2) by blast then have T7A: "P InAngle A B C" by simp have T7B: "D E F CongA A B P" by (simp add: T7) have T13: "A B Q CongA A B P" using T6 T7 conga_trans by blast have T14: "Bet A B Q \ B Out A Q \ \ Col A B Q" using not_out_bet by auto { assume R1: "Bet A B Q" then have "A B C CongA D E F" using T13 T5 T7 bet_col bet_conga__bet bet_in_angle_bet by blast } { assume R2: "B Out A Q" then have "A B C CongA D E F" using T6 in_angle_out l11_21_a l11_21_b by blast } { assume R3: "\ Col A B Q" have R3A: "Bet A B P \ B Out A P \ \ Col A B P" using not_out_bet by blast { assume R3AA: "Bet A B P" then have "A B C CongA D E F" using T5 T7 bet_col bet_in_angle_bet by blast } { assume R3AB: "B Out A P" then have "A B C CongA D E F" by (meson Col_cases R3 T13 ncol_conga_ncol out_col) } { assume R3AC: "\ Col A B P" have R3AD: "B Out P Q \ A B TS P Q" proof - have "Coplanar A B P Q" using T6A T7A coplanar_perm_8 in_angle_trans inangle__coplanar by blast thus ?thesis by (simp add: T13 conga_sym conga_cop__or_out_ts) qed { assume "B Out P Q" then have "C InAngle A B P" by (meson R3 T6A bet_col between_symmetry l11_24 l11_25_aux) then have "A B C CongA A B P" by (simp add: T7A in_angle_asym) then have "A B C CongA D E F" by (meson T7B Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms) } { assume W1: "A B TS P Q" have "A B OS P Q" using Col_perm R3 R3AC T6A T7A in_angle_one_side in_angle_trans by blast then have "A B C CongA D E F" using W1 l9_9 by blast } then have "A B C CongA D E F" using R3AD \B Out P Q \ A B C CongA D E F\ by blast } then have "A B C CongA D E F" using R3A \B Out A P \ A B C CongA D E F\ \Bet A B P \ A B C CongA D E F\ by blast } thus ?thesis using T14 \B Out A Q \ A B C CongA D E F\ \Bet A B Q \ A B C CongA D E F\ by blast qed lemma col_lta__bet: assumes "Col X Y Z" and "A B C LtA X Y Z" shows "Bet X Y Z" proof - have "A B C LeA X Y Z \ \ A B C CongA X Y Z" using LtA_def assms(2) by auto then have "Y Out X Z \ False" using Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_distincts Tarski_neutral_dimensionless_axioms l11_31_1 by fastforce thus ?thesis using not_out_bet assms(1) by blast qed lemma col_lta__out: assumes "Col A B C" and "A B C LtA X Y Z" shows "B Out A C" proof - have "A B C LeA X Y Z \ \ A B C CongA X Y Z" using LtA_def assms(2) by auto thus ?thesis by (metis assms(1) l11_31_2 lea_asym lea_distincts or_bet_out) qed lemma lta_distincts: assumes "A B C LtA D E F" shows "A\B \ C\B \ D\E \ F\E \ D \ F" by (metis LtA_def assms bet_neq12__neq col_lta__bet lea_distincts not_col_distincts) lemma gta_distincts: assumes "A B C GtA D E F" shows "A\B \ C\B \ D\E \ F\E \ A \ C" using GtA_def assms lta_distincts by presburger lemma acute_distincts: assumes "Acute A B C" shows "A\B \ C\B" using Acute_def assms lta_distincts by blast lemma obtuse_distincts: assumes "Obtuse A B C" shows "A\B \ C\B \ A \ C" using Obtuse_def assms lta_distincts by blast lemma two_sides_in_angle: assumes "B \ P'" and "B P TS A C" and "Bet P B P'" shows "P InAngle A B C \ P' InAngle A B C" proof - obtain T where P1: "Col T B P \ Bet A T C" using TS_def assms(2) by auto have P2: "A \ B" using assms(2) ts_distincts by blast have P3: "C \ B" using assms(2) ts_distincts by blast show ?thesis proof cases assume "B = T" thus ?thesis using P1 P2 P3 assms(1) in_angle_line by auto next assume "B \ T" thus ?thesis by (metis InAngle_def P1 assms(1) assms(2) assms(3) between_symmetry l6_3_2 or_bet_out ts_distincts) qed qed lemma in_angle_reverse: assumes "A' \ B" and "Bet A B A'" and "C InAngle A B D" shows "D InAngle A' B C" proof - have P1: "A \ B" using assms(3) inangle_distincts by auto have P2: "D \ B" using assms(3) inangle_distincts by blast have P3: "C \ B" using assms(3) inangle_distincts by auto show ?thesis proof cases assume "Col B A C" thus ?thesis by (smt P1 P2 P3 assms(1) assms(2) assms(3) bet_in_angle_bet between_inner_transitivity between_symmetry in_angle_line l6_3_2 out321__inangle outer_transitivity_between third_point) next assume P4: "\ Col B A C" thus ?thesis proof cases assume "Col B D C" thus ?thesis by (smt P2 P4 assms(1) assms(2) assms(3) bet_col1 col2__eq col_permutation_2 in_angle_one_side l9_19_R1 out341__inangle) next assume P5: "\ Col B D C" have P6: "C B TS A D" using P4 P5 assms(3) in_angle_two_sides by auto obtain X where P7: "Bet A X D \ (X = B \ B Out X C)" using InAngle_def assms(3) by auto have P8: "X = B \ D InAngle A' B C" using Out_def P1 P2 P3 P7 assms(1) assms(2) l5_2 out321__inangle by auto { assume P9: "B Out X C" have P10: "C \ B" by (simp add: P3) have P10A: "\ Col B A C" by (simp add: P4) have P10B: "\ Col B D C" by (simp add: P5) have P10C: "C InAngle D B A" by (simp add: assms(3) l11_24) { assume "Col D B A" have "Col B A C" proof - have "B \ X" using P9 out_distinct by blast have "Col B X A" by (meson Bet_perm P10C P5 P7 \Col D B A\ bet_col1 col_permutation_3 in_angle_out or_bet_out out_col) have "Col B X C" by (simp add: P9 out_col) thus ?thesis using \B \ X\ \Col B X A\ col_transitivity_1 by blast qed then have False by (simp add: P4) } then have P10E: "\ Col D B A" by auto have P11: "D B OS C A" by (simp add: P10C P10E P5 in_angle_one_side) have P12: "\ Col A D B" using Col_cases P10E by auto have P13: "\ Col A' D B" by (metis Col_def \Col D B A \ False\ assms(1) assms(2) col_transitivity_1) have P14: "D B TS A A'" using P12 P13 TS_def assms(2) col_trivial_3 by blast have P15: "D B TS C A'" using P11 P14 l9_8_2 one_side_symmetry by blast have P16: "\ Col C D B" by (simp add: P5 not_col_permutation_3) obtain Y where P17: "Col Y D B \ Bet C Y A'" using P15 TS_def by auto have P18: "Bet A' Y C" using Bet_perm P17 by blast { assume S1: "Y \ B" have S2: "Col D B Y" using P17 not_col_permutation_2 by blast then have S3: "Bet D B Y \ Bet B Y D \ Bet Y D B" using Col_def S2 by auto { assume S4: "Bet D B Y" have S5: "C B OS A' Y" by (metis P17 P18 P5 S1 bet_out_1 col_transitivity_2 l6_6 not_col_permutation_3 not_col_permutation_5 out_one_side) have S6: "C B TS Y D" by (metis Bet_perm P16 P17 S1 S4 bet__ts col3 col_trivial_3 invert_two_sides not_col_permutation_1) have "C B TS A A'" by (metis (full_types) P4 assms(1) assms(2) bet__ts invert_two_sides not_col_permutation_5) then have "C B TS Y A" using S5 l9_2 l9_8_2 by blast then have S9: "C B OS A D" using P6 S6 l9_8_1 l9_9 by blast then have "B Out Y D" using P6 S9 l9_9 by auto } { assume "Bet B Y D" then have "B Out Y D" by (simp add: S1 bet_out) } { assume "Bet Y D B" then have "B Out Y D" by (simp add: P2 bet_out_1 l6_6) } have "B Out Y D" using S3 \Bet B Y D \ B Out Y D\ \Bet D B Y \ B Out Y D\ \Bet Y D B \ B Out Y D\ by blast } then have P19: "(Y = B \ B Out Y D)" by auto have "D InAngle A' B C" using InAngle_def P18 P19 P2 P3 assms(1) by auto } thus ?thesis using P7 P8 by blast qed qed qed lemma in_angle_trans2: assumes "C InAngle A B D" and "D InAngle A B E" shows "D InAngle C B E" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. Bet p pa (pp p pa) \ pa \ (pp p pa)" using point_construction_different by moura then have f2: "\p. C InAngle D B (pp p B) \ \ D InAngle p B A" by (metis assms(1) in_angle_reverse in_angle_trans l11_24) have f3: "D InAngle E B A" using assms(2) l11_24 by blast then have "E \ B" by (simp add: inangle_distincts) thus ?thesis using f3 f2 f1 by (meson Bet_perm in_angle_reverse l11_24) qed lemma l11_36_aux1: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "A B C LeA D E F" shows "D' E F LeA A' B C" proof - obtain P where P1: "C InAngle A B P \ A B P CongA D E F" using assms(7) l11_29_a by blast thus ?thesis by (metis LeA_def Tarski_neutral_dimensionless.l11_13 Tarski_neutral_dimensionless_axioms assms(2) assms(4) assms(5) assms(6) conga_sym in_angle_reverse) qed lemma l11_36_aux2: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "D' E F LeA A' B C" shows "A B C LeA D E F" by (metis Bet_cases assms(1) assms(3) assms(5) assms(6) assms(7) l11_36_aux1 lea_distincts) lemma l11_36: assumes "A \ B" and "A' \ B" and "D \ E" and "D' \ E" and "Bet A B A'" and "Bet D E D'" shows "A B C LeA D E F \ D' E F LeA A' B C" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l11_36_aux1 l11_36_aux2 by auto lemma l11_41_aux: assumes "\ Col A B C" and "Bet B A D" and "A \ D" shows "A C B LtA C A D" proof - obtain M where P1: "M Midpoint A C" using midpoint_existence by auto obtain P where P2: "M Midpoint B P" using symmetric_point_construction by auto have P3: "A C B Cong3 C A P" by (smt Cong3_def P1 P2 assms(1) l7_13_R1 l7_2 midpoint_distinct_1 not_col_distincts) have P4: "A \ C" using assms(1) col_trivial_3 by blast have P5: "B \ C" using assms(1) col_trivial_2 by blast have P7: "A \ M" using P1 P4 is_midpoint_id by blast have P8: "A C B CongA C A P" by (simp add: P3 P4 P5 cong3_conga) have P8A: "Bet D A B" using Bet_perm assms(2) by blast have P8B: "Bet P M B" by (simp add: P2 between_symmetry midpoint_bet) then obtain X where P9: "Bet A X P \ Bet M X D" using P8A inner_pasch by blast have P9A: "Bet A X P" by (simp add: P9) have P9B: "Bet M X D" by (simp add: P9) have P10A: "P InAngle C A D" proof - have K1: "P InAngle M A D" by (metis InAngle_def P3 P5 P7 P9 assms(3) bet_out cong3_diff2) have K2: "A Out C M" using Out_def P1 P4 P7 midpoint_bet by auto have K3: "A Out D D" using assms(3) out_trivial by auto have "A Out P P" using K1 inangle_distincts out_trivial by auto thus ?thesis using K1 K2 K3 l11_25 by blast qed then have P10: "A C B LeA C A D" using LeA_def P8 by auto { assume K5: "A C B CongA C A D" then have K6: "C A D CongA C A P" using P8 conga_sym conga_trans by blast have K7: "Coplanar C A D P" using P10A inangle__coplanar ncoplanar_perm_18 by blast then have K8: "A Out D P \ C A TS D P" by (simp add: K6 conga_cop__or_out_ts) { assume "A Out D P" then have "Col M B A" by (meson P8A P8B bet_col1 bet_out__bet between_symmetry not_col_permutation_4) then have K8F: "Col A M B" using not_col_permutation_1 by blast have "Col A M C" by (simp add: P1 bet_col midpoint_bet) then have "False" using K8F P7 assms(1) col_transitivity_1 by blast } then have K9: "\ A Out D P" by auto { assume V1: "C A TS D P" then have V3: "A C TS B P" by (metis P10A P8A assms(1) col_trivial_1 col_trivial_2 in_angle_reverse in_angle_two_sides invert_two_sides l11_24 l9_18 not_col_permutation_5) have "A C TS B D" by (simp add: assms(1) assms(2) assms(3) bet__ts not_col_permutation_5) then have "A C OS D P" using V1 V3 invert_two_sides l9_8_1 l9_9 by blast then have "False" using V1 invert_one_side l9_9 by blast } then have "\ C A TS D P" by auto then have "False" using K8 K9 by auto } then have "\ A C B CongA C A D" by auto thus ?thesis by (simp add: LtA_def P10) qed lemma l11_41: assumes "\ Col A B C" and "Bet B A D" and "A \ D" shows "A C B LtA C A D \ A B C LtA C A D" proof - have P1: "A C B LtA C A D" using assms(1) assms(2) assms(3) l11_41_aux by auto have "A B C LtA C A D" proof - obtain E where T1: "Bet C A E \ Cong A E C A" using segment_construction by blast have T1A: "Bet C A E" using T1 by simp have T1B: "Cong A E C A" using T1 by simp have T2: "A B C LtA B A E" using T1 assms(1) cong_reverse_identity l11_41_aux not_col_distincts not_col_permutation_5 by blast have T3: "B A C CongA C A B" by (metis assms(1) conga_pseudo_refl not_col_distincts) have T3A: "D A C CongA E A B" by (metis CongA_def T1 T3 assms(2) assms(3) cong_reverse_identity l11_13) then have T4: "B A E CongA C A D" using conga_comm conga_sym by blast have "A B C CongA A B C" using T2 Tarski_neutral_dimensionless.conga_refl Tarski_neutral_dimensionless.lta_distincts Tarski_neutral_dimensionless_axioms by fastforce then have T5: "A B C LeA C A D" by (meson T2 T4 Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless_axioms) have "\ A B C CongA C A D" by (meson T2 Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless_axioms T3A) thus ?thesis by (simp add: LtA_def T5) qed thus ?thesis by (simp add: P1) qed lemma not_conga: assumes "A B C CongA A' B' C'" and "\ A B C CongA D E F" shows "\ A' B' C' CongA D E F" by (meson assms(1) assms(2) conga_trans) lemma not_conga_sym: assumes "\ A B C CongA D E F" shows "\ D E F CongA A B C" using assms conga_sym by blast lemma not_and_lta: shows "\ (A B C LtA D E F \ D E F LtA A B C)" proof - { assume P1: "A B C LtA D E F \ D E F LtA A B C" then have "A B C CongA D E F" using LtA_def lea_asym by blast then have "False" using LtA_def P1 by blast } thus ?thesis by auto qed lemma conga_preserves_lta: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "A B C LtA D E F" shows "A' B' C' LtA D' E' F'" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.conga_trans Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3)) lemma lta_trans: assumes "A B C LtA A1 B1 C1" and "A1 B1 C1 LtA A2 B2 C2" shows "A B C LtA A2 B2 C2" proof - have P1: "A B C LeA A2 B2 C2" by (meson LtA_def assms(1) assms(2) lea_trans) { assume "A B C CongA A2 B2 C2" then have "False" by (meson Tarski_neutral_dimensionless.LtA_def Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_trans Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga__lea456123) } thus ?thesis using LtA_def P1 by blast qed lemma obtuse_sym: assumes "Obtuse A B C" shows "Obtuse C B A" by (meson Obtuse_def Tarski_neutral_dimensionless.lta_right_comm Tarski_neutral_dimensionless_axioms assms) lemma acute_sym: assumes "Acute A B C" shows "Acute C B A" by (meson Acute_def Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms) lemma acute_col__out: assumes "Col A B C" and "Acute A B C" shows "B Out A C" by (meson Tarski_neutral_dimensionless.Acute_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_lta__out) lemma col_obtuse__bet: assumes "Col A B C" and "Obtuse A B C" shows "Bet A B C" using Obtuse_def assms(1) assms(2) col_lta__bet by blast lemma out__acute: assumes "B Out A C" shows "Acute A B C" proof - have P1: "A \ B" using assms out_diff1 by auto then obtain D where P3: "B D Perp A B" using perp_exists by blast then have P4: "B \ D" using perp_distinct by auto have P5: "Per A B D" by (simp add: P3 l8_2 perp_per_1) have P6: "A B C LeA A B D" using P1 P4 assms l11_31_1 by auto { assume "A B C CongA A B D" then have "False" by (metis Col_cases P1 P4 P5 assms col_conga_col l8_9 out_col) } then have "A B C LtA A B D" using LtA_def P6 by auto thus ?thesis using P5 Acute_def by auto qed lemma bet__obtuse: assumes "Bet A B C" and "A \ B" and "B \ C" shows "Obtuse A B C" proof - obtain D where P1: "B D Perp A B" using assms(2) perp_exists by blast have P5: "B \ D" using P1 perp_not_eq_1 by auto have P6: "Per A B D" using P1 Perp_cases perp_per_1 by blast have P7: "A B D LeA A B C" using assms(2) assms(3) P5 assms(1) l11_31_2 by auto { assume "A B D CongA A B C" then have "False" using assms(2) P5 P6 assms(1) bet_col ncol_conga_ncol per_not_col by blast } then have "A B D LtA A B C" using LtA_def P7 by blast thus ?thesis using Obtuse_def P6 by blast qed lemma l11_43_aux: assumes "A \ B" and "A \ C" and "Per B A C \ Obtuse B A C" shows "Acute A B C" proof cases assume P1: "Col A B C" { assume "Per B A C" then have "Acute A B C" using Col_cases P1 assms(1) assms(2) per_col_eq by blast } { assume "Obtuse B A C" then have "Bet B A C" using P1 col_obtuse__bet col_permutation_4 by blast then have "Acute A B C" by (simp add: assms(1) bet_out out__acute) } thus ?thesis using \Per B A C \ Acute A B C\ assms(3) by blast next assume P2: "\ Col A B C" then have P3: "B \ C" using col_trivial_2 by auto obtain B' where P4: "Bet B A B' \ Cong A B' B A" using segment_construction by blast have P5: "\ Col B' A C" by (metis Col_def P2 P4 col_transitivity_2 cong_reverse_identity) then have P6: "B' \ A \ B' \ C" using not_col_distincts by blast then have P7: "A C B LtA C A B' \ A B C LtA C A B'" using P2 P4 l11_41 by auto then have P7A: "A C B LtA C A B'" by simp have P7B: "A B C LtA C A B'" by (simp add: P7) { assume "Per B A C" have "Acute A B C" by (metis Acute_def P4 P7B \Per B A C\ assms(1) bet_col col_per2__per col_trivial_3 l8_3 lta_right_comm) } { assume T1: "Obtuse B A C" then obtain a b c where T2: "Per a b c \ a b c LtA B A C" using Obtuse_def by blast then have T2A: "Per a b c" by simp have T2B: "a b c LtA B A C" by (simp add: T2) then have T3: "a b c LeA B A C \ \ a b c CongA B A C" by (simp add: LtA_def) then have T3A: "a b c LeA B A C" by simp have T3B: "\ a b c CongA B A C" by (simp add: T3) obtain P where T4: "P InAngle B A C \ a b c CongA B A P" using LeA_def T3 by blast then have T5: "Per B A P" using T4 T2 l11_17 by blast then have T6: "Per P A B" using l8_2 by blast have "Col A B B'" by (simp add: P4 bet_col col_permutation_4) then have "Per P A B'" using T6 assms(1) per_col by blast then have S3: "B A P CongA B' A P" using l8_2 P6 T5 T4 CongA_def assms(1) l11_16 by auto have "C A B' LtA P A B" proof - have S4: "B A P LeA B A C \ B' A C LeA B' A P" using P4 P6 assms(1) l11_36 by auto have S5: "C A B' LeA P A B" proof - have S6: "B A P LeA B A C" using T4 inangle__lea by auto have "B' A P CongA P A B" using S3 conga_left_comm not_conga_sym by blast thus ?thesis using P6 S4 S6 assms(2) conga_pseudo_refl l11_30 by auto qed { assume T10: "C A B' CongA P A B" have "Per B' A C" proof - have "B A P CongA B' A C" using T10 conga_comm conga_sym by blast thus ?thesis using T5 l11_17 by blast qed then have "Per C A B" using Col_cases P6 \Col A B B'\ l8_2 l8_3 by blast have "a b c CongA B A C" proof - have "a \ b" using T3A lea_distincts by auto have "c \ b" using T2B lta_distincts by blast have "Per B A C" using Per_cases \Per C A B\ by blast thus ?thesis using T2 \a \ b\ \c \ b\ assms(1) assms(2) l11_16 by auto qed then have "False" using T3B by blast } then have "\ C A B' CongA P A B" by blast thus ?thesis by (simp add: LtA_def S5) qed then have "A B C LtA B A P" by (meson P7 lta_right_comm lta_trans) then have "Acute A B C" using T5 using Acute_def by blast } thus ?thesis using \Per B A C \ Acute A B C\ assms(3) by blast qed lemma l11_43: assumes "A \ B" and "A \ C" and "Per B A C \ Obtuse B A C" shows "Acute A B C \ Acute A C B" using Per_perm assms(1) assms(2) assms(3) l11_43_aux obtuse_sym by blast lemma acute_lea_acute: assumes "Acute D E F" and "A B C LeA D E F" shows "Acute A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ D E F LtA A' B' C'" using Acute_def assms(1) by auto have P2: "A B C LeA A' B' C'" using LtA_def P1 assms(2) lea_trans by blast have "\ A B C CongA A' B' C'" by (meson LtA_def P1 assms(2) conga__lea456123 lea_asym lea_trans) then have "A B C LtA A' B' C'" by (simp add: LtA_def P2) thus ?thesis using Acute_def P1 by auto qed lemma lea_obtuse_obtuse: assumes "Obtuse D E F" and "D E F LeA A B C" shows "Obtuse A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ A' B' C' LtA D E F" using Obtuse_def assms(1) by auto then have P2: "A' B' C' LeA A B C" using LtA_def assms(2) lea_trans by blast have "\ A' B' C' CongA A B C" by (meson LtA_def P1 assms(2) conga__lea456123 lea_asym lea_trans) then have "A' B' C' LtA A B C" by (simp add: LtA_def P2) thus ?thesis using Obtuse_def P1 by auto qed lemma l11_44_1_a: assumes "A \ B" and "A \ C" and "Cong B A B C" shows "B A C CongA B C A" by (metis (no_types, opaque_lifting) Cong3_def assms(1) assms(2) assms(3) cong3_conga cong_inner_transitivity cong_pseudo_reflexivity) lemma l11_44_2_a: assumes "\ Col A B C" and "B A Lt B C" shows "B C A LtA B A C" proof - have T1: "A \ B" using assms(1) col_trivial_1 by auto have T3: "A \ C" using assms(1) col_trivial_3 by auto have "B A Le B C" by (simp add: assms(2) lt__le) then obtain C' where P1: "Bet B C' C \ Cong B A B C'" using assms(2) Le_def by blast have T5: "C \ C'" using P1 assms(2) cong__nlt by blast have T5A: "C' \ A" using Col_def Col_perm P1 assms(1) by blast then have T6: "C' InAngle B A C" using InAngle_def P1 T1 T3 out_trivial by auto have T7: "C' A C LtA A C' B \ C' C A LtA A C' B" proof - have W1: "\ Col C' C A" by (metis Col_def P1 T5 assms(1) col_transitivity_2) have W2: "Bet C C' B" using Bet_perm P1 by blast have "C' \ B" using P1 T1 cong_identity by blast thus ?thesis using l11_41 W1 W2 by simp qed have T90: "B A C' LtA B A C" proof - have T90A: "B A C' LeA B A C" by (simp add: T6 inangle__lea) have "B A C' CongA B A C'" using T1 T5A conga_refl by auto { assume "B A C' CongA B A C" then have R1: "A Out C' C" by (metis P1 T7 assms(1) bet_out conga_os__out lta_distincts not_col_permutation_4 out_one_side) have "B A OS C' C" by (metis Col_perm P1 T1 assms(1) bet_out cong_diff_2 out_one_side) then have "False" using Col_perm P1 T5 R1 bet_col col2__eq one_side_not_col123 out_col by blast } then have "\ B A C' CongA B A C" by blast thus ?thesis by (simp add: LtA_def T90A) qed have "B A C' CongA B C' A" using P1 T1 T5A l11_44_1_a by auto then have K2: "A C' B CongA B A C'" using conga_left_comm not_conga_sym by blast have "B C A LtA B A C'" proof - have K1: "B C A CongA B C A" using assms(1) conga_refl not_col_distincts by blast have "B C A LtA A C' B" proof - have "C' C A CongA B C A" proof - have K2: "C Out B C'" using P1 T5 bet_out_1 l6_6 by auto have "C Out A A" by (simp add: T3 out_trivial) thus ?thesis by (simp add: K2 out2__conga) qed have "A C' B CongA A C' B" using CongA_def K2 conga_refl by auto thus ?thesis using T7 \C' C A CongA B C A\ conga_preserves_lta by auto qed thus ?thesis using K1 K2 conga_preserves_lta by auto qed thus ?thesis using T90 lta_trans by blast qed lemma not_lta_and_conga: "\ ( A B C LtA D E F \ A B C CongA D E F)" by (simp add: LtA_def) lemma conga_sym_equiv: "A B C CongA A' B' C' \ A' B' C' CongA A B C" using not_conga_sym by blast lemma conga_dec: "A B C CongA D E F \ \ A B C CongA D E F" by auto lemma lta_not_conga: assumes "A B C LtA D E F" shows "\ A B C CongA D E F" using assms not_lta_and_conga by auto lemma lta__lea: assumes "A B C LtA D E F" shows "A B C LeA D E F" using LtA_def assms by auto lemma nlta: "\ A B C LtA A B C" using not_and_lta by blast lemma lea__nlta: assumes "A B C LeA D E F" shows "\ D E F LtA A B C" by (meson Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.not_lta_and_conga Tarski_neutral_dimensionless_axioms assms lta__lea) lemma lta__nlea: assumes "A B C LtA D E F" shows "\ D E F LeA A B C" using assms lea__nlta by blast lemma l11_44_1_b: assumes "\ Col A B C" and "B A C CongA B C A" shows "Cong B A B C" proof - have "B A Lt B C \ B A Gt B C \ Cong B A B C" by (simp add: or_lt_cong_gt) thus ?thesis by (meson Gt_def assms(1) assms(2) conga_sym l11_44_2_a not_col_permutation_3 not_lta_and_conga) qed lemma l11_44_2_b: assumes "B A C LtA B C A" shows "B C Lt B A" proof cases assume "Col A B C" thus ?thesis using Col_perm assms bet__lt1213 col_lta__bet lta_distincts by blast next assume P1: "\ Col A B C" then have P2: "A \ B" using col_trivial_1 by blast have P3: "A \ C" using P1 col_trivial_3 by auto have "B A Lt B C \ B A Gt B C \ Cong B A B C" by (simp add: or_lt_cong_gt) { assume "B A Lt B C" then have "B C Lt B A" using P1 assms l11_44_2_a not_and_lta by blast } { assume "B A Gt B C" then have "B C Lt B A" using Gt_def P1 assms l11_44_2_a not_and_lta by blast } { assume "Cong B A B C" then have "B A C CongA B C A" by (simp add: P2 P3 l11_44_1_a) then have "B C Lt B A" using assms not_lta_and_conga by blast } thus ?thesis by (meson P1 Tarski_neutral_dimensionless.not_and_lta Tarski_neutral_dimensionless_axioms \B A Gt B C \ B C Lt B A\ \B A Lt B C \ B A Gt B C \ Cong B A B C\ assms l11_44_2_a) qed lemma l11_44_1: assumes "\ Col A B C" shows "B A C CongA B C A \ Cong B A B C" using assms l11_44_1_a l11_44_1_b not_col_distincts by blast lemma l11_44_2: assumes "\ Col A B C" shows "B A C LtA B C A \ B C Lt B A" using assms l11_44_2_a l11_44_2_b not_col_permutation_3 by blast lemma l11_44_2bis: assumes "\ Col A B C" shows "B A C LeA B C A \ B C Le B A" proof - { assume P1: "B A C LeA B C A" { assume "B A Lt B C" then have "B C A LtA B A C" by (simp add: assms l11_44_2_a) then have "False" using P1 lta__nlea by auto } then have "\ B A Lt B C" by blast have "B C Le B A" using \\ B A Lt B C\ nle__lt by blast } { assume P2: "B C Le B A" have "B A C LeA B C A" proof cases assume "Cong B C B A" then have "B A C CongA B C A" by (metis assms conga_sym l11_44_1_a not_col_distincts) thus ?thesis by (simp add: conga__lea) next assume "\ Cong B C B A" then have "B A C LtA B C A" by (simp add: l11_44_2 assms Lt_def P2) thus ?thesis by (simp add: lta__lea) qed } thus ?thesis using \B A C LeA B C A \ B C Le B A\ by blast qed lemma l11_46: assumes "A \ B" and "B \ C" and "Per A B C \ Obtuse A B C" shows "B A Lt A C \ B C Lt A C" proof cases assume "Col A B C" thus ?thesis by (meson assms(1) assms(2) assms(3) bet__lt1213 bet__lt2313 col_obtuse__bet lt_left_comm per_not_col) next assume P1: "\ Col A B C" have P2: "A \ C" using P1 col_trivial_3 by auto have P3: "Acute B A C \ Acute B C A" using assms(1) assms(2) assms(3) l11_43 by auto then obtain A' B' C' where P4: "Per A' B' C' \ B C A LtA A' B' C'" using Acute_def P3 by auto { assume P5: "Per A B C" have P5A: "A C B CongA A C B" by (simp add: P2 assms(2) conga_refl) have S1: "A \ B" by (simp add: assms(1)) have S2: "B \ C" by (simp add: assms(2)) have S3: "A' \ B'" using P4 lta_distincts by blast have S4: "B' \ C'" using P4 lta_distincts by blast then have "A' B' C' CongA A B C" using l11_16 using S1 S2 S3 S4 P4 P5 by blast then have "A C B LtA A B C" using P5A P4 conga_preserves_lta lta_left_comm by blast } { assume "Obtuse A B C" obtain A'' B'' C'' where P6: "Per A'' B'' C'' \ A'' B'' C'' LtA A B C" using Obtuse_def \Obtuse A B C\ by auto have "B C A LtA A' B' C'" by (simp add: P4) then have P7: "A C B LtA A' B' C'" by (simp add: lta_left_comm) have "A' B' C' LtA A B C" proof - have U1: "A'' B'' C'' CongA A' B' C'" proof - have V2: "A'' \ B''" using P6 lta_distincts by blast have V3: "C'' \ B''" using P6 lta_distincts by blast have V5: "A' \ B'" using P7 lta_distincts by blast have "C' \ B'" using P4 lta_distincts by blast thus ?thesis using P6 V2 V3 P4 V5 by (simp add: l11_16) qed have U2: "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto have U3: "A'' B'' C'' LtA A B C" by (simp add: P6) thus ?thesis using U1 U2 conga_preserves_lta by auto qed then have "A C B LtA A B C" using P7 lta_trans by blast } then have "A C B LtA A B C" using \Per A B C \ A C B LtA A B C\ assms(3) by blast then have "A B Lt A C" by (simp add: l11_44_2_b) then have "B A Lt A C" using Lt_cases by blast have "C A B LtA C B A" proof - obtain A' B' C' where U4: "Per A' B' C' \ B A C LtA A' B' C'" using Acute_def P3 by blast { assume "Per A B C" then have W3: "A' B' C' CongA C B A" using U4 assms(2) l11_16 l8_2 lta_distincts by blast have W2: "C A B CongA C A B" using P2 assms(1) conga_refl by auto have "C A B LtA A' B' C'" by (simp add: U4 lta_left_comm) then have "C A B LtA C B A" using W2 W3 conga_preserves_lta by blast } { assume "Obtuse A B C" then obtain A'' B'' C'' where W4: "Per A'' B'' C'' \ A'' B'' C'' LtA A B C" using Obtuse_def by auto have W5: "C A B LtA A' B' C'" by (simp add: U4 lta_left_comm) have "A' B' C' LtA C B A" proof - have W6: "A'' B'' C'' CongA A' B' C'" using l11_16 W4 U4 using lta_distincts by blast have "C B A CongA C B A" using assms(1) assms(2) conga_refl by auto thus ?thesis using W4 W6 conga_left_comm conga_preserves_lta by blast qed then have "C A B LtA C B A" using W5 lta_trans by blast } thus ?thesis using \Per A B C \ C A B LtA C B A\ assms(3) by blast qed then have "C B Lt C A" by (simp add: l11_44_2_b) then have "C B Lt A C" using Lt_cases by auto then have "B C Lt A C" using Lt_cases by blast thus ?thesis by (simp add: \B A Lt A C\) qed lemma l11_47: assumes "Per A C B" and "H PerpAt C H A B" shows "Bet A H B \ A \ H \ B \ H" proof - have P1: "Per C H A" using assms(2) perp_in_per_1 by auto have P2: "C H Perp A B" using assms(2) perp_in_perp by auto thus ?thesis proof cases assume "Col A C B" thus ?thesis by (metis P1 assms(1) assms(2) per_distinct_1 per_not_col perp_in_distinct perp_in_id) next assume P3: "\ Col A C B" have P4: "A \ H" by (metis P2 Per_perm Tarski_neutral_dimensionless.l8_7 Tarski_neutral_dimensionless_axioms assms(1) assms(2) col_trivial_1 perp_in_per_2 perp_not_col2) have P5: "Per C H B" using assms(2) perp_in_per_2 by auto have P6: "B \ H" using P1 P2 assms(1) l8_2 l8_7 perp_not_eq_1 by blast have P7: "H A Lt A C \ H C Lt A C" by (metis P1 P2 P4 l11_46 l8_2 perp_distinct) have P8: "C A Lt A B \ C B Lt A B" using P3 assms(1) l11_46 not_col_distincts by blast have P9: "H B Lt B C \ H C Lt B C" by (metis P2 P5 P6 Per_cases l11_46 perp_not_eq_1) have P10: "Bet A H B" proof - have T1: "Col A H B" using assms(2) col_permutation_5 perp_in_col by blast have T2: "A H Le A B" using P7 P8 by (meson lt_comm lt_transitivity nlt__le not_and_lt) have "H B Le A B" by (meson Lt_cases P8 P9 le_transitivity local.le_cases lt__nle) thus ?thesis using T1 T2 l5_12_b by blast qed thus ?thesis by (simp add: P4 P6) qed qed lemma l11_49: assumes "A B C CongA A' B' C'" and "Cong B A B' A'" and "Cong B C B' C'" shows "Cong A C A' C' \ (A \ C \ (B A C CongA B' A' C' \ B C A CongA B' C' A'))" proof - have T1:" Cong A C A' C'" using assms(1) assms(2) assms(3) cong2_conga_cong not_cong_2143 by blast { assume P1: "A \ C" have P2: "A \ B" using CongA_def assms(1) by blast have P3: "C \ B" using CongA_def assms(1) by blast have "B A C Cong3 B' A' C'" by (simp add: Cong3_def T1 assms(2) assms(3)) then have T2: "B A C CongA B' A' C'" using P1 P2 cong3_conga by auto have "B C A Cong3 B' C' A'" using Cong3_def T1 assms(2) assms(3) cong_3_swap_2 by blast then have "B C A CongA B' C' A'" using P1 P3 cong3_conga by auto then have "B A C CongA B' A' C' \ B C A CongA B' C' A'" using T2 by blast } thus ?thesis by (simp add: T1) qed lemma l11_50_1: assumes "\ Col A B C" and "B A C CongA B' A' C'" and "A B C CongA A' B' C'" and "Cong A B A' B'" shows "Cong A C A' C' \ Cong B C B' C' \ A C B CongA A' C' B'" proof - obtain C'' where P1: "B' Out C'' C' \ Cong B' C'' B C" by (metis Col_perm assms(1) assms(3) col_trivial_3 conga_diff56 l6_11_existence) have P2: "B' \ C''" using P1 out_diff1 by auto have P3: "\ Col A' B' C'" using assms(1) assms(3) ncol_conga_ncol by blast have P4: "\ Col A' B' C''" by (meson P1 P2 P3 col_transitivity_1 not_col_permutation_2 out_col) have P5: "Cong A C A' C''" proof - have Q1: "B Out A A" using assms(1) not_col_distincts out_trivial by auto have Q2: "B Out C C" using assms(1) col_trivial_2 out_trivial by force have Q3: "B' Out A' A'" using P3 not_col_distincts out_trivial by auto have Q5: "Cong B A B' A'" using assms(4) not_cong_2143 by blast have "Cong B C B' C''" using P1 not_cong_3412 by blast thus ?thesis using l11_4_1 P1 Q1 Q2 Q3 Q5 assms(3) by blast qed have P6: "B A C Cong3 B' A' C''" using Cong3_def Cong_perm P1 P5 assms(4) by blast have P7: "B A C CongA B' A' C''" by (metis P6 assms(1) cong3_conga not_col_distincts) have P8: "B' A' C' CongA B' A' C''" by (meson P7 assms(2) conga_sym conga_trans) have "B' A' OS C' C''" using Col_perm Out_cases P1 P3 out_one_side by blast then have "A' Out C' C''" using P8 conga_os__out by auto then have "Col A' C' C''" using out_col by auto then have P9: "C' = C''" using Col_perm P1 out_col P3 col_transitivity_1 by blast have T1: "Cong A C A' C'" by (simp add: P5 P9) have T2: "Cong B C B' C'" using Cong_perm P1 P9 by blast then have "A C B CongA A' C' B'" using T1 assms(1) assms(2) assms(4) col_trivial_2 l11_49 by blast thus ?thesis using T1 T2 by blast qed lemma l11_50_2: assumes "\ Col A B C" and "B C A CongA B' C' A'" and "A B C CongA A' B' C'" and "Cong A B A' B'" shows "Cong A C A' C' \ Cong B C B' C' \ C A B CongA C' A' B'" proof - have P1: "A \ B" using assms(1) col_trivial_1 by auto have P2: "B \ C" using assms(1) col_trivial_2 by auto have P3: "A' \ B'" using P1 assms(4) cong_diff by blast have P4: "B' \ C'" using assms(2) conga_diff45 by auto then obtain C'' where P5: "B' Out C'' C' \ Cong B' C'' B C" using P2 l6_11_existence by presburger have P5BIS: "B' \ C''" using P5 out_diff1 by auto have P5A: "Col B' C'' C'" using P5 out_col by auto have P6: "\ Col A' B' C'" using assms(1) assms(3) ncol_conga_ncol by blast { assume "Col A' B' C''" then have "Col B' C'' A'" using not_col_permutation_2 by blast then have "Col B' C' A'" using col_transitivity_1 P5BIS P5A by blast then have "Col A' B' C'" using Col_perm by blast then have False using P6 by auto } then have P7: "\ Col A' B' C''" by blast have P8: "Cong A C A' C''" proof - have "B Out A A" by (simp add: P1 out_trivial) have K1: "B Out C C" using P2 out_trivial by auto have K2: "B' Out A' A'" using P3 out_trivial by auto have "Cong B A B' A'" by (simp add: Cong_perm assms(4)) have "Cong B C B' C''" using Cong_perm P5 by blast thus ?thesis using P5 \Cong B A B' A'\ P1 out_trivial K1 K2 assms(3) l11_4_1 by blast qed have P9: "B C A Cong3 B' C'' A'" using Cong3_def Cong_perm P5 P8 assms(4) by blast then have P10: "B C A CongA B' C'' A'" using assms(1) cong3_conga not_col_distincts by auto have P11: "B' C' A' CongA B' C'' A'" using P9 assms(2) cong3_conga2 conga_sym by blast show ?thesis proof cases assume L1: "C' = C''" then have L2: "Cong A C A' C'" by (simp add: P8) have L3: "Cong B C B' C'" using Cong_perm L1 P5 by blast have "C A B Cong3 C' A' B'" by (simp add: L1 P9 cong_3_swap cong_3_swap_2) then have "C A B CongA C' A' B'" by (metis CongA_def P1 assms(2) cong3_conga) thus ?thesis using L2 L3 by auto next assume R1: "C' \ C''" have R1A: "\ Col C'' C' A'" by (metis P5A P7 R1 col_permutation_2 col_trivial_2 colx) have R1B: "Bet B' C'' C' \ Bet B' C' C''" using Out_def P5 by auto { assume S1: "Bet B' C'' C'" then have S2: "C'' A' C' LtA A' C'' B' \ C'' C' A' LtA A' C'' B'" using P5BIS R1A between_symmetry l11_41 by blast have "B' C' A' CongA C'' C' A'" by (metis P11 R1 Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms S1 bet_out_1 conga_diff45 not_conga_sym out2__conga out_trivial) then have "B' C' A' LtA A' C'' B'" by (meson P11 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.not_conga Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms S2 not_lta_and_conga) then have "Cong A C A' C' \ Cong B C B' C'" by (meson P11 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms not_lta_and_conga) } { assume Z1: "Bet B' C' C''" have Z2: "\ Col C' C'' A'" by (simp add: R1A not_col_permutation_4) have Z3: "C'' Out C' B'" by (simp add: R1 Z1 bet_out_1) have Z4: "C'' Out A' A'" using P7 not_col_distincts out_trivial by blast then have Z4A: "B' C'' A' CongA C' C'' A'" by (simp add: Z3 out2__conga) have Z4B: "B' C'' A' LtA A' C' B'" proof - have Z5: "C' C'' A' CongA B' C'' A'" using Z4A not_conga_sym by blast have Z6: "A' C' B' CongA A' C' B'" using P11 P4 conga_diff2 conga_refl by blast have "C' C'' A' LtA A' C' B'" using P4 Z1 Z2 between_symmetry l11_41 by blast thus ?thesis using Z5 Z6 conga_preserves_lta by auto qed have "B' C'' A' CongA B' C' A'" using P11 not_conga_sym by blast then have "Cong A C A' C' \ Cong B C B' C'" by (meson Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms Z4B not_lta_and_conga) } then have R2: "Cong A C A' C' \ Cong B C B' C'" using R1B \Bet B' C'' C' \ Cong A C A' C' \ Cong B C B' C'\ by blast then have "C A B CongA C' A' B'" using P1 assms(2) l11_49 not_cong_2143 by blast thus ?thesis using R2 by auto qed qed lemma l11_51: assumes "A \ B" and "A \ C" and "B \ C" and "Cong A B A' B'" and "Cong A C A' C'" and "Cong B C B' C'" shows "B A C CongA B' A' C' \ A B C CongA A' B' C' \ B C A CongA B' C' A'" proof - have "B A C Cong3 B' A' C' \ A B C Cong3 A' B' C' \ B C A Cong3 B' C' A'" using Cong3_def Cong_perm assms(4) assms(5) assms(6) by blast thus ?thesis using assms(1) assms(2) assms(3) cong3_conga by auto qed lemma conga_distinct: assumes "A B C CongA D E F" shows "A \ B \ C \ B \ D \ E \ F \ E" using CongA_def assms by auto lemma l11_52: assumes "A B C CongA A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" and "B C Le A C" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have P1: "A \ B" using CongA_def assms(1) by blast have P2: "C \ B" using CongA_def assms(1) by blast have P3: "A' \ B'" using CongA_def assms(1) by blast have P4: "C' \ B'" using assms(1) conga_diff56 by auto have P5: "Cong B A B' A'" proof cases assume P6: "Col A B C" then have P7: "Bet A B C \ Bet B C A \ Bet C A B" using Col_def by blast { assume P8: "Bet A B C" then have "Bet A' B' C'" using assms(1) bet_conga__bet by blast then have "Cong B A B' A'" using P8 assms(2) assms(3) l4_3 not_cong_2143 by blast } { assume P9: "Bet B C A" then have P10: "B' Out A' C'" using Out_cases P2 assms(1) bet_out l11_21_a by blast then have P11: "Bet B' A' C' \ Bet B' C' A'" by (simp add: Out_def) { assume "Bet B' A' C'" then have "Cong B A B' A'" using P3 assms(2) assms(3) assms(4) bet_le_eq l5_6 by blast } { assume "Bet B' C' A'" then have "Cong B A B' A'" using Cong_perm P9 assms(2) assms(3) l2_11_b by blast } then have "Cong B A B' A'" using P11 \Bet B' A' C' \ Cong B A B' A'\ by blast } { assume "Bet C A B" then have "Cong B A B' A'" using P1 assms(4) bet_le_eq between_symmetry by blast } thus ?thesis using P7 \Bet A B C \ Cong B A B' A'\ \Bet B C A \ Cong B A B' A'\ by blast next assume Z1: "\ Col A B C" obtain A'' where Z2: "B' Out A'' A' \ Cong B' A'' B A" using P1 P3 l6_11_existence by force then have Z3: "A' B' C' CongA A'' B' C'" by (simp add: P4 out2__conga out_trivial) have Z4: "A B C CongA A'' B' C'" using Z3 assms(1) not_conga by blast have Z5: "Cong A'' C' A C" using Z2 Z4 assms(3) cong2_conga_cong cong_4321 cong_symmetry by blast have Z6: "A'' B' C' Cong3 A B C" using Cong3_def Cong_perm Z2 Z5 assms(3) by blast have Z7: "Cong A'' C' A' C'" using Z5 assms(2) cong_transitivity by blast have Z8: "\ Col A' B' C'" by (metis Z1 assms(1) ncol_conga_ncol) then have Z9: "\ Col A'' B' C'" by (metis Z2 col_transitivity_1 not_col_permutation_4 out_col out_diff1) { assume Z9A: "A'' \ A'" have Z10: "Bet B' A'' A' \ Bet B' A' A''" using Out_def Z2 by auto { assume Z11: "Bet B' A'' A'" have Z12: "A'' C' B' LtA C' A'' A' \ A'' B' C' LtA C' A'' A'" by (simp add: Z11 Z9 Z9A l11_41) have Z13: "Cong A' C' A'' C'" using Cong_perm Z7 by blast have Z14: "\ Col A'' C' A'" by (metis Col_def Z11 Z9 Z9A col_transitivity_1) have Z15: "C' A'' A' CongA C' A' A'' \ Cong C' A'' C' A'" by (simp add: Z14 l11_44_1) have Z16: "Cong C' A' C' A''" using Cong_perm Z7 by blast then have Z17: "Cong C' A'' C' A'" using Cong_perm by blast then have Z18: "C' A'' A' CongA C' A' A''" by (simp add: Z15) have Z19: "\ Col B' C' A''" using Col_perm Z9 by blast have Z20: "B' A' C' CongA A'' A' C'" by (metis Tarski_neutral_dimensionless.col_conga_col Tarski_neutral_dimensionless_axioms Z11 Z3 Z9 Z9A bet_out_1 col_trivial_3 out2__conga out_trivial) have Z21: "\ Col B' C' A'" using Col_perm Z8 by blast then have Z22: "C' B' A' LtA C' A' B' \ C' A' Lt C' B'" by (simp add: l11_44_2) have "A'' B' C' CongA C' B' A'" using Z3 conga_right_comm not_conga_sym by blast then have U1: "C' B' A' LtA C' A' B'" proof - have f1: "\p pa pb pc pd pe pf pg ph pi pj pk pl pm. \ Tarski_neutral_dimensionless p pa \ \ Tarski_neutral_dimensionless.CongA p pa (pb::'p) pc pd pe pf pg \ \ Tarski_neutral_dimensionless.CongA p pa ph pi pj pk pl pm \ \ Tarski_neutral_dimensionless.LtA p pa pb pc pd ph pi pj \ Tarski_neutral_dimensionless.LtA p pa pe pf pg pk pl pm" by (simp add: Tarski_neutral_dimensionless.conga_preserves_lta) have f2: "C' A'' A' CongA C' A' A''" by (metis Z15 Z17) have f3: "\p pa pb pc pd pe pf pg. \ Tarski_neutral_dimensionless p pa \ \ Tarski_neutral_dimensionless.CongA p pa (pb::'p) pc pd pe pf pg \ Tarski_neutral_dimensionless.CongA p pa pe pf pg pb pc pd" by (metis (no_types) Tarski_neutral_dimensionless.conga_sym) then have "\ C' B' A' LtA C' A'' A' \ A'' B' C' LtA C' A' A''" using f2 f1 by (meson Tarski_neutral_dimensionless_axioms \A'' B' C' CongA C' B' A'\) then have "C' B' A' LtA C' A' B' \ A'' B' C' LtA A'' A' C' \ A'' = B'" using f2 f1 by (metis (no_types) Tarski_neutral_dimensionless.conga_refl Tarski_neutral_dimensionless_axioms Z12 \A'' B' C' CongA C' B' A'\ lta_right_comm) thus ?thesis using f3 f2 f1 by (metis (no_types) Tarski_neutral_dimensionless_axioms Z12 Z20 \A'' B' C' CongA C' B' A'\ lta_right_comm) qed then have Z23: "C' A' Lt C' B'" using Z22 by auto have Z24: "C' A'' Lt C' B'" using Z16 Z23 cong2_lt__lt cong_reflexivity by blast have Z25: "C A Le C B" proof - have Z26: "Cong C' A'' C A" using Z5 not_cong_2143 by blast have "Cong C' B' C B" using assms(3) not_cong_4321 by blast thus ?thesis using l5_6 Z24 Z26 lt__le by blast qed then have Z27: "Cong C A C B" by (simp add: assms(4) le_anti_symmetry le_comm) have "Cong C' A'' C' B'" by (metis Cong_perm Z13 Z27 assms(2) assms(3) cong_transitivity) then have "False" using Z24 cong__nlt by blast then have "Cong B A B' A'" by simp } { assume W1: "Bet B' A' A''" have W2: "A' \ A''" using Z9A by auto have W3: "A' C' B' LtA C' A' A'' \ A' B' C' LtA C' A' A''" using W1 Z8 Z9A l11_41 by blast have W4: "Cong A' C' A'' C'" using Z7 not_cong_3412 by blast have "\ Col A'' C' A'" by (metis Col_def W1 Z8 Z9A col_transitivity_1) then have W6: "C' A'' A' CongA C' A' A'' \ Cong C' A'' C' A'" using l11_44_1 by auto have W7: "Cong C' A' C' A''" using Z7 not_cong_4321 by blast then have W8: "Cong C' A'' C' A'" using W4 not_cong_4321 by blast have W9: "\ Col B' C' A''" by (simp add: Z9 not_col_permutation_1) have W10: "B' A'' C' CongA A' A'' C'" by (metis Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms W1 Z9 Z9A bet_out_1 between_trivial not_col_distincts out2__conga) have W12: "C' B' A'' LtA C' A'' B' \ C' A'' Lt C' B'" by (simp add: W9 l11_44_2) have W12A: "C' B' A'' LtA C' A'' B'" proof - have V1: "A' B' C' CongA C' B' A''" by (simp add: Z3 conga_right_comm) have "A' A'' C' CongA B' A'' C'" by (metis Tarski_neutral_dimensionless.Out_def Tarski_neutral_dimensionless_axioms W1 \\ Col A'' C' A'\ between_equality_2 not_col_distincts or_bet_out out2__conga out_col) then have "C' A' A'' CongA C' A'' B'" by (meson W6 W8 conga_left_comm not_conga not_conga_sym) thus ?thesis using W3 V1 conga_preserves_lta by auto qed then have "C' A'' Lt C' B'" using W12 by auto then have W14: "C' A' Lt C' B'" using W8 cong2_lt__lt cong_reflexivity by blast have W15: "C A Le C B" proof - have Q1: "C' A'' Le C' B'" using W12 W12A lt__le by blast have Q2: "Cong C' A'' C A" using Z5 not_cong_2143 by blast have "Cong C' B' C B" using assms(3) not_cong_4321 by blast thus ?thesis using Q1 Q2 l5_6 by blast qed have "C B Le C A" by (simp add: assms(4) le_comm) then have "Cong C A C B" by (simp add: W15 le_anti_symmetry) then have "Cong C' A' C' B'" by (metis Cong_perm assms(2) assms(3) cong_inner_transitivity) then have "False" using W14 cong__nlt by blast then have "Cong B A B' A'" by simp } then have "Cong B A B' A'" using Z10 \Bet B' A'' A' \ Cong B A B' A'\ by blast } { assume "A'' = A'" then have "Cong B A B' A'" using Z2 not_cong_3412 by blast } thus ?thesis using \A'' \ A' \ Cong B A B' A'\ by blast qed have P6: "A B C Cong3 A' B' C'" using Cong3_def Cong_perm P5 assms(2) assms(3) by blast thus ?thesis using P2 P5 assms(1) assms(3) assms(4) l11_49 le_zero by blast qed lemma l11_53: assumes "Per D C B" and "C \ D" and "A \ B" and "B \ C" and "Bet A B C" shows "C A D LtA C B D \ B D Lt A D" proof - have P1: "C \ A" using assms(3) assms(5) between_identity by blast have P2: "\ Col B A D" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) bet_col bet_col1 col3 col_permutation_4 l8_9) have P3: "A \ D" using P2 col_trivial_2 by blast have P4: "C A D LtA C B D" proof - have P4A: "B D A LtA D B C \ B A D LtA D B C" by (simp add: P2 assms(4) assms(5) l11_41) have P4AA:"A Out B C" using assms(3) assms(5) bet_out by auto have "A Out D D" using P3 out_trivial by auto then have P4B: "C A D CongA B A D" using P4AA by (simp add: out2__conga) then have P4C: "B A D CongA C A D" by (simp add: P4B conga_sym) have "D B C CongA C B D" using assms(1) assms(4) conga_pseudo_refl per_distinct_1 by auto thus ?thesis using P4A P4C conga_preserves_lta by blast qed obtain B' where P5: "C Midpoint B B' \ Cong D B D B'" using Per_def assms(1) by auto have K2: "A \ B'" using Bet_cases P5 assms(4) assms(5) between_equality_2 midpoint_bet by blast { assume "Col B D B'" then have "Col B A D" by (metis Col_cases P5 assms(1) assms(2) assms(4) col2__eq midpoint_col midpoint_distinct_2 per_not_col) then have "False" by (simp add: P2) } then have P6: "\ Col B D B'" by blast then have "D B B' CongA D B' B \ Cong D B D B'" by (simp add: l11_44_1) then have "D B B' CongA D B' B" using P5 by simp { assume K1: "Col A D B'" have "Col B' A B" using Col_def P5 assms(4) assms(5) midpoint_bet outer_transitivity_between by blast then have "Col B' B D" using K1 K2 Col_perm col_transitivity_2 by blast then have "Col B D B'" using Col_perm by blast then have "False" by (simp add: P6) } then have K3B: "\ Col A D B'" by blast then have K4: "D A B' LtA D B' A \ D B' Lt D A" by (simp add: l11_44_2) have K4A: "C A D LtA C B' D" by (metis Midpoint_def P1 P3 P4 P5 P5 P6 assms(2) assms(4) col_trivial_1 cong_reflexivity conga_preserves_lta conga_refl l11_51 not_cong_2134) have "D B' Lt D A" proof - have "D A B' LtA D B' A" proof - have K5A: "A Out D D" using P3 out_trivial by auto have K5AA: "A Out B' C" by (smt K2 Out_def P1 P5 assms(4) assms(5) midpoint_bet outer_transitivity_between2) then have K5: "D A C CongA D A B'" by (simp add: K5A out2__conga) have K6A: "B' Out D D" using K3B not_col_distincts out_trivial by blast have "B' Out A C" by (smt P5 K5AA assms(4) assms(5) between_equality_2 l6_4_2 midpoint_bet midpoint_distinct_2 out_col outer_transitivity_between2) then have K6: "D B' C CongA D B' A" by (simp add: K6A out2__conga) have "D A C LtA D B' C" by (simp add: K4A lta_comm) thus ?thesis using K5 K6 conga_preserves_lta by auto qed thus ?thesis by (simp add: K4) qed thus ?thesis using P4 P5 cong2_lt__lt cong_pseudo_reflexivity not_cong_4312 by blast qed lemma cong2_conga_obtuse__cong_conga2: assumes "Obtuse A B C" and "A B C CongA A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have "B C Le A C" proof cases assume "Col A B C" thus ?thesis by (simp add: assms(1) col_obtuse__bet l5_12_a) next assume "\ Col A B C" thus ?thesis using l11_46 assms(1) lt__le not_col_distincts by auto qed thus ?thesis using l11_52 assms(2) assms(3) assms(4) by blast qed lemma cong2_per2__cong_conga2: assumes "A \ B" and "B \ C" and "Per A B C" and "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A' \ B A C CongA B' A' C' \ B C A CongA B' C' A'" proof - have P1: "B C Le A C \ \ Cong B C A C" using assms(1) assms(2) assms(3) cong__nlt l11_46 lt__le by blast then have "A B C CongA A' B' C'" using assms(2) assms(3) assms(4) assms(5) assms(6) cong_diff cong_inner_transitivity cong_symmetry l11_16 by blast thus ?thesis using P1 assms(5) assms(6) l11_52 by blast qed lemma cong2_per2__cong: assumes "Per A B C" and "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "Cong B A B' A'" proof cases assume "B = C" thus ?thesis using assms(3) assms(4) cong_reverse_identity not_cong_2143 by blast next assume "B \ C" show ?thesis proof cases assume "A = B" thus ?thesis proof - have "Cong A C B' C'" using \A = B\ assms(4) by blast then have "B' = A'" by (meson Cong3_def Per_perm assms(2) assms(3) cong_inner_transitivity cong_pseudo_reflexivity l8_10 l8_7) thus ?thesis using \A = B\ cong_trivial_identity by blast qed next assume "A \ B" show ?thesis proof cases assume "A' = B'" thus ?thesis by (metis Cong3_def Per_perm \A \ B\ assms(1) assms(3) assms(4) cong_inner_transitivity cong_pseudo_reflexivity l8_10 l8_7) next assume "A' \ B'" thus ?thesis using cong2_per2__cong_conga2 \A \ B\ \B \ C\ assms(1) assms(2) assms(3) assms(4) by blast qed qed qed lemma cong2_per2__cong_3: assumes "Per A B C" "Per A' B' C'" and "Cong A C A' C'" and "Cong B C B' C'" shows "A B C Cong3 A' B' C'" by (metis Tarski_neutral_dimensionless.Cong3_def Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) cong2_per2__cong cong_3_swap) lemma cong_lt_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "Cong A B A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" proof cases assume "A = B" thus ?thesis using assms(3) assms(4) cong_reverse_identity by blast next assume "A \ B" show ?thesis proof cases assume "B = C" thus ?thesis by (smt assms(2) assms(3) assms(4) cong2_lt__lt cong_4312 cong_diff cong_reflexivity l11_46 lt_diff) next assume P0: "B \ C" have "B C Lt B' C'" by (simp add: assms(4)) then have R1: "B C Le B' C' \ \ Cong B C B' C'" by (simp add: Lt_def) then obtain C0 where P1: "Bet B' C0 C' \ Cong B C B' C0" using Le_def by auto then have P2: "Per A' B' C0" by (metis Col_def Per_cases assms(2) bet_out_1 col_col_per_per col_trivial_1 l8_5 out_diff2) have "C0 A' Lt C' A'" using l11_53 by (metis P1 P2 R1 P0 bet__lt2313 between_symmetry cong_diff) then have P3: "A' C0 Lt A' C'" using Lt_cases by blast have P4: "Cong A' C0 A C" using P1 P2 assms(1) assms(3) l10_12 not_cong_3412 by blast have "Cong A' C' A' C'" by (simp add: cong_reflexivity) thus ?thesis using cong2_lt__lt P3 P4 by blast qed qed lemma cong_le_per2__le: assumes "Per A B C" and "Per A' B' C'" and "Cong A B A' B'" and "B C Le B' C'" shows "A C Le A' C'" proof cases assume "Cong B C B' C'" thus ?thesis using assms(1) assms(2) assms(3) cong__le l10_12 by blast next assume "\ Cong B C B' C'" then have "B C Lt B' C'" using Lt_def assms(4) by blast thus ?thesis using assms(1) assms(2) assms(3) cong_lt_per2__lt lt__le by auto qed lemma lt2_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "A B Lt A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" proof - have P2: "B A Lt B' A'" by (simp add: assms(3) lt_comm) have P3: "B C Le B' C' \ \ Cong B C B' C'" using assms(4) cong__nlt lt__le by auto then obtain C0 where P4: "Bet B' C0 C' \ Cong B C B' C0" using Le_def by auto have P4A: "B' \ C'" using assms(4) lt_diff by auto have "Col B' C' C0" using P4 bet_col not_col_permutation_5 by blast then have P5: "Per A' B' C0" using assms(2) P4A per_col by blast have P6: "A C Lt A' C0" by (meson P2 P4 P5 assms(1) cong_lt_per2__lt l8_2 lt_comm not_cong_2143) have "B' C0 Lt B' C'" by (metis P4 assms(4) bet__lt1213 cong__nlt) then have "A' C0 Lt A' C'" using P5 assms(2) cong_lt_per2__lt cong_reflexivity by blast thus ?thesis using P6 lt_transitivity by blast qed lemma le_lt_per2__lt: assumes "Per A B C" and "Per A' B' C'" and "A B Le A' B'" and "B C Lt B' C'" shows "A C Lt A' C'" using Lt_def assms(1) assms(2) assms(3) assms(4) cong_lt_per2__lt lt2_per2__lt by blast lemma le2_per2__le: assumes "Per A B C" and "Per A' B' C'" and "A B Le A' B'" and "B C Le B' C'" shows "A C Le A' C'" proof cases assume "Cong B C B' C'" thus ?thesis by (meson Per_cases Tarski_neutral_dimensionless.cong_le_per2__le Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) le_comm not_cong_2143) next assume "\ Cong B C B' C'" then have "B C Lt B' C'" by (simp add: Lt_def assms(4)) thus ?thesis using assms(1) assms(2) assms(3) le_lt_per2__lt lt__le by blast qed lemma cong_lt_per2__lt_1: assumes "Per A B C" and "Per A' B' C'" and "A B Lt A' B'" and "Cong A C A' C'" shows "B' C' Lt B C" by (meson Gt_def assms(1) assms(2) assms(3) assms(4) cong2_per2__cong cong_4321 cong__nlt cong_symmetry lt2_per2__lt or_lt_cong_gt) lemma symmetry_preserves_conga: assumes "A \ B" and "C \ B" and "M Midpoint A A'" and "M Midpoint B B'" and "M Midpoint C C'" shows "A B C CongA A' B' C'" by (metis Mid_perm assms(1) assms(2) assms(3) assms(4) assms(5) conga_trivial_1 l11_51 l7_13 symmetric_point_uniqueness) lemma l11_57: assumes "A A' OS B B'" and "Per B A A'" and "Per B' A' A" and "A A' OS C C'" and "Per C A A'" and "Per C' A' A" shows "B A C CongA B' A' C'" proof - obtain M where P1: "M Midpoint A A'" using midpoint_existence by auto obtain B'' where P2: "M Midpoint B B''" using symmetric_point_construction by auto obtain C'' where P3: "M Midpoint C C''" using symmetric_point_construction by auto have P4: "\ Col A A' B" using assms(1) col123__nos by auto have P5: "\ Col A A' C" using assms(4) col123__nos by auto have P6: "B A C CongA B'' A' C''" by (metis P1 P2 P3 assms(1) assms(4) os_distincts symmetry_preserves_conga) have "B'' A' C'' CongA B' A' C'" proof - have "B \ M" using P1 P4 midpoint_col not_col_permutation_2 by blast then have P7: "\ Col B'' A A'" using Mid_cases P1 P2 P4 mid_preserves_col not_col_permutation_3 by blast have K3: "Bet B'' A' B'" proof - have "Per B'' A' A" using P1 P2 assms(2) per_mid_per by blast have "Col B B'' M \ Col A A' M" using P1 P2 midpoint_col not_col_permutation_2 by blast then have "Coplanar B A A' B''" using Coplanar_def by auto then have "Coplanar A B' B'' A'" by (meson assms(1) between_trivial2 coplanar_trans_1 ncoplanar_perm_4 ncoplanar_perm_8 one_side_chara os__coplanar) then have P8: "Col B' B'' A'" using cop_per2__col P1 P2 P7 assms(2) assms(3) not_col_distincts per_mid_per by blast have "A A' TS B B''" using P1 P2 P4 mid_two_sides by auto then have "A' A TS B'' B'" using assms(1) invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using Col_cases P8 col_two_sides_bet by blast qed have "\ Col C'' A A'" by (smt Col_def P1 P3 P5 l7_15 l7_2 not_col_permutation_5) have "Bet C'' A' C'" proof - have Z2: "Col C' C'' A'" proof - have "Col C C'' M \ Col A A' M" using P1 P3 col_permutation_1 midpoint_col by blast then have "Coplanar C A A' C''" using Coplanar_def by blast then have Z1: "Coplanar A C' C'' A'" by (meson assms(4) between_trivial2 coplanar_trans_1 ncoplanar_perm_4 ncoplanar_perm_8 one_side_chara os__coplanar) have "Per C'' A' A" using P1 P3 assms(5) per_mid_per by blast thus ?thesis using Z1 P5 assms(6) col_trivial_1 cop_per2__col by blast qed have "A A' TS C C''" using P1 P3 P5 mid_two_sides by auto then have "A' A TS C'' C'" using assms(4) invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using Col_cases Z2 col_two_sides_bet by blast qed thus ?thesis by (metis P6 K3 assms(1) assms(4) conga_diff45 conga_diff56 l11_14 os_distincts) qed thus ?thesis using P6 conga_trans by blast qed lemma cop3_orth_at__orth_at: assumes "\ Col D E F" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar A B C F" and "X OrthAt A B C U V" shows "X OrthAt D E F U V" proof - have P1: "\ Col A B C \ Coplanar A B C X" using OrthAt_def assms(5) by blast then have P2: "Coplanar D E F X" using assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast { fix M assume "Coplanar A B C M" then have "Coplanar D E F M" using P1 assms(2) assms(3) assms(4) coplanar_pseudo_trans by blast } have T1: "U \ V" using OrthAt_def assms(5) by blast have T2: "Col U V X" using OrthAt_def assms(5) by auto { fix P Q assume P7: "Coplanar D E F P \ Col U V Q" then have "Coplanar A B C P" by (meson \\M. Coplanar A B C M \ Coplanar D E F M\ assms(1) assms(2) assms(3) assms(4) l9_30) then have "Per P X Q" using P7 OrthAt_def assms(5) by blast } thus ?thesis using assms(1) by (simp add: OrthAt_def P2 T1 T2) qed lemma col2_orth_at__orth_at: assumes "U \ V" and "Col P Q U" and "Col P Q V" and "X OrthAt A B C P Q" shows "X OrthAt A B C U V" proof - have "Col P Q X" using OrthAt_def assms(4) by auto then have "Col U V X" by (metis OrthAt_def assms(2) assms(3) assms(4) col3) thus ?thesis using OrthAt_def assms(1) assms(2) assms(3) assms(4) colx by presburger qed lemma col_orth_at__orth_at: assumes "U \ W" and "Col U V W" and "X OrthAt A B C U V" shows "X OrthAt A B C U W" using assms(1) assms(2) assms(3) col2_orth_at__orth_at col_trivial_3 by blast lemma orth_at_symmetry: assumes "X OrthAt A B C U V" shows "X OrthAt A B C V U" by (metis assms col2_orth_at__orth_at col_trivial_2 col_trivial_3) lemma orth_at_distincts: assumes "X OrthAt A B C U V" shows "A \ B \ B \ C \ A \ C \ U \ V" using OrthAt_def assms not_col_distincts by fastforce lemma orth_at_chara: "X OrthAt A B C X P \ (\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P)))" proof - { assume "X OrthAt A B C X P" then have "\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P))" using OrthAt_def col_trivial_2 by auto } { assume T1: "\ Col A B C \ X \ P \ Coplanar A B C X \ (\ D.(Coplanar A B C D \ Per D X P))" { fix P0 Q assume "Coplanar A B C P0 \ Col X P Q" then have "Per P0 X Q" using T1 OrthAt_def per_col by auto } then have "X OrthAt A B C X P" by (simp add: T1 \\Q P0. Coplanar A B C P0 \ Col X P Q \ Per P0 X Q\ Tarski_neutral_dimensionless.OrthAt_def Tarski_neutral_dimensionless_axioms col_trivial_3) } thus ?thesis using \X OrthAt A B C X P \ \ Col A B C \ X \ P \ Coplanar A B C X \ (\D. Coplanar A B C D \ Per D X P)\ by blast qed lemma cop3_orth__orth: assumes "\ Col D E F" and "Coplanar A B C D" and "Coplanar A B C E" and "Coplanar A B C F" and "A B C Orth U V" shows "D E F Orth U V" using Orth_def assms(1) assms(2) assms(3) assms(4) assms(5) cop3_orth_at__orth_at by blast lemma col2_orth__orth: assumes "U \ V" and "Col P Q U" and "Col P Q V" and "A B C Orth P Q" shows "A B C Orth U V" by (meson Orth_def Tarski_neutral_dimensionless.col2_orth_at__orth_at Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4)) lemma col_orth__orth: assumes "U \ W" and "Col U V W" and "A B C Orth U V" shows "A B C Orth U W" by (meson assms(1) assms(2) assms(3) col2_orth__orth col_trivial_3) lemma orth_symmetry: assumes "A B C Orth U V" shows "A B C Orth V U" by (meson Orth_def assms orth_at_symmetry) lemma orth_distincts: assumes "A B C Orth U V" shows "A \ B \ B \ C \ A \ C \ U \ V" using Orth_def assms orth_at_distincts by blast lemma col_cop_orth__orth_at: assumes "A B C Orth U V" and "Coplanar A B C X" and "Col U V X" shows "X OrthAt A B C U V" proof - obtain Y where P1: "\ Col A B C \ U \ V \ Coplanar A B C Y \ Col U V Y \ (\ P Q. (Coplanar A B C P \ Col U V Q) \ Per P Y Q)" by (metis OrthAt_def Tarski_neutral_dimensionless.Orth_def Tarski_neutral_dimensionless_axioms assms(1)) then have P2: "X = Y" using assms(2) assms(3) per_distinct_1 by blast { fix P Q assume "Coplanar A B C P \ Col U V Q" then have "Per P X Q" using P1 P2 by auto } thus ?thesis using OrthAt_def Orth_def assms(1) assms(2) assms(3) by auto qed lemma l11_60_aux: assumes "\ Col A B C" and "Cong A P A Q" and "Cong B P B Q" and "Cong C P C Q" and "Coplanar A B C D" shows "Cong D P D Q" proof - obtain M where P1: "Bet P M Q \ Cong P M M Q" by (meson Midpoint_def Tarski_neutral_dimensionless.midpoint_existence Tarski_neutral_dimensionless_axioms) obtain X where P2: " (Col A B X \ Col C D X) \ (Col A C X \ Col B D X) \ (Col A D X \ Col B C X)" using assms(5) Coplanar_def by auto { assume "Col A B X \ Col C D X" then have "Cong D P D Q" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_5) } { assume "Col A C X \ Col B D X" then have "Cong D P D Q" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_5) } { assume "Col A D X \ Col B C X" then have "Cong D P D Q" by (smt assms(1) assms(2) assms(3) assms(4) l4_17 not_col_distincts not_col_permutation_1) } thus ?thesis using P2 \Col A B X \ Col C D X \ Cong D P D Q\ \Col A C X \ Col B D X \ Cong D P D Q\ by blast qed lemma l11_60: assumes "\ Col A B C" and "Per A D P" and "Per B D P" and "Per C D P" and "Coplanar A B C E" shows "Per E D P" by (meson Per_def assms(1) assms(2) assms(3) assms(4) assms(5) l11_60_aux per_double_cong) lemma l11_60_bis: assumes "\ Col A B C" and "D \ P" and "Coplanar A B C D" and "Per A D P" and "Per B D P" and "Per C D P" shows "D OrthAt A B C D P" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l11_60 orth_at_chara by auto lemma l11_61: assumes "A \ A'" and "A \ B" and "A \ C" and "Coplanar A A' B B'" and "Per B A A'" and "Per B' A' A" and "Coplanar A A' C C'" and "Per C A A'" and "Per B A C" shows "Per B' A' C'" proof - have P1: "\ Col C A A'" using assms(1) assms(3) assms(8) per_col_eq by blast obtain C'' where P2: "A A' Perp C'' A' \ A A' OS C C''" using l10_15 using Col_perm P1 col_trivial_2 by blast have P6: "B' \ A" using assms(1) assms(6) per_distinct by blast have P8: "\ Col A' A C''" using P2 not_col_permutation_4 one_side_not_col124 by blast have P9: "Per A' A' B'" by (simp add: l8_2 l8_5) have P10: "Per A A' B'" by (simp add: assms(6) l8_2) { fix B' assume "A A' OS B B' \ Per B' A' A" then have "B A C CongA B' A' C''" using l11_17 by (meson P2 Perp_cases Tarski_neutral_dimensionless.l11_57 Tarski_neutral_dimensionless_axioms assms(5) assms(8) perp_per_1) then have "Per B' A' C''" using assms(9) l11_17 by blast } then have Q1: "\ B'. (A A' OS B B' \ Per B' A' A) \ Per B' A' C''" by simp { fix B' assume P12: "Coplanar A A' B B' \ Per B' A' A \ B' \ A" have "Per B' A' C''" proof cases assume "B' = A'" thus ?thesis by (simp add: Per_perm l8_5) next assume P13: "B' \ A'" have P14: "\ Col B' A' A" using P12 P13 assms(1) l8_9 by auto have P15: "\ Col B A A'" using assms(1) assms(2) assms(5) per_not_col by auto then have Z1: "A A' TS B B' \ A A' OS B B'" using P12 P14 cop__one_or_two_sides not_col_permutation_5 by blast { assume "A A' OS B B'" then have "Per B' A' C''" by (simp add: P12 \\B'a. A A' OS B B'a \ Per B'a A' A \ Per B'a A' C''\) } { assume Q2: "A A' TS B B'" obtain B'' where Z2: "Bet B' A' B'' \ Cong A' B'' A' B'" using segment_construction by blast have "B' \ B''" using P13 Z2 bet_neq12__neq by blast then have Z4: "A' \ B''" using Z2 cong_diff_4 by blast then have "A A' TS B'' B'" by (meson TS_def Z2 Q2 bet__ts invert_two_sides l9_2 not_col_permutation_1) then have Z5: "A A' OS B B''" using Q2 l9_8_1 by auto have "Per B'' A' A" using P12 P13 Z2 bet_col col_per2__per l8_2 l8_5 by blast then have "Per C'' A' B''" using l8_2 Q1 Z5 by blast then have "Per B' A' C''" by (metis Col_def Per_perm Tarski_neutral_dimensionless.l8_3 Tarski_neutral_dimensionless_axioms Z2 Z4) } thus ?thesis using Z1 using \A A' OS B B' \ Per B' A' C''\ by blast qed } then have "\ B'. (Coplanar A A' B B' \ Per B' A' A \ B' \ A) \ Per B' A' C''" by simp then have "Per B' A' C''" using P6 assms(4) assms(6) by blast then have P11: "Per C'' A' B'" using Per_cases by auto have "Coplanar A' A C'' C'" by (meson P1 P2 assms(7) coplanar_trans_1 ncoplanar_perm_6 ncoplanar_perm_8 os__coplanar) thus ?thesis using P8 P9 P10 P11 l8_2 l11_60 by blast qed lemma l11_61_bis: assumes "D OrthAt A B C D P" and "D E Perp E Q" and "Coplanar A B C E" and "Coplanar D E P Q" shows "E OrthAt A B C E Q" proof - have P4: "D \ E" using assms(2) perp_not_eq_1 by auto have P5: "E \ Q" using assms(2) perp_not_eq_2 by auto have "\ D'. (D E Perp D' D \ Coplanar A B C D')" proof - obtain F where T1: "Coplanar A B C F \ \ Col D E F" using P4 ex_ncol_cop by blast obtain D' where T2: "D E Perp D' D \ Coplanar D E F D'" using P4 ex_perp_cop by blast have "Coplanar A B C D'" proof - have T3A: "\ Col A B C" using OrthAt_def assms(1) by auto have T3B: "Coplanar A B C D" using OrthAt_def assms(1) by blast then have T4: "Coplanar D E F A" by (meson T1 T3A assms(3) coplanar_pseudo_trans ncop_distincts) have T5: "Coplanar D E F B" using T1 T3A T3B assms(3) coplanar_pseudo_trans ncop_distincts by blast have "Coplanar D E F C" using T1 T3A T3B assms(3) coplanar_pseudo_trans ncop_distincts by blast thus ?thesis using T1 T2 T4 T5 coplanar_pseudo_trans by blast qed thus ?thesis using T2 by auto qed then obtain D' where R1: "D E Perp D' D \ Coplanar A B C D'" by auto then have R2: "D \ D'" using perp_not_eq_2 by blast { fix M assume R3: "Coplanar A B C M" have "Col D P P" by (simp add: col_trivial_2) then have "Per E D P" using assms(1) assms(3) orth_at_chara by auto then have R4: "Per P D E" using l8_2 by auto have R5: "Per Q E D" using Perp_cases assms(2) perp_per_2 by blast have R6: "Coplanar D E D' M" proof - have S1: "\ Col A B C" using OrthAt_def assms(1) by auto have "Coplanar A B C D" using OrthAt_def assms(1) by auto thus ?thesis using S1 assms(3) R1 R3 coplanar_pseudo_trans by blast qed have R7: "Per D' D E" using Perp_cases R1 perp_per_1 by blast have "Per D' D P" using R1 assms(1) orth_at_chara by blast then have "Per P D D'" using Per_cases by blast then have "Per Q E M" using l11_61 R4 R5 R6 R7 OrthAt_def P4 R2 assms(1) assms(4) by blast then have "Per M E Q" using l8_2 by auto } { fix P0 Q0 assume "Coplanar A B C P0 \ Col E Q Q0" then have "Per P0 E Q0" using P5 \\M. Coplanar A B C M \ Per M E Q\ per_col by blast } thus ?thesis using OrthAt_def P5 assms(1) assms(3) col_trivial_3 by auto qed lemma l11_62_unicity: assumes "Coplanar A B C D" and "Coplanar A B C D'" and "\ E. Coplanar A B C E \ Per E D P" and "\ E. Coplanar A B C E \ Per E D' P" shows "D = D'" by (metis assms(1) assms(2) assms(3) assms(4) l8_8 not_col_distincts per_not_colp) lemma l11_62_unicity_bis: assumes "X OrthAt A B C X U" and "Y OrthAt A B C Y U" shows "X = Y" proof - have P1: "Coplanar A B C X" using assms(1) orth_at_chara by blast have P2: "Coplanar A B C Y" using assms(2) orth_at_chara by blast { fix E assume "Coplanar A B C E" then have "Per E X U" using OrthAt_def assms(1) col_trivial_2 by auto } { fix E assume "Coplanar A B C E" then have "Per E Y U" using assms(2) orth_at_chara by auto } thus ?thesis by (meson P1 P2 \\E. Coplanar A B C E \ Per E X U\ l8_2 l8_7) qed lemma orth_at2__eq: assumes "X OrthAt A B C U V" and "Y OrthAt A B C U V" shows "X = Y" proof - have P1: "Coplanar A B C X" using assms(1) by (simp add: OrthAt_def) have P2: "Coplanar A B C Y" using OrthAt_def assms(2) by auto { fix E assume "Coplanar A B C E" then have "Per E X U" using OrthAt_def assms(1) col_trivial_3 by auto } { fix E assume "Coplanar A B C E" then have "Per E Y U" using OrthAt_def assms(2) col_trivial_3 by auto } thus ?thesis by (meson P1 P2 Per_perm \\E. Coplanar A B C E \ Per E X U\ l8_7) qed lemma col_cop_orth_at__eq: assumes "X OrthAt A B C U V" and "Coplanar A B C Y" and "Col U V Y" shows "X = Y" proof - have "Y OrthAt A B C U V" using Orth_def assms(1) assms(2) assms(3) col_cop_orth__orth_at by blast thus ?thesis using assms(1) orth_at2__eq by auto qed lemma orth_at__ncop1: assumes "U \ X" and "X OrthAt A B C U V" shows "\ Coplanar A B C U" using assms(1) assms(2) col_cop_orth_at__eq not_col_distincts by blast lemma orth_at__ncop2: assumes "V \ X" and "X OrthAt A B C U V" shows "\ Coplanar A B C V" using assms(1) assms(2) col_cop_orth_at__eq not_col_distincts by blast lemma orth_at__ncop: assumes "X OrthAt A B C X P" shows "\ Coplanar A B C P" by (metis assms orth_at__ncop2 orth_at_distincts) lemma l11_62_existence: "\ D. (Coplanar A B C D \ (\ E. (Coplanar A B C E \ Per E D P)))" proof cases assume "Coplanar A B C P" thus ?thesis using l8_5 by auto next assume P1: "\ Coplanar A B C P" then have P2: "\ Col A B C" using ncop__ncol by auto have "\ Col A B P" using P1 ncop__ncols by auto then obtain D0 where P4: "Col A B D0 \ A B Perp P D0" using l8_18_existence by blast have P5: "Coplanar A B C D0" using P4 ncop__ncols by auto have "A \ B" using P2 not_col_distincts by auto then obtain D1 where P10: "A B Perp D1 D0 \ Coplanar A B C D1" using ex_perp_cop by blast have P11: "\ Col A B D1" using P10 P4 perp_not_col2 by blast { fix D assume "Col D0 D1 D" then have "Coplanar A B C D" by (metis P10 P5 col_cop2__cop perp_not_eq_2) } obtain A0 where P11: "A \ A0 \ B \ A0 \ D0 \ A0 \ Col A B A0" using P4 diff_col_ex3 by blast have P12: "Coplanar A B C A0" using P11 ncop__ncols by blast have P13: "Per P D0 A0" using l8_16_1 P11 P4 by blast show ?thesis proof cases assume Z1: "Per P D0 D1" { fix E assume R1: "Coplanar A B C E" have R2: "\ Col A0 D1 D0" by (metis P10 P11 P4 col_permutation_5 colx perp_not_col2) have R3: "Per A0 D0 P" by (simp add: P13 l8_2) have R4: "Per D1 D0 P" by (simp add: Z1 l8_2) have R5: "Per D0 D0 P" by (simp add: l8_2 l8_5) have "Coplanar A0 D1 D0 E" using R1 P2 P12 P10 P5 coplanar_pseudo_trans by blast then have "Per E D0 P" using l11_60 R2 R3 R4 R5 by blast } thus ?thesis using P5 by auto next assume S1: "\ Per P D0 D1" { assume S2: "Col D0 D1 P" have "\ D. Col D0 D1 D \ Coplanar A B C D" by (simp add: \\Da. Col D0 D1 Da \ Coplanar A B C Da\) then have "False" using P1 S2 by blast } then have S2A: "\ Col D0 D1 P" by blast then obtain D where S3: "Col D0 D1 D \ D0 D1 Perp P D" using l8_18_existence by blast have S4: "Coplanar A B C D" by (simp add: S3 \\Da. Col D0 D1 Da \ Coplanar A B C Da\) { fix E assume S5: "Coplanar A B C E" have S6: "D \ D0" using S1 S3 l8_2 perp_per_1 by blast have S7: "Per D0 D P" by (metis Perp_cases S3 S6 perp_col perp_per_1) have S8: "Per D D0 A0" proof - have V4: "D0 \ D1" using P10 perp_not_eq_2 by blast have V6: "Per A0 D0 D1" using P10 P11 P4 l8_16_1 l8_2 by blast thus ?thesis using S3 V4 V6 l8_2 per_col by blast qed have S9: "Per A0 D P" proof - obtain A0' where W1: "D Midpoint A0 A0'" using symmetric_point_construction by auto obtain D0' where W2: "D Midpoint D0 D0'" using symmetric_point_construction by auto have "Cong P A0 P A0'" proof - have V3: "Cong P D0 P D0'" using S7 W2 l8_2 per_double_cong by blast have V4: "Cong D0 A0 D0' A0'" using W1 W2 cong_4321 l7_13 by blast have "Per P D0' A0'" proof - obtain P' where V5: "D Midpoint P P'" using symmetric_point_construction by blast have "Per P' D0 A0" proof - have "\ Col P D D0" by (metis S2A S3 S6 col2__eq col_permutation_1) thus ?thesis by (metis (full_types) P13 S3 S8 V5 S2A col_per2__per midpoint_col) qed thus ?thesis using midpoint_preserves_per V5 Mid_cases W1 W2 by blast qed thus ?thesis using l10_12 P13 V3 V4 by blast qed thus ?thesis using Per_def Per_perm W1 by blast qed have S13: "Per D D P" using Per_perm l8_5 by blast have S14: "\ Col D0 A0 D" using P11 S7 S9 per_not_col Col_perm S6 S8 by blast have "Coplanar A B C D" using S4 by auto then have "Coplanar D0 A0 D E" using P12 P2 P5 S5 coplanar_pseudo_trans by blast then have "Per E D P" using S13 S14 S7 S9 l11_60 by blast } thus ?thesis using S4 by blast qed qed lemma l11_62_existence_bis: assumes "\ Coplanar A B C P" shows "\ X. X OrthAt A B C X P" proof - obtain X where P1: "Coplanar A B C X \ (\ E. Coplanar A B C E \ Per E X P)" using l11_62_existence by blast then have P2: "X \ P" using assms by auto have P3: "\ Col A B C" using assms ncop__ncol by auto thus ?thesis using P1 P2 P3 orth_at_chara by auto qed lemma l11_63_aux: assumes "Coplanar A B C D" and "D \ E" and "E OrthAt A B C E P" shows "\ Q. (D E OS P Q \ A B C Orth D Q)" proof - have P1: "\ Col A B C" using OrthAt_def assms(3) by blast have P2: "E \ P" using OrthAt_def assms(3) by blast have P3: "Coplanar A B C E" using OrthAt_def assms(3) by blast have P4: "\ P0 Q. (Coplanar A B C P0 \ Col E P Q) \ Per P0 E Q" using OrthAt_def assms(3) by blast have P5: "\ Coplanar A B C P" using assms(3) orth_at__ncop by auto have P6: "Col D E D" by (simp add: col_trivial_3) have "\ Col D E P" using P3 P5 assms(1) assms(2) col_cop2__cop by blast then obtain Q where P6: "D E Perp Q D \ D E OS P Q" using P6 l10_15 by blast have "A B C Orth D Q" proof - obtain F where P7: "Coplanar A B C F \ \ Col D E F" using assms(2) ex_ncol_cop by blast obtain D' where P8: "D E Perp D' D \ Coplanar D E F D'" using assms(2) ex_perp_cop by presburger have P9: "\ Col D' D E" using P8 col_permutation_1 perp_not_col by blast have P10: "Coplanar D E F A" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) have P11: "Coplanar D E F B" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) have P12: "Coplanar D E F C" by (meson P1 P3 P7 assms(1) coplanar_pseudo_trans ncop_distincts) then have "D OrthAt A B C D Q" proof - have "Per D' D Q" proof - obtain E' where Y1: "D E Perp E' E \ Coplanar D E F E'" using assms(2) ex_perp_cop by blast have Y2: "E \ E'" using Y1 perp_distinct by auto have Y5: "Coplanar E D E' D'" by (meson P7 P8 Y1 coplanar_perm_12 coplanar_perm_7 coplanar_trans_1 not_col_permutation_2) have Y6: "Per E' E D" by (simp add: Perp_perm Tarski_neutral_dimensionless.perp_per_2 Tarski_neutral_dimensionless_axioms Y1) have Y7: "Per D' D E" using P8 col_trivial_2 col_trivial_3 l8_16_1 by blast have Y8: "Coplanar E D P Q" using P6 ncoplanar_perm_6 os__coplanar by blast have Y9: "Per P E D" using P4 using assms(1) assms(3) l8_2 orth_at_chara by blast have Y10: "Coplanar A B C E'" using P10 P11 P12 P7 Y1 coplanar_pseudo_trans by blast then have Y11: "Per E' E P" using P4 Y10 col_trivial_2 by auto have "E \ D" using assms(2) by blast thus ?thesis using l11_61 Y2 assms(2) P2 Y5 Y6 Y7 Y8 Y9 Y10 Y11 by blast qed then have X1: "D OrthAt D' D E D Q" using l11_60_bis by (metis OS_def P6 P9 Per_perm TS_def Tarski_neutral_dimensionless.l8_5 Tarski_neutral_dimensionless_axioms col_trivial_3 invert_one_side ncop_distincts perp_per_1) have X3: "Coplanar D' D E A" using P10 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast have X4: "Coplanar D' D E B" using P11 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast have "Coplanar D' D E C" using P12 P7 P8 coplanar_perm_14 coplanar_trans_1 not_col_permutation_3 by blast thus ?thesis using X1 P1 X3 X4 cop3_orth_at__orth_at by blast qed thus ?thesis using Orth_def by blast qed thus ?thesis using P6 by blast qed lemma l11_63_existence: assumes "Coplanar A B C D" and "\ Coplanar A B C P" shows "\ Q. A B C Orth D Q" using Orth_def assms(1) assms(2) l11_62_existence_bis l11_63_aux by fastforce lemma l8_21_3: assumes "Coplanar A B C D" and "\ Coplanar A B C X" shows "\ P T. (A B C Orth D P \ Coplanar A B C T \ Bet X T P)" proof - obtain E where P1: "E OrthAt A B C E X" using assms(2) l11_62_existence_bis by blast thus ?thesis proof cases assume P2: "D = E" obtain Y where P3: "Bet X D Y \ Cong D Y D X" using segment_construction by blast have P4: "D \ X" using assms(1) assms(2) by auto have P5: "A B C Orth D X" using Orth_def P1 P2 by auto have P6: "D \ Y" using P3 P4 cong_reverse_identity by blast have "Col D X Y" using Col_def Col_perm P3 by blast then have "A B C Orth D Y" using P5 P6 col_orth__orth by auto thus ?thesis using P3 assms(1) by blast next assume K1: "D \ E" then obtain P' where P7: "D E OS X P' \ A B C Orth D P'" using P1 assms(1) l11_63_aux by blast have P8: "\ Col A B C" using assms(2) ncop__ncol by auto have P9: "E \ X" using P7 os_distincts by auto have P10: "\ P Q. (Coplanar A B C P \ Col E X Q) \ Per P E Q" using OrthAt_def P1 by auto have P11: "D OrthAt A B C D P'" by (simp add: P7 assms(1) col_cop_orth__orth_at col_trivial_3) have P12: "D \ P'" using P7 os_distincts by auto have P13: "\ Coplanar A B C P'" using P11 orth_at__ncop by auto have P14: "\ P Q. (Coplanar A B C P \ Col D P' Q) \ Per P D Q" using OrthAt_def P11 by auto obtain P where P15: "Bet P' D P \ Cong D P D P'" using segment_construction by blast have P16: "D E TS X P" proof - have P16A: "D E OS P' X" using P7 one_side_symmetry by blast then have "D E TS P' P" by (metis P12 P15 Tarski_neutral_dimensionless.bet__ts Tarski_neutral_dimensionless_axioms cong_diff_3 one_side_not_col123) thus ?thesis using l9_8_2 P16A by blast qed obtain T where P17: "Col T D E \ Bet X T P" using P16 TS_def by blast have P18: "D \ P" using P16 ts_distincts by blast have "Col D P' P" using Col_def Col_perm P15 by blast then have "A B C Orth D P" using P18 P7 col_orth__orth by blast thus ?thesis using col_cop2__cop by (meson P1 P17 Tarski_neutral_dimensionless.orth_at_chara Tarski_neutral_dimensionless_axioms K1 assms(1) col_permutation_1) qed qed lemma mid2_orth_at2__cong: assumes "X OrthAt A B C X P" and "Y OrthAt A B C Y Q" and "X Midpoint P P'" and "Y Midpoint Q Q'" shows "Cong P Q P' Q'" proof - have Q1: "\ Col A B C" using assms(1) col__coplanar orth_at__ncop by blast have Q2: "X \ P" using assms(1) orth_at_distincts by auto have Q3: "Coplanar A B C X" using OrthAt_def assms(1) by auto have Q4: "\ P0 Q. (Coplanar A B C P0 \ Col X P Q) \ Per P0 X Q" using OrthAt_def assms(1) by blast have Q5: "Y \ P" by (metis assms(1) assms(2) orth_at__ncop2 orth_at_chara) have Q6: "Coplanar A B C Y" using OrthAt_def assms(2) by blast have Q7: "\ P Q0. (Coplanar A B C P \ Col Y Q Q0) \ Per P Y Q0" using OrthAt_def assms(2) by blast obtain Z where P1: "Z Midpoint X Y" using midpoint_existence by auto obtain R where P2: "Z Midpoint P R" using symmetric_point_construction by auto obtain R' where P3: "Z Midpoint P' R'" using symmetric_point_construction by auto have T1: "Coplanar A B C Z" using P1 Q3 Q6 bet_cop2__cop midpoint_bet by blast then have "Per Z X P" using Q4 assms(1) orth_at_chara by auto then have T2: "Cong Z P Z P'" using assms(3) per_double_cong by blast have T3: "Cong R Z R' Z" by (metis Cong_perm Midpoint_def P2 P3 T2 cong_transitivity) have T4: "Cong R Q R' Q'" by (meson P1 P2 P3 assms(3) assms(4) l7_13 not_cong_4321 symmetry_preserves_midpoint) have "Per Z Y Q" using Q7 T1 assms(2) orth_at_chara by auto then have T5: "Cong Z Q Z Q'" using assms(4) per_double_cong by auto have "R \ Z" by (metis P2 P3 Q2 T3 assms(3) cong_diff_3 l7_17_bis midpoint_not_midpoint) thus ?thesis using P2 P3 T2 T3 T4 T5 five_segment l7_2 midpoint_bet by blast qed lemma orth_at2_tsp__ts: assumes "P \ Q" and "P OrthAt A B C P X" and "Q OrthAt A B C Q Y" and "A B C TSP X Y" shows "P Q TS X Y" proof - obtain T where P0: "Coplanar A B C T \ Bet X T Y" using TSP_def assms(4) by auto have P1: "\ Col A B C" using assms(4) ncop__ncol tsp__ncop1 by blast have P2: "P \ X " using assms(2) orth_at_distincts by auto have P3: "Coplanar A B C P" using OrthAt_def assms(2) by blast have P4: "\ D. Coplanar A B C D \ Per D P X" using assms(2) orth_at_chara by blast have P5: "Q \ Y" using assms(3) orth_at_distincts by auto have P6: "Coplanar A B C Q" using OrthAt_def assms(3) by blast have P7: "\ D. Coplanar A B C D \ Per D Q Y" using assms(3) orth_at_chara by blast have P8: "\ Col X P Q" using P3 P6 assms(1) assms(4) col_cop2__cop not_col_permutation_2 tsp__ncop1 by blast have P9: "\ Col Y P Q" using P3 P6 assms(1) assms(4) col_cop2__cop not_col_permutation_2 tsp__ncop2 by blast have "Col T P Q" proof - obtain X' where Q1: "P Midpoint X X'" using symmetric_point_construction by auto obtain Y' where Q2: "Q Midpoint Y Y'" using symmetric_point_construction by auto have "Per T P X" using P0 P4 by auto then have Q3: "Cong T X T X'" using Q1 per_double_cong by auto have "Per T Q Y" using P0 P7 by auto then have Q4: "Cong T Y T Y'" using Q2 per_double_cong by auto have "Cong X Y X' Y'" using P1 Q1 Q2 assms(2) assms(3) mid2_orth_at2__cong by blast then have "X T Y Cong3 X' T Y'" using Cong3_def Q3 Q4 not_cong_2143 by blast then have "Bet X' T Y'" using l4_6 P0 by blast thus ?thesis using Q1 Q2 Q3 Q4 Col_def P0 between_symmetry l7_22 by blast qed thus ?thesis using P0 P8 P9 TS_def by blast qed lemma orth_dec: shows "A B C Orth U V \ \ A B C Orth U V" by auto lemma orth_at_dec: shows "X OrthAt A B C U V \ \ X OrthAt A B C U V" by auto lemma tsp_dec: shows "A B C TSP X Y \ \ A B C TSP X Y" by auto lemma osp_dec: shows "A B C OSP X Y \ \ A B C OSP X Y" by auto lemma ts2__inangle: assumes "A C TS B P" and "B P TS A C" shows "P InAngle A B C" by (metis InAngle_def assms(1) assms(2) bet_out ts2__ex_bet2 ts_distincts) lemma os_ts__inangle: assumes "B P TS A C" and "B A OS C P" shows "P InAngle A B C" proof - have P1: "\ Col A B P" using TS_def assms(1) by auto have P2: "\ Col B A C" using assms(2) col123__nos by blast obtain P' where P3: "B Midpoint P P'" using symmetric_point_construction by blast then have P4: "\ Col B A P'" by (metis assms(2) col_one_side col_permutation_5 midpoint_col midpoint_distinct_2 one_side_not_col124) have P5: "(B \ P' \ B P TS A C \ Bet P B P') \ (P InAngle A B C \ P' InAngle A B C)" using two_sides_in_angle by auto then have P6: "P InAngle A B C \ P' InAngle A B C" using P3 P4 assms(1) midpoint_bet not_col_distincts by blast { assume "P' InAngle A B C" then have P7: "A B OS P' C" using Col_cases P2 P4 in_angle_one_side by blast then have P8: "\ A B TS P' C" using l9_9 by auto have "B A TS P P'" using P1 P3 P4 bet__ts midpoint_bet not_col_distincts not_col_permutation_4 by auto then have "B A TS C P'" using P7 assms(2) invert_one_side l9_2 l9_8_2 l9_9 by blast then have "B A TS P' C" using l9_2 by blast then have "A B TS P' C" by (simp add: invert_two_sides) then have "P InAngle A B C" using P8 by auto } thus ?thesis using P6 by blast qed lemma os2__inangle: assumes "B A OS C P" and "B C OS A P" shows "P InAngle A B C" using assms(1) assms(2) col124__nos l9_9_bis os_ts__inangle two_sides_cases by blast lemma acute_conga__acute: assumes "Acute A B C" and "A B C CongA D E F" shows "Acute D E F" proof - have "D E F LeA A B C" by (simp add: assms(2) conga__lea456123) thus ?thesis using acute_lea_acute assms(1) by blast qed lemma acute_out2__acute: assumes "B Out A' A" and "B Out C' C" and "Acute A B C" shows "Acute A' B C'" by (meson Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms acute_conga__acute assms(1) assms(2) assms(3)) lemma conga_obtuse__obtuse: assumes "Obtuse A B C" and "A B C CongA D E F" shows "Obtuse D E F" using assms(1) assms(2) conga__lea lea_obtuse_obtuse by blast lemma obtuse_out2__obtuse: assumes "B Out A' A" and "B Out C' C" and "Obtuse A B C" shows "Obtuse A' B C'" by (meson Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) conga_obtuse__obtuse) lemma bet_lea__bet: assumes "Bet A B C" and "A B C LeA D E F" shows "Bet D E F" proof - have "A B C CongA D E F" by (metis assms(1) assms(2) l11_31_2 lea_asym lea_distincts) thus ?thesis using assms(1) bet_conga__bet by blast qed lemma out_lea__out: assumes "E Out D F" and "A B C LeA D E F" shows "B Out A C" proof - have "D E F CongA A B C" using Tarski_neutral_dimensionless.l11_31_1 Tarski_neutral_dimensionless.lea_asym Tarski_neutral_dimensionless.lea_distincts Tarski_neutral_dimensionless_axioms assms(1) assms(2) by fastforce thus ?thesis using assms(1) out_conga_out by blast qed lemma bet2_lta__lta: assumes "A B C LtA D E F" and "Bet A B A'" and "A' \ B" and "Bet D E D'" and "D' \ E" shows "D' E F LtA A' B C" proof - have P1: "D' E F LeA A' B C" by (metis Bet_cases assms(1) assms(2) assms(3) assms(4) assms(5) l11_36_aux2 lea_distincts lta__lea) have "\ D' E F CongA A' B C" by (metis assms(1) assms(2) assms(4) between_symmetry conga_sym l11_13 lta_distincts not_lta_and_conga) thus ?thesis by (simp add: LtA_def P1) qed lemma lea123456_lta__lta: assumes "A B C LeA D E F" and "D E F LtA G H I" shows "A B C LtA G H I" proof - have "\ G H I LeA F E D" by (metis (no_types) Tarski_neutral_dimensionless.lea__nlta Tarski_neutral_dimensionless.lta_left_comm Tarski_neutral_dimensionless_axioms assms(2)) then have "\ A B C CongA G H I" by (metis Tarski_neutral_dimensionless.lta_distincts Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga_pseudo_refl l11_30) thus ?thesis by (meson LtA_def Tarski_neutral_dimensionless.lea_trans Tarski_neutral_dimensionless_axioms assms(1) assms(2)) qed lemma lea456789_lta__lta: assumes "A B C LtA D E F" and "D E F LeA G H I" shows "A B C LtA G H I" by (meson LtA_def assms(1) assms(2) conga__lea456123 lea_trans lta__nlea) lemma acute_per__lta: assumes "Acute A B C" and "D \ E" and "E \ F" and "Per D E F" shows "A B C LtA D E F" proof - obtain G H I where P1: "Per G H I \ A B C LtA G H I" using Acute_def assms(1) by auto then have "G H I CongA D E F" using assms(2) assms(3) assms(4) l11_16 lta_distincts by blast thus ?thesis by (metis P1 conga_preserves_lta conga_refl lta_distincts) qed lemma obtuse_per__lta: assumes "Obtuse A B C" and "D \ E" and "E \ F" and "Per D E F" shows "D E F LtA A B C" proof - obtain G H I where P1: "Per G H I \ G H I LtA A B C" using Obtuse_def assms(1) by auto then have "G H I CongA D E F" using assms(2) assms(3) assms(4) l11_16 lta_distincts by blast thus ?thesis by (metis P1 Tarski_neutral_dimensionless.l11_51 Tarski_neutral_dimensionless_axioms assms(1) cong_reflexivity conga_preserves_lta obtuse_distincts) qed lemma acute_obtuse__lta: assumes "Acute A B C" and "Obtuse D E F" shows "A B C LtA D E F" by (metis Acute_def assms(1) assms(2) lta_distincts lta_trans obtuse_per__lta) lemma lea_in_angle: assumes "A B P LeA A B C" and "A B OS C P" shows "P InAngle A B C" proof - obtain T where P3: "T InAngle A B C \ A B P CongA A B T" using LeA_def assms(1) by blast thus ?thesis by (metis assms(2) conga_preserves_in_angle conga_refl not_conga_sym one_side_symmetry os_distincts) qed lemma acute_bet__obtuse: assumes "Bet A B A'" and "A' \ B" and "Acute A B C" shows "Obtuse A' B C" proof cases assume P1: "Col A B C" show ?thesis proof cases assume "Bet A B C" thus ?thesis using P1 acute_col__out assms(3) not_bet_and_out by blast next assume "\ Bet A B C" thus ?thesis by (smt P1 assms(1) assms(2) bet__obtuse between_inner_transitivity between_symmetry outer_transitivity_between third_point) qed next assume P2: "\ Col A B C" then obtain D where P3: "A B Perp D B \ A B OS C D" using col_trivial_2 l10_15 by blast { assume P4: "Col C B D" then have "Per A B C" proof - have P5: "B \ D" using P3 perp_not_eq_2 by auto have "Per A B D" using P3 Perp_perm perp_per_2 by blast thus ?thesis using P4 P5 not_col_permutation_2 per_col by blast qed then have "A B C LtA A B C" by (metis Acute_def acute_per__lta assms(3) lta_distincts) then have "False" by (simp add: nlta) } then have P6: "\ Col C B D" by auto have P7: "B A' OS C D" by (metis P3 assms(1) assms(2) bet_col col2_os__os l5_3) have T1: "Per A B D" by (simp add: P3 perp_left_comm perp_per_1) have Q1: "B C TS A' A" using P2 assms(1) assms(2) bet__ts l9_2 not_col_permutation_1 by auto have "A B C LtA A B D" using P2 P6 T1 acute_per__lta assms(3) not_col_distincts by auto then have "A B C LeA A B D" by (simp add: lta__lea) then have "C InAngle A B D" by (simp add: P3 lea_in_angle one_side_symmetry) then have "C InAngle D B A" using l11_24 by blast then have "C B TS D A" by (simp add: P2 P6 in_angle_two_sides not_col_permutation_1 not_col_permutation_4) then have "B C TS D A" using invert_two_sides by blast then have "B C OS A' D" using Q1 l9_8_1 by auto then have T1A: "D InAngle A' B C" by (simp add: P7 os2__inangle) then have "A B D CongA A' B D" by (metis Per_cases T1 Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.l11_18_1 Tarski_neutral_dimensionless_axioms acute_distincts assms(1) assms(3) inangle_distincts) then have T2: "A B D LeA A' B C" using LeA_def T1A by auto { assume "A B D CongA A' B C" then have "False" by (metis OS_def P7 T1 TS_def Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms \A B C LtA A B D\ \A B D CongA A' B D\ \\thesis. (\D. A B Perp D B \ A B OS C D \ thesis) \ thesis\ col_trivial_2 invert_one_side l11_17 l11_19 not_lta_and_conga out_trivial) } then have "\ A B D CongA A' B C" by auto then have "A B D LtA A' B C" using T2 LtA_def by auto thus ?thesis using T1 Obtuse_def by blast qed lemma bet_obtuse__acute: assumes "Bet A B A'" and "A' \ B" and "Obtuse A B C" shows "Acute A' B C" proof cases assume P1: "Col A B C" thus ?thesis proof cases assume "Bet A B C" then have "B Out A' C" by (smt Out_def assms(1) assms(2) assms(3) l5_2 obtuse_distincts) thus ?thesis by (simp add: out__acute) next assume "\ Bet A B C" thus ?thesis using P1 assms(3) col_obtuse__bet by blast qed next assume P2: "\ Col A B C" then obtain D where P3: "A B Perp D B \ A B OS C D" using col_trivial_2 l10_15 by blast { assume P3A: "Col C B D" have P3B: "B \ D" using P3 perp_not_eq_2 by blast have P3C: "Per A B D" using P3 Perp_perm perp_per_2 by blast then have "Per A B C" using P3A P3B not_col_permutation_2 per_col by blast then have "A B C LtA A B C" using P2 assms(3) not_col_distincts obtuse_per__lta by auto then have "False" by (simp add: nlta) } then have P4: "\ Col C B D" by auto have "Col B A A'" using Col_def Col_perm assms(1) by blast then have P5: "B A' OS C D" using P3 assms(2) invert_one_side col2_os__os col_trivial_3 by blast have P7: "Per A' B D" by (meson Col_def P3 Tarski_neutral_dimensionless.Per_perm Tarski_neutral_dimensionless_axioms assms(1) col_trivial_2 l8_16_1) have "A' B C LtA A' B D" proof - have P8: "A' B C LeA A' B D" proof - have P9: "C InAngle A' B D" proof - have P10: "B A' OS D C" by (simp add: P5 one_side_symmetry) have "B D OS A' C" proof - have P6: "\ Col A B D" using P3 col124__nos by auto then have P11: "B D TS A' A" using Col_perm P5 assms(1) bet__ts l9_2 os_distincts by blast have "A B D LtA A B C" proof - have P11A: "A \ B" using P2 col_trivial_1 by auto have P11B: "B \ D" using P4 col_trivial_2 by blast have "Per A B D" using P3 Perp_cases perp_per_2 by blast thus ?thesis by (simp add: P11A P11B Tarski_neutral_dimensionless.obtuse_per__lta Tarski_neutral_dimensionless_axioms assms(3)) qed then have "A B D LeA A B C" by (simp add: lta__lea) then have "D InAngle A B C" by (simp add: P3 lea_in_angle) then have "D InAngle C B A" using l11_24 by blast then have "D B TS C A" by (simp add: P4 P6 in_angle_two_sides not_col_permutation_4) then have "B D TS C A" by (simp add: invert_two_sides) thus ?thesis using OS_def P11 by blast qed thus ?thesis by (simp add: P10 os2__inangle) qed have "A' B C CongA A' B C" using assms(2) assms(3) conga_refl obtuse_distincts by blast thus ?thesis by (simp add: P9 inangle__lea) qed { assume "A' B C CongA A' B D" then have "B Out C D" using P5 conga_os__out invert_one_side by blast then have "False" using P4 not_col_permutation_4 out_col by blast } then have "\ A' B C CongA A' B D" by auto thus ?thesis by (simp add: LtA_def P8) qed thus ?thesis using Acute_def P7 by blast qed lemma inangle_dec: "P InAngle A B C \ \ P InAngle A B C" by blast lemma lea_dec: "A B C LeA D E F \ \ A B C LeA D E F" by blast lemma lta_dec: "A B C LtA D E F \ \ A B C LtA D E F" by blast lemma lea_total: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "A B C LeA D E F \ D E F LeA A B C" proof cases assume P1: "Col A B C" show ?thesis proof cases assume "B Out A C" thus ?thesis using assms(3) assms(4) l11_31_1 by auto next assume "\ B Out A C" thus ?thesis by (metis P1 assms(1) assms(2) assms(3) assms(4) l11_31_2 or_bet_out) qed next assume P2: "\ Col A B C" show ?thesis proof cases assume P3: "Col D E F" show ?thesis proof cases assume "E Out D F" thus ?thesis using assms(1) assms(2) l11_31_1 by auto next assume "\ E Out D F" thus ?thesis by (metis P3 assms(1) assms(2) assms(3) assms(4) l11_31_2 l6_4_2) qed next assume P4: "\ Col D E F" show ?thesis proof cases assume "A B C LeA D E F" thus ?thesis by simp next assume P5: "\ A B C LeA D E F" obtain P where P6: "D E F CongA A B P \ A B OS P C" using P2 P4 angle_construction_1 by blast then have P7: "B A OS C P" using invert_one_side one_side_symmetry by blast have "B C OS A P" proof - { assume "Col P B C" then have P7B: "B Out C P" using Col_cases P7 col_one_side_out by blast have "A B C CongA D E F" proof - have P7C: "A B P CongA D E F" by (simp add: P6 conga_sym) have P7D: "B Out A A" by (simp add: assms(1) out_trivial) have P7E: "E Out D D" by (simp add: assms(3) out_trivial) have "E Out F F" using assms(4) out_trivial by auto thus ?thesis using P7B P7C P7D P7E l11_10 by blast qed then have "A B C LeA D E F" by (simp add: conga__lea) then have "False" by (simp add: P5) } then have P8: "\ Col P B C" by auto { assume T0: "B C TS A P" have "A B C CongA D E F" proof - have T1: "A B C LeA A B P" proof - have T1A: "C InAngle A B P" by (simp add: P7 T0 one_side_symmetry os_ts__inangle) have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto thus ?thesis by (simp add: T1A inangle__lea) qed have T2: "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto have "A B P CongA D E F" by (simp add: P6 conga_sym) thus ?thesis using P5 T1 T2 l11_30 by blast qed then have "A B C LeA D E F" by (simp add: conga__lea) then have "False" by (simp add: P5) } then have "\ B C TS A P" by auto thus ?thesis using Col_perm P7 P8 one_side_symmetry os_ts1324__os two_sides_cases by blast qed then have "P InAngle A B C" using P7 os2__inangle by blast then have "D E F LeA A B C" using P6 LeA_def by blast thus ?thesis by simp qed qed qed lemma or_lta2_conga: assumes "A \ B" and "C \ B" and "D \ E" and "F \ E" shows "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" proof - have P1: "A B C LeA D E F \ D E F LeA A B C" using assms(1) assms(2) assms(3) assms(4) lea_total by auto { assume "A B C LeA D E F" then have "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" using LtA_def by blast } { assume "D E F LeA A B C" then have "A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F" using LtA_def conga_sym by blast } thus ?thesis using P1 \A B C LeA D E F \ A B C LtA D E F \ D E F LtA A B C \ A B C CongA D E F\ by blast qed lemma angle_partition: assumes "A \ B" and "B \ C" shows "Acute A B C \ Per A B C \ Obtuse A B C" proof - obtain A' B' D' where P1: "\ (Bet A' B' D' \ Bet B' D' A' \ Bet D' A' B')" using lower_dim by auto then have "\ Col A' B' D'" by (simp add: Col_def) obtain C' where P3: "A' B' Perp C' B'" by (metis P1 Perp_perm between_trivial2 perp_exists) then have P4: "A B C LtA A' B' C' \ A' B' C' LtA A B C \ A B C CongA A' B' C'" by (metis P1 assms(1) assms(2) between_trivial2 or_lta2_conga perp_not_eq_2) { assume "A B C LtA A' B' C'" then have "Acute A B C \ Per A B C \ Obtuse A B C" using Acute_def P3 Perp_cases perp_per_2 by blast } { assume "A' B' C' LtA A B C" then have "Acute A B C \ Per A B C \ Obtuse A B C" using Obtuse_def P3 Perp_cases perp_per_2 by blast } { assume "A B C CongA A' B' C'" then have "Acute A B C \ Per A B C \ Obtuse A B C" by (metis P3 Perp_cases Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.l11_17 Tarski_neutral_dimensionless_axioms perp_per_2) } thus ?thesis using P4 \A B C LtA A' B' C' \ Acute A B C \ Per A B C \ Obtuse A B C\ \A' B' C' LtA A B C \ Acute A B C \ Per A B C \ Obtuse A B C\ by auto qed lemma acute_chara_1: assumes "Bet A B A'" and "B \ A'" and "Acute A B C" shows "A B C LtA A' B C" proof - have "Obtuse A' B C" using acute_bet__obtuse assms(1) assms(2) assms(3) by blast thus ?thesis by (simp add: acute_obtuse__lta assms(3)) qed lemma acute_chara_2: assumes "Bet A B A'" and "A B C LtA A' B C" shows "Acute A B C" by (metis Tarski_neutral_dimensionless.Acute_def Tarski_neutral_dimensionless_axioms acute_chara_1 angle_partition assms(1) assms(2) bet_obtuse__acute between_symmetry lta_distincts not_and_lta) lemma acute_chara: assumes "Bet A B A'" and "B \ A'" shows "Acute A B C \ A B C LtA A' B C" using acute_chara_1 acute_chara_2 assms(1) assms(2) by blast lemma obtuse_chara: assumes "Bet A B A'" and "B \ A'" shows "Obtuse A B C \ A' B C LtA A B C" by (metis Tarski_neutral_dimensionless.Obtuse_def Tarski_neutral_dimensionless_axioms acute_bet__obtuse acute_chara assms(1) assms(2) bet_obtuse__acute between_symmetry lta_distincts) lemma conga__acute: assumes "A B C CongA A C B" shows "Acute A B C" by (metis acute_sym angle_partition assms conga_distinct conga_obtuse__obtuse l11_17 l11_43) lemma cong__acute: assumes "A \ B" and "B \ C" and "Cong A B A C" shows "Acute A B C" using angle_partition assms(1) assms(2) assms(3) cong__nlt l11_46 lt_left_comm by blast lemma nlta__lea: assumes "\ A B C LtA D E F" and "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "D E F LeA A B C" by (metis LtA_def assms(1) assms(2) assms(3) assms(4) assms(5) conga__lea456123 or_lta2_conga) lemma nlea__lta: assumes "\ A B C LeA D E F" and "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "D E F LtA A B C" using assms(1) assms(2) assms(3) assms(4) assms(5) nlta__lea by blast lemma triangle_strict_inequality: assumes "Bet A B D" and "Cong B C B D" and "\ Bet A B C" shows "A C Lt A D" proof cases assume P1: "Col A B C" then have P2: "B Out A C" using assms(3) not_out_bet by auto { assume "Bet B A C" then have P3: "A C Le A D" by (meson assms(1) assms(2) cong__le l5_12_a le_transitivity) have "\ Cong A C A D" by (metis Out_def P1 P2 assms(1) assms(2) assms(3) l4_18) then have "A C Lt A D" by (simp add: Lt_def P3) } { assume "Bet A C B" then have P5: "Bet A C D" using assms(1) between_exchange4 by blast then have P6: "A C Le A D" by (simp add: bet__le1213) have "\ Cong A C A D" using P5 assms(1) assms(3) between_cong by blast then have "A C Lt A D" by (simp add: Lt_def P6) } thus ?thesis using P1 \Bet B A C \ A C Lt A D\ assms(3) between_symmetry third_point by blast next assume T1: "\ Col A B C" have T2: "A \ D" using T1 assms(1) between_identity col_trivial_1 by auto have T3: "\ Col A C D" by (metis Col_def T1 T2 assms(1) col_transitivity_2) have T4: "Bet D B A" using Bet_perm assms(1) by blast have T5: "C D A CongA D C B" proof - have T6: "C D B CongA D C B" by (metis assms(1) assms(2) assms(3) between_trivial conga_comm l11_44_1_a not_conga_sym) have T7: "D Out C C" using T3 not_col_distincts out_trivial by blast have T8: "D Out A B" by (metis assms(1) assms(2) assms(3) bet_out_1 cong_diff l6_6) have T9: "C Out D D" using T7 out_trivial by force have "C Out B B" using T1 not_col_distincts out_trivial by auto thus ?thesis using T6 T7 T8 T9 l11_10 by blast qed have "A D C LtA A C D" proof - have "B InAngle D C A" by (metis InAngle_def T1 T3 T4 not_col_distincts out_trivial) then have "C D A LeA D C A" using T5 LeA_def by auto then have T10: "A D C LeA A C D" by (simp add: lea_comm) have "\ A D C CongA A C D" by (metis Col_perm T3 assms(1) assms(2) assms(3) bet_col l11_44_1_b l4_18 not_bet_distincts not_cong_3412) thus ?thesis using LtA_def T10 by blast qed thus ?thesis by (simp add: l11_44_2_b) qed lemma triangle_inequality: assumes "Bet A B D" and "Cong B C B D" shows "A C Le A D" proof cases assume "Bet A B C" thus ?thesis by (metis assms(1) assms(2) between_cong_3 cong__le le_reflexivity) next assume "\ Bet A B C" then have "A C Lt A D" using assms(1) assms(2) triangle_strict_inequality by auto thus ?thesis by (simp add: Lt_def) qed lemma triangle_strict_inequality_2: assumes "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" and "\ Bet A B C" shows "A C Lt A' C'" proof - obtain D where P1: "Bet A B D \ Cong B D B C" using segment_construction by blast then have P2: "A C Lt A D" using P1 assms(4) cong_symmetry triangle_strict_inequality by blast have "Cong A D A' C'" using P1 assms(1) assms(2) assms(3) cong_transitivity l2_11_b by blast thus ?thesis using P2 cong2_lt__lt cong_reflexivity by blast qed lemma triangle_inequality_2: assumes "Bet A' B' C'" and "Cong A B A' B'" and "Cong B C B' C'" shows "A C Le A' C'" proof - obtain D where P1: "Bet A B D \ Cong B D B C" using segment_construction by blast have P2: "A C Le A D" by (meson P1 Tarski_neutral_dimensionless.triangle_inequality Tarski_neutral_dimensionless_axioms not_cong_3412) have "Cong A D A' C'" using P1 assms(1) assms(2) assms(3) cong_transitivity l2_11_b by blast thus ?thesis using P2 cong__le le_transitivity by blast qed lemma triangle_strict_reverse_inequality: assumes "A Out B D" and "Cong A C A D" and "\ A Out B C" shows "B D Lt B C" proof cases assume "Col A B C" thus ?thesis using assms(1) assms(2) assms(3) col_permutation_4 cong_symmetry not_bet_and_out or_bet_out triangle_strict_inequality by blast next assume P1: "\ Col A B C" show ?thesis proof cases assume "B = D" thus ?thesis using P1 lt1123 not_col_distincts by auto next assume P2: "B \ D" then have P3: "\ Col B C D" using P1 assms(1) col_trivial_2 colx not_col_permutation_5 out_col by blast have P4: "\ Col A C D" using P1 assms(1) col2__eq col_permutation_4 out_col out_distinct by blast have P5: "C \ D" using assms(1) assms(3) by auto then have P6: "A C D CongA A D C" by (metis P1 assms(2) col_trivial_3 l11_44_1_a) { assume T1: "Bet A B D" then have T2: "Bet D B A" using Bet_perm by blast have "B C D LtA B D C" proof - have T3: "D C B CongA B C D" by (metis P3 conga_pseudo_refl not_col_distincts) have T3A: "D Out B A" by (simp add: P2 T1 bet_out_1) have T3B: "C Out D D" using P5 out_trivial by auto have T3C: "C Out A A" using P1 not_col_distincts out_trivial by blast have "D Out C C" by (simp add: P5 out_trivial) then have T4: "D C A CongA B D C" using T3A T3B T3C by (meson Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.l11_10 Tarski_neutral_dimensionless_axioms P6) have "D C B LtA D C A" proof - have T4A: "D C B LeA D C A" proof - have T4AA: "B InAngle D C A" using InAngle_def P1 P5 T2 not_col_distincts out_trivial by auto have "D C B CongA D C B" using T3 conga_right_comm by blast thus ?thesis by (simp add: T4AA inangle__lea) qed { assume T5: "D C B CongA D C A" have "D C OS B A" using Col_perm P3 T3A out_one_side by blast then have "C Out B A" using T5 conga_os__out by blast then have "False" using Col_cases P1 out_col by blast } then have "\ D C B CongA D C A" by auto thus ?thesis using LtA_def T4A by blast qed thus ?thesis using T3 T4 conga_preserves_lta by auto qed } { assume Q1: "Bet B D A" obtain E where Q2: "Bet B C E \ Cong B C C E" using Cong_perm segment_construction by blast have "A D C LtA E C D" proof - have Q3: "D C OS A E" proof - have Q4: "D C TS A B" by (metis Col_perm P2 Q1 P4 bet__ts between_symmetry) then have "D C TS E B" by (metis Col_def Q1 Q2 TS_def bet__ts cong_identity invert_two_sides l9_2) thus ?thesis using OS_def Q4 by blast qed have Q5: "A C D LtA E C D" proof - have "D C A LeA D C E" proof - have "B Out D A" using P2 Q1 bet_out by auto then have "B C OS D A" by (simp add: P3 out_one_side) then have "C B OS D A" using invert_one_side by blast then have "C E OS D A" by (metis Col_def Q2 Q3 col124__nos col_one_side diff_col_ex not_col_permutation_5) then have Q5A: "A InAngle D C E" by (simp add: \C E OS D A\ Q3 invert_one_side one_side_symmetry os2__inangle) have "D C A CongA D C A" using CongA_def P6 conga_refl by auto thus ?thesis by (simp add: Q5A inangle__lea) qed then have Q6: "A C D LeA E C D" using lea_comm by blast { assume "A C D CongA E C D" then have "D C A CongA D C E" by (simp add: conga_comm) then have "C Out A E" using Q3 conga_os__out by auto then have "False" by (meson Col_def Out_cases P1 Q2 not_col_permutation_3 one_side_not_col123 out_one_side) } then have "\ A C D CongA E C D" by auto thus ?thesis by (simp add: LtA_def Q6) qed have "E C D CongA E C D" by (metis P1 P5 Q2 cong_diff conga_refl not_col_distincts) thus ?thesis using Q5 P6 conga_preserves_lta by auto qed then have "B C D LtA B D C" using Bet_cases P1 P2 Q1 Q2 bet2_lta__lta not_col_distincts by blast } then have "B C D LtA B D C" by (meson Out_def \Bet A B D \ B C D LtA B D C\ assms(1) between_symmetry) thus ?thesis by (simp add: l11_44_2_b) qed qed lemma triangle_reverse_inequality: assumes "A Out B D" and "Cong A C A D" shows "B D Le B C" proof cases assume "A Out B C" thus ?thesis by (metis assms(1) assms(2) bet__le1213 cong_pseudo_reflexivity l6_11_uniqueness l6_6 not_bet_distincts not_cong_4312) next assume "\ A Out B C" thus ?thesis using triangle_strict_reverse_inequality assms(1) assms(2) lt__le by auto qed lemma os3__lta: assumes "A B OS C D" and "B C OS A D" and "A C OS B D" shows "B A C LtA B D C" proof - have P1: "D InAngle A B C" by (simp add: assms(1) assms(2) invert_one_side os2__inangle) then obtain E where P2: "Bet A E C \ (E = B \ B Out E D)" using InAngle_def by auto have P3: "\ Col A B C" using assms(1) one_side_not_col123 by auto have P4: "\ Col A C D" using assms(3) col124__nos by auto have P5: "\ Col B C D" using assms(2) one_side_not_col124 by auto have P6: "\ Col A B D" using assms(1) one_side_not_col124 by auto { assume "E = B" then have "B A C LtA B D C" using P2 P3 bet_col by blast } { assume P7: "B Out E D" have P8: "A \ E" using P6 P7 not_col_permutation_4 out_col by blast have P9: "C \ E" using P5 P7 out_col by blast have P10: "B A C LtA B E C" proof - have P10A: "\ Col E A B" by (metis Col_def P2 P3 P8 col_transitivity_1) then have P10B: "E B A LtA B E C" using P2 P9 Tarski_neutral_dimensionless.l11_41_aux Tarski_neutral_dimensionless_axioms by fastforce have P10C: "E A B LtA B E C" using P2 P9 P10A l11_41 by auto have P11: "E A B CongA B A C" proof - have P11A: "A Out B B" using assms(2) os_distincts out_trivial by auto have "A Out C E" using P2 P8 bet_out l6_6 by auto thus ?thesis using P11A conga_right_comm out2__conga by blast qed have P12: "B E C CongA B E C" by (metis Col_def P2 P3 P9 conga_refl) thus ?thesis using P11 P10C conga_preserves_lta by auto qed have "B E C LtA B D C" proof - have U1: "E Out D B" proof - obtain pp :: "'p \ 'p \ 'p" where f1: "\p pa. p \ (pp p pa) \ pa \ (pp p pa) \ Col p pa (pp p pa)" using diff_col_ex by moura then have "\p pa pb. Col pb pa p \ \ Col pb pa (pp p pa)" by (meson l6_16_1) then have f2: "\p pa. Col pa p pa" using f1 by metis have f3: "(E = B \ D = E) \ Col D E B" using f1 by (metis Col_def P2 col_out2_col l6_16_1 out_trivial) have "\p. (A = E \ Col p A C) \ \ Col p A E" using Col_def P2 l6_16_1 by blast thus ?thesis using f3 f2 by (metis (no_types) Col_def assms(3) not_out_bet one_side_chara one_side_symmetry) qed have U2: "D \ E" using P2 P4 bet_col not_col_permutation_5 by blast have U3: "\ Col D E C" by (metis Col_def P2 P4 P9 col_transitivity_1) have U4: "Bet E D B" by (simp add: P7 U1 out2__bet) have "D C E LtA C D B" using P5 U3 U4 l11_41_aux not_col_distincts by blast have U5: "D E C LtA C D B" using P7 U4 U3 l11_41 out_diff2 by auto have "D E C CongA B E C" by (simp add: P9 U1 l6_6 out2__conga out_trivial) thus ?thesis by (metis U5 conga_preserves_lta conga_pseudo_refl lta_distincts) qed then have "B A C LtA B D C" using P10 lta_trans by blast } thus ?thesis using P2 \E = B \ B A C LtA B D C\ by blast qed lemma bet_le__lt: assumes "Bet A D B" and "A \ D" and "D \ B" and "A C Le B C" shows "D C Lt B C" proof - have P1: "A \ B" using assms(1) assms(2) between_identity by blast have "C D Lt C B" proof cases assume P3: "Col A B C" thus ?thesis proof cases assume "Bet C D B" thus ?thesis by (simp add: assms(3) bet__lt1213) next assume "\ Bet C D B" then have "D Out C B" by (metis Out_def P1 P3 assms(1) col_transitivity_2 not_col_permutation_3 not_out_bet out_col) thus ?thesis by (smt Le_cases P3 assms(1) assms(2) assms(4) bet2_le2__le bet_le_eq bet_out_1 l6_6 l6_7 nle__lt or_bet_out out2__bet out_bet__out) qed next assume Q0A: "\ Col A B C" then have Q0B: "\ Col B C D" by (meson Col_def assms(1) assms(3) col_transitivity_2) have "C B D LtA C D B" proof - have Q1: "B Out C C" using Q0A not_col_distincts out_trivial by force have Q2: "B Out A D" using Out_cases assms(1) assms(3) bet_out_1 by blast have Q3: "A Out C C" by (metis Q0A col_trivial_3 out_trivial) have Q4: "A Out B B" using P1 out_trivial by auto have "C B A LeA C A B" using Col_perm Le_cases Q0A assms(4) l11_44_2bis by blast then have T9: "C B D LeA C A B" using Q1 Q2 Q3 Q4 lea_out4__lea by blast have "C A B LtA C D B" proof - have U2: "\ Col D A C" using Q0B assms(1) assms(2) bet_col col_transitivity_2 not_col_permutation_3 not_col_permutation_4 by blast have U3: "D \ C" using Q0B col_trivial_2 by blast have U4: "D C A LtA C D B" using U2 assms(1) assms(3) l11_41_aux by auto have U5: "D A C LtA C D B" by (simp add: U2 assms(1) assms(3) l11_41) have "A Out B D" using Out_def P1 assms(1) assms(2) by auto then have "D A C CongA C A B" using Q3 conga_right_comm out2__conga by blast thus ?thesis by (metis U5 U3 assms(3) conga_preserves_lta conga_refl) qed thus ?thesis using T9 lea123456_lta__lta by blast qed thus ?thesis by (simp add: l11_44_2_b) qed thus ?thesis using Lt_cases by auto qed lemma cong2__ncol: assumes "A \ B" and "B \ C" and "A \ C" and "Cong A P B P" and "Cong A P C P" shows "\ Col A B C" proof - have "Cong B P C P" using assms(4) assms(5) cong_inner_transitivity by blast thus ?thesis using bet_le__lt by (metis assms(1) assms(2) assms(3) assms(4) assms(5) cong__le cong__nlt lt__nle not_col_permutation_5 third_point) qed lemma cong4_cop2__eq: assumes "A \ B" and "B \ C" and "A \ C" and "Cong A P B P" and "Cong A P C P" and "Coplanar A B C P" and "Cong A Q B Q" and "Cong A Q C Q" and "Coplanar A B C Q" shows "P = Q" proof - have P1: "\ Col A B C" using assms(1) assms(2) assms(3) assms(4) assms(5) cong2__ncol by auto { assume P2: "P \ Q" have P3: "\ R. Col P Q R \ (Cong A R B R \ Cong A R C R)" using P2 assms(4) assms(5) assms(7) assms(8) l4_17 not_cong_4321 by blast obtain D where P4: "D Midpoint A B" using midpoint_existence by auto have P5: "Coplanar A B C D" using P4 coplanar_perm_9 midpoint__coplanar by blast have P6: "Col P Q D" proof - have P6A: "Coplanar P Q D A" using P1 P5 assms(6) assms(9) coplanar_pseudo_trans ncop_distincts by blast then have P6B: "Coplanar P Q D B" by (metis P4 col_cop__cop midpoint_col midpoint_distinct_1) have P6D: "Cong P A P B" using assms(4) not_cong_2143 by blast have P6E: "Cong Q A Q B" using assms(7) not_cong_2143 by blast have "Cong D A D B" using Midpoint_def P4 not_cong_2134 by blast thus ?thesis using cong3_cop2__col P6A P6B assms(1) P6D P6E by blast qed obtain R1 where P7: "P \ R1 \ Q \ R1 \ D \ R1 \ Col P Q R1" using P6 diff_col_ex3 by blast obtain R2 where P8: "Bet R1 D R2 \ Cong D R2 R1 D" using segment_construction by blast have P9: "Col P Q R2" by (metis P6 P7 P8 bet_col colx) have P9A: "Cong R1 A R1 B" using P3 P7 not_cong_2143 by blast then have "Per R1 D A" using P4 Per_def by auto then have "Per A D R1" using l8_2 by blast have P10: "Cong A R1 A R2" proof - have f1: "Bet R2 D R1 \Bet R1 R2 D" by (metis (full_types) Col_def P7 P8 between_equality not_col_permutation_5) have f2: "Cong B R1 A R1" using Cong_perm \Cong R1 A R1 B\ by blast have "Cong B R1 A R2 \ Bet R1 R2 D" using f1 Cong_perm Midpoint_def P4 P8 l7_13 by blast thus ?thesis using f2 P8 bet_cong_eq cong_inner_transitivity by blast qed have "Col A B C" proof - have P11: "Cong B R1 B R2" by (metis Cong_perm P10 P3 P9 P9A cong_inner_transitivity) have P12: "Cong C R1 C R2" using P10 P3 P7 P9 cong_inner_transitivity by blast have P12A: "Coplanar A B C R1" using P2 P7 assms(6) assms(9) col_cop2__cop by blast have P12B: "Coplanar A B C R2" using P2 P9 assms(6) assms(9) col_cop2__cop by blast have "R1 \ R2" using P7 P8 between_identity by blast thus ?thesis using P10 P11 P12A P12B P12 cong3_cop2__col by blast qed then have False by (simp add: P1) } thus ?thesis by auto qed lemma t18_18_aux: assumes "Cong A B D E" and "Cong A C D F" and "F D E LtA C A B" and "\ Col A B C" and "\ Col D E F" and "D F Le D E" shows "E F Lt B C" proof - obtain G0 where P1: "C A B CongA F D G0 \ F D OS G0 E" using angle_construction_1 assms(4) assms(5) not_col_permutation_2 by blast then have P2: "\ Col F D G0" using col123__nos by auto then obtain G where P3: "D Out G0 G \ Cong D G A B" by (metis assms(4) bet_col between_trivial2 col_trivial_2 segment_construction_3) have P4: "C A B CongA F D G" proof - have P4B: "A Out C C" by (metis assms(4) col_trivial_3 out_trivial) have P4C: "A Out B B" by (metis assms(4) col_trivial_1 out_trivial) have P4D: "D Out F F" using P2 not_col_distincts out_trivial by blast have "D Out G G0" by (simp add: P3 l6_6) thus ?thesis using P1 P4B P4C P4D using l11_10 by blast qed have "D Out G G0" by (simp add: P3 l6_6) then have "D F OS G G0" using P2 not_col_permutation_4 out_one_side by blast then have "F D OS G G0" by (simp add: invert_one_side) then have P5: "F D OS G E" using P1 one_side_transitivity by blast have P6: "\ Col F D G" by (meson P5 one_side_not_col123) have P7: "Cong C B F G" using P3 P4 assms(2) cong2_conga_cong cong_commutativity cong_symmetry by blast have P8: "F E Lt F G" proof - have P9: "F D E LtA F D G" by (metis P4 assms(3) assms(5) col_trivial_1 col_trivial_3 conga_preserves_lta conga_refl) have P10: "Cong D G D E" using P3 assms(1) cong_transitivity by blast { assume P11: "Col E F G" have P12: "F D E LeA F D G" by (simp add: P9 lta__lea) have P13: "\ F D E CongA F D G" using P9 not_lta_and_conga by blast have "F D E CongA F D G" proof - have "F Out E G" using Col_cases P11 P5 col_one_side_out l6_6 by blast then have "E F D CongA G F D" by (metis assms(5) conga_os__out conga_refl l6_6 not_col_distincts one_side_reflexivity out2__conga) thus ?thesis by (meson P10 assms(2) assms(6) cong_4321 cong_inner_transitivity l11_52 le_comm) qed then have "False" using P13 by blast } then have P15: "\ Col E F G" by auto { assume P20: "Col D E G" have P21: "F D E LeA F D G" by (simp add: P9 lta__lea) have P22: "\ F D E CongA F D G" using P9 not_lta_and_conga by blast have "F D E CongA F D G" proof - have "D Out E G" by (meson Out_cases P5 P20 col_one_side_out invert_one_side not_col_permutation_5) thus ?thesis using P10 P15 \D Out G G0\ cong_inner_transitivity l6_11_uniqueness l6_7 not_col_distincts by blast qed then have "False" by (simp add: P22) } then have P16: "\ Col D E G" by auto have P17: "E InAngle F D G" using lea_in_angle by (simp add: P5 P9 lta__lea) then obtain H where P18: "Bet F H G \ (H = D \ D Out H E)" using InAngle_def by auto { assume "H = D" then have "F G E LtA F E G" using P18 P6 bet_col by blast } { assume P19: "D Out H E" have P20: "H \ F" using Out_cases P19 assms(5) out_col by blast have P21: "H \ G" using P19 P16 l6_6 out_col by blast have "F D Le G D" using P10 assms(6) cong_pseudo_reflexivity l5_6 not_cong_4312 by blast then have "H D Lt G D" using P18 P20 P21 bet_le__lt by blast then have P22: "D H Lt D E" using Lt_cases P10 cong2_lt__lt cong_reflexivity by blast then have P23: "D H Le D E \ \ Cong D H D E" using Lt_def by blast have P24: "H \ E" using P23 cong_reflexivity by blast have P25: "Bet D H E" by (simp add: P19 P23 l6_13_1) have P26: "E G OS F D" by (metis InAngle_def P15 P16 P18 P25 bet_out_1 between_symmetry in_angle_one_side not_col_distincts not_col_permutation_1) have "F G E LtA F E G" proof - have P27: "F G E LtA D E G" proof - have P28: "D G E CongA D E G" by (metis P10 P16 l11_44_1_a not_col_distincts) have "F G E LtA D G E" proof - have P29: "F G E LeA D G E" by (metis OS_def P17 P26 P5 TS_def in_angle_one_side inangle__lea_1 invert_one_side l11_24 os2__inangle) { assume "F G E CongA D G E" then have "E G F CongA E G D" by (simp add: conga_comm) then have "G Out F D" using P26 conga_os__out by auto then have "False" using P6 not_col_permutation_2 out_col by blast } then have "\ F G E CongA D G E" by auto thus ?thesis by (simp add: LtA_def P29) qed thus ?thesis by (metis P28 P6 col_trivial_3 conga_preserves_lta conga_refl) qed have "G E D LtA G E F" proof - have P30: "G E D LeA G E F" proof - have P31: "D InAngle G E F" by (simp add: P16 P17 P26 assms(5) in_angle_two_sides l11_24 not_col_permutation_5 os_ts__inangle) have "G E D CongA G E D" by (metis P16 col_trivial_1 col_trivial_2 conga_refl) thus ?thesis using P31 inangle__lea by auto qed have "\ G E D CongA G E F" by (metis OS_def P26 P5 TS_def conga_os__out invert_one_side out_col) thus ?thesis by (simp add: LtA_def P30) qed then have "D E G LtA F E G" using lta_comm by blast thus ?thesis using P27 lta_trans by blast qed } then have "F G E LtA F E G" using P18 \H = D \ F G E LtA F E G\ by blast thus ?thesis by (simp add: l11_44_2_b) qed then have "E F Lt F G" using lt_left_comm by blast thus ?thesis using P7 cong2_lt__lt cong_pseudo_reflexivity not_cong_4312 by blast qed lemma t18_18: assumes "Cong A B D E" and "Cong A C D F" and "F D E LtA C A B" shows "E F Lt B C" proof - have P1: "F \ D" using assms(3) lta_distincts by blast have P2: "E \ D" using assms(3) lta_distincts by blast have P3: "C \ A" using assms(3) lta_distincts by auto have P4: "B \ A" using assms(3) lta_distincts by blast { assume P6: "Col A B C" { assume P7: "Bet B A C" obtain C' where P8:"Bet E D C' \ Cong D C' A C" using segment_construction by blast have P9: "Cong E F E F" by (simp add: cong_reflexivity) have P10: "Cong E C' B C" using P7 P8 assms(1) l2_11_b not_cong_4321 by blast have "E F Lt E C'" proof - have P11: "Cong D F D C'" using P8 assms(2) cong_transitivity not_cong_3412 by blast have "\ Bet E D F" using Bet_perm Col_def assms(3) col_lta__out not_bet_and_out by blast thus ?thesis using P11 P8 triangle_strict_inequality by blast qed then have "E F Lt B C" using P9 P10 cong2_lt__lt by blast } { assume "\ Bet B A C" then have "E F Lt B C" using P6 assms(3) between_symmetry col_lta__bet col_permutation_2 by blast } then have "E F Lt B C" using \Bet B A C \ E F Lt B C\ by auto } { assume P12: "\ Col A B C" { assume P13: "Col D E F" { assume P14: "Bet F D E" then have "C A B LeA F D E" by (simp add: P1 P2 P3 P4 l11_31_2) then have "F D E LtA F D E" using assms(3) lea__nlta by auto then have "False" by (simp add: nlta) then have "E F Lt B C" by auto } { assume "\ Bet F D E" then have P16: "D Out F E" using P13 not_col_permutation_1 not_out_bet by blast obtain F' where P17: "A Out B F' \ Cong A F' A C" using P3 P4 segment_construction_3 by fastforce then have P18: "B F' Lt B C" by (meson P12 Tarski_neutral_dimensionless.triangle_strict_reverse_inequality Tarski_neutral_dimensionless_axioms not_cong_3412 out_col) have "Cong B F' E F" by (meson Out_cases P16 P17 assms(1) assms(2) cong_transitivity out_cong_cong) then have "E F Lt B C" using P18 cong2_lt__lt cong_reflexivity by blast } then have "E F Lt B C" using \Bet F D E \ E F Lt B C\ by blast } { assume P20: "\ Col D E F" { assume "D F Le D E" then have "E F Lt B C" by (meson P12 Tarski_neutral_dimensionless.t18_18_aux Tarski_neutral_dimensionless_axioms P20 assms(1) assms(2) assms(3)) } { assume "D E Le D F" then have "E F Lt B C" by (meson P12 P20 Tarski_neutral_dimensionless.lta_comm Tarski_neutral_dimensionless.t18_18_aux Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) lt_comm not_col_permutation_5) } then have "E F Lt B C" using \D F Le D E \ E F Lt B C\ local.le_cases by blast } then have "E F Lt B C" using \Col D E F \ E F Lt B C\ by blast } thus ?thesis using \Col A B C \ E F Lt B C\ by auto qed lemma t18_19: assumes "A \ B" and "A \ C" and "Cong A B D E" and "Cong A C D F" and "E F Lt B C" shows "F D E LtA C A B" proof - { assume P1: "C A B LeA F D E" { assume "C A B CongA F D E" then have "False" using Cong_perm assms(3) assms(4) assms(5) cong__nlt l11_49 by blast } { assume P2: "\ C A B CongA F D E" then have "C A B LtA F E D" by (metis P1 assms(3) assms(4) assms(5) cong_symmetry lea_distincts lta__nlea not_and_lt or_lta2_conga t18_18) then have "B C Lt E F" by (metis P1 P2 assms(3) assms(4) cong_symmetry lta__nlea lta_distincts or_lta2_conga t18_18) then have "False" using assms(5) not_and_lt by auto } then have "False" using \C A B CongA F D E \ False\ by auto } then have "\ C A B LeA F D E" by auto thus ?thesis using assms(1) assms(2) assms(3) assms(4) cong_identity nlea__lta by blast qed lemma acute_trivial: assumes "A \ B" shows "Acute A B A" by (metis Tarski_neutral_dimensionless.acute_distincts Tarski_neutral_dimensionless_axioms angle_partition assms l11_43) lemma acute_not_per: assumes "Acute A B C" shows "\ Per A B C" proof - obtain A' B' C' where P1: "Per A' B' C' \ A B C LtA A' B' C'" using Acute_def assms by auto thus ?thesis using acute_distincts acute_per__lta assms nlta by fastforce qed lemma angle_bisector: assumes "A \ B" and "C \ B" shows "\ P. (P InAngle A B C \ P B A CongA P B C)" proof cases assume P1: "Col A B C" thus ?thesis proof cases assume P2: "Bet A B C" then obtain Q where P3: "\ Col A B Q" using assms(1) not_col_exists by auto then obtain P where P4: "A B Perp P B \ A B OS Q P" using P1 l10_15 os_distincts by blast then have P5: "P InAngle A B C" by (metis P2 assms(2) in_angle_line os_distincts) have "P B A CongA P B C" proof - have P9: "P \ B" using P4 os_distincts by blast have "Per P B A" by (simp add: P4 Perp_perm Tarski_neutral_dimensionless.perp_per_2 Tarski_neutral_dimensionless_axioms) thus ?thesis using P2 assms(1) assms(2) P9 l11_18_1 by auto qed thus ?thesis using P5 by auto next assume T1: "\ Bet A B C" then have T2: "B Out A C" by (simp add: P1 l6_4_2) have T3: "C InAngle A B C" by (simp add: assms(1) assms(2) inangle3123) have "C B A CongA C B C" using T2 between_trivial2 l6_6 out2__conga out2_bet_out by blast thus ?thesis using T3 by auto qed next assume T4: "\ Col A B C" obtain C0 where T5: "B Out C0 C \ Cong B C0 B A" using assms(1) assms(2) l6_11_existence by fastforce obtain P where T6: "P Midpoint A C0" using midpoint_existence by auto have T6A: "\ Col A B C0" by (metis T4 T5 col3 l6_3_1 not_col_distincts out_col) have T6B: "P \ B" using Col_def Midpoint_def T6 T6A by auto have T6D: "P \ A" using T6 T6A is_midpoint_id not_col_distincts by blast have "P InAngle A B C0" using InAngle_def T5 T6 T6B assms(1) l6_3_1 midpoint_bet out_trivial by fastforce then have T7: "P InAngle A B C" using T5 T6B in_angle_trans2 l11_24 out341__inangle by blast have T8: "(P = B) \ B Out P P" using out_trivial by auto have T9: "B Out A A" by (simp add: assms(1) out_trivial) { assume T9A: "B Out P P" have "P B A CongA P B C0 \ B P A CongA B P C0 \ P A B CongA P C0 B" proof - have T9B: "Cong B P B P" by (simp add: cong_reflexivity) have T9C: "Cong B A B C0" using Cong_perm T5 by blast have "Cong P A P C0" using Midpoint_def T6 not_cong_2134 by blast thus ?thesis using l11_51 T6B assms(1) T9B T9C T6D by presburger qed then have "P B A CongA P B C0" by auto then have "P B A CongA P B C" using l11_10 T9A T9 by (meson T5 l6_6) then have "\ P. (P InAngle A B C \ P B A CongA P B C)" using T7 by auto } thus ?thesis using T6B T8 by blast qed lemma reflectl__conga: assumes "A \ B" and "B \ P" and "P P' ReflectL A B" shows "A B P CongA A B P'" proof - obtain A' where P1: "A' Midpoint P' P \ Col A B A' \ (A B Perp P' P \ P = P')" using ReflectL_def assms(3) by auto { assume P2: "A B Perp P' P" then have P3: "P \ P'" using perp_not_eq_2 by blast then have P4: "A' \ P'" using P1 is_midpoint_id by blast have P5: "A' \ P" using P1 P3 is_midpoint_id_2 by auto have "A B P CongA A B P'" proof cases assume P6: "A' = B" then have P8: "B \ P'" using P4 by auto have P9: "Per A B P" by (smt P1 P3 P6 Perp_cases col_transitivity_2 midpoint_col midpoint_distinct_1 not_col_permutation_2 perp_col2_bis perp_per_2) have "Per A B P'" by (smt Mid_cases P1 P2 P6 P8 assms(1) col_trivial_3 midpoint_col not_col_permutation_3 perp_col4 perp_per_2) thus ?thesis using l11_16 P4 P5 P6 P9 assms(1) by auto next assume T1: "A' \ B" have T2: "B A' P CongA B A' P'" proof - have T2A: "Cong B P B P'" using assms(3) col_trivial_2 is_image_spec_col_cong l10_4_spec not_cong_4321 by blast then have T2B: "A' B P CongA A' B P'" by (metis Cong_perm Midpoint_def P1 P5 T1 Tarski_neutral_dimensionless.l11_51 Tarski_neutral_dimensionless_axioms assms(2) cong_reflexivity) have "A' P B CongA A' P' B" by (simp add: P5 T2A T2B cong_reflexivity conga_comm l11_49) thus ?thesis using P5 T2A T2B cong_reflexivity l11_49 by blast qed have T3: "Cong A' B A' B" by (simp add: cong_reflexivity) have "Cong A' P A' P'" using Midpoint_def P1 not_cong_4312 by blast then have T4: "A' B P CongA A' B P' \ A' P B CongA A' P' B" using l11_49 using assms(2) T2 T3 by blast show ?thesis proof cases assume "Bet A' B A" thus ?thesis using T4 assms(1) l11_13 by blast next assume "\ Bet A' B A" then have T5: "B Out A' A" using P1 not_col_permutation_3 or_bet_out by blast have T6: "B \ P'" using T4 conga_distinct by blast have T8: "B Out A A'" by (simp add: T5 l6_6) have T9: "B Out P P" using assms(2) out_trivial by auto have "B Out P' P'" using T6 out_trivial by auto thus ?thesis using l11_10 T4 T8 T9 by blast qed qed } { assume "P = P'" then have "A B P CongA A B P'" using assms(1) assms(2) conga_refl by auto } thus ?thesis using P1 \A B Perp P' P \ A B P CongA A B P'\ by blast qed lemma conga_cop_out_reflectl__out: assumes "\ B Out A C" and "Coplanar A B C P" and "P B A CongA P B C" and "B Out A T" and "T T' ReflectL B P" shows "B Out C T'" proof - have P1: "P B T CongA P B T'" by (metis assms(3) assms(4) assms(5) conga_distinct is_image_spec_rev out_distinct reflectl__conga) have P2: "T T' Reflect B P" by (metis P1 assms(5) conga_distinct is_image_is_image_spec) have P3: "B \ T'" using CongA_def P1 by blast have P4: "P B C CongA P B T'" proof - have P5: "P B C CongA P B A" by (simp add: assms(3) conga_sym) have "P B A CongA P B T'" proof - have P7: "B Out P P" using assms(3) conga_diff45 out_trivial by blast have P8: "B Out A T" by (simp add: assms(4)) have "B Out T' T'" using P3 out_trivial by auto thus ?thesis using P1 P7 P8 l11_10 by blast qed thus ?thesis using P5 not_conga by blast qed have "P B OS C T'" proof - have P9: "P B TS A C" using assms(1) assms(2) assms(3) conga_cop__or_out_ts coplanar_perm_20 by blast then have "T \ T'" by (metis Col_perm P2 P3 TS_def assms(4) col_transitivity_2 l10_8 out_col) then have "P B TS T T'" by (metis P2 P4 conga_diff45 invert_two_sides l10_14) then have "P B TS A T'" using assms(4) col_trivial_2 out_two_sides_two_sides by blast thus ?thesis using OS_def P9 l9_2 by blast qed thus ?thesis using P4 conga_os__out by auto qed lemma col_conga_cop_reflectl__col: assumes "\ B Out A C" and "Coplanar A B C P" and "P B A CongA P B C" and "Col B A T" and "T T' ReflectL B P" shows "Col B C T'" proof cases assume "B = T" thus ?thesis using assms(5) col_image_spec__eq not_col_distincts by blast next assume P1: "B \ T" thus ?thesis proof cases assume "B Out A T" thus ?thesis using out_col conga_cop_out_reflectl__out assms(1) assms(2) assms(3) assms(5) by blast next assume P2: "\ B Out A T" obtain A' where P3: "Bet A B A' \ Cong B A' A B" using segment_construction by blast obtain C' where P4: "Bet C B C' \ Cong B C' C B" using segment_construction by blast have P5: "B Out C' T'" proof - have P6: "\ B Out A' C'" by (metis P3 P4 assms(1) between_symmetry cong_diff_2 l6_2 out_diff1 out_diff2) have P7: "Coplanar A' B C' P" proof cases assume "Col A B C" thus ?thesis by (smt P3 P4 assms(1) assms(2) assms(3) bet_col bet_neq32__neq col2_cop__cop col_transitivity_1 colx conga_diff2 conga_diff56 l6_4_2 ncoplanar_perm_15 not_col_permutation_5) next assume P7B: "\ Col A B C" have P7C: "Coplanar A B C A'" using P3 bet_col ncop__ncols by blast have P7D: "Coplanar A B C B" using ncop_distincts by blast have "Coplanar A B C C'" using P4 bet__coplanar coplanar_perm_20 by blast thus ?thesis using P7B P7C P7D assms(2) coplanar_pseudo_trans by blast qed have P8: "P B A' CongA P B C'" by (metis CongA_def P3 P4 assms(3) cong_reverse_identity conga_left_comm l11_13 not_conga_sym) have P9: "B Out A' T" by (smt Out_def P1 P2 P3 P8 assms(3) assms(4) conga_distinct l5_2 l6_4_2 not_col_permutation_4) thus ?thesis using P6 P7 P8 P9 assms(5) conga_cop_out_reflectl__out by blast qed thus ?thesis by (metis Col_def P4 col_transitivity_1 out_col out_diff1) qed qed lemma conga2_cop2__col: assumes "\ B Out A C" and "P B A CongA P B C" and "P' B A CongA P' B C" and "Coplanar A B P P'" and "Coplanar B C P P'" shows "Col B P P'" proof - obtain C' where P1: "B Out C' C \ Cong B C' B A" by (metis assms(2) conga_distinct l6_11_existence) have P1A: "Cong P A P C' \ (P \ A \ (B P A CongA B P C' \ B A P CongA B C' P))" proof - have P2: "P B A CongA P B C'" proof - have P2A: "B Out P P" using assms(2) conga_diff45 out_trivial by auto have "B Out A A" using assms(2) conga_distinct out_trivial by auto thus ?thesis using P1 P2A assms(2) l11_10 by blast qed have P3: "Cong B P B P" by (simp add: cong_reflexivity) have "Cong B A B C'" using Cong_perm P1 by blast thus ?thesis using l11_49 P2 cong_reflexivity by blast qed have P4: "P' B A CongA P' B C'" proof - have P4A: "B Out P' P'" using assms(3) conga_diff1 out_trivial by auto have "B Out A A" using assms(2) conga_distinct out_trivial by auto thus ?thesis using P1 P4A assms(3) l11_10 by blast qed have P5: "Cong B P' B P'" by (simp add: cong_reflexivity) have P5A: "Cong B A B C'" using Cong_perm P1 by blast then have P6: "P' \ A \ (B P' A CongA B P' C' \ B A P' CongA B C' P')" using P4 P5 l11_49 by blast have P7: "Coplanar B P P' A" using assms(4) ncoplanar_perm_18 by blast have P8: "Coplanar B P P' C'" by (smt Col_cases P1 assms(5) col_cop__cop ncoplanar_perm_16 ncoplanar_perm_8 out_col out_diff2) have "A \ C'" using P1 assms(1) by auto thus ?thesis using P4 P5 P7 P8 P5A P1A cong3_cop2__col l11_49 by blast qed lemma conga2_cop2__col_1: assumes "\ Col A B C" and "P B A CongA P B C" and "P' B A CongA P' B C" and "Coplanar A B C P" and "Coplanar A B C P'" shows "Col B P P'" proof - have P1: "\ B Out A C" using Col_cases assms(1) out_col by blast have P2: "Coplanar A B P P'" by (meson assms(1) assms(4) assms(5) coplanar_perm_12 coplanar_trans_1 not_col_permutation_2) have "Coplanar B C P P'" using assms(1) assms(4) assms(5) coplanar_trans_1 by auto thus ?thesis using P1 P2 conga2_cop2__col assms(2) assms(3) conga2_cop2__col by auto qed lemma col_conga__conga: assumes "P B A CongA P B C" and "Col B P P'" and "B \ P'" shows "P' B A CongA P' B C" proof cases assume "Bet P B P'" thus ?thesis using assms(1) assms(3) l11_13 by blast next assume "\ Bet P B P'" then have P1: "B Out P P'" using Col_cases assms(2) or_bet_out by blast then have P2: "B Out P' P" by (simp add: l6_6) have P3: "B Out A A" using CongA_def assms(1) out_trivial by auto have "B Out C C" using assms(1) conga_diff56 out_trivial by blast thus ?thesis using P2 P3 assms(1) l11_10 by blast qed lemma cop_inangle__ex_col_inangle: assumes "\ B Out A C" and "P InAngle A B C" and "Coplanar A B C Q" shows "\ R. (R InAngle A B C \ P \ R \ Col P Q R)" proof - have P1: "A \ B" using assms(2) inangle_distincts by blast then have P4: "A \ C" using assms(1) out_trivial by blast have P2: "C \ B" using assms(2) inangle_distincts by auto have P3: "P \ B" using InAngle_def assms(2) by auto thus ?thesis proof cases assume "P = Q" thus ?thesis using P1 P2 P4 col_trivial_1 inangle1123 inangle3123 by blast next assume P5: "P \ Q" thus ?thesis proof cases assume P6: "Col B P Q" obtain R where P7: "Bet B P R \ Cong P R B P" using segment_construction by blast have P8: "R InAngle A B C" using Out_cases P1 P2 P3 P7 assms(2) bet_out l11_25 out_trivial by blast have "P \ R" using P3 P7 cong_reverse_identity by blast thus ?thesis by (metis P3 P6 P7 P8 bet_col col_transitivity_2) next assume T1: "\ Col B P Q" thus ?thesis proof cases assume T2: "Col A B C" have T3: "Q InAngle A B C" by (metis P1 P2 T1 T2 assms(1) in_angle_line l6_4_2 not_col_distincts) thus ?thesis using P5 col_trivial_2 by blast next assume Q1: "\ Col A B C" thus ?thesis proof cases assume Q2: "Col B C P" have Q3: "\ Col B A P" using Col_perm P3 Q1 Q2 col_transitivity_2 by blast have Q4: "Coplanar B P Q A" using P2 Q2 assms(3) col2_cop__cop col_trivial_3 ncoplanar_perm_22 ncoplanar_perm_3 by blast have Q5: "Q \ P" using P5 by auto have Q6: "Col B P P" using not_col_distincts by blast have Q7: "Col Q P P" using not_col_distincts by auto have "\ Col B P A" using Col_cases Q3 by auto then obtain Q0 where P10: "Col Q P Q0 \ B P OS A Q0" using cop_not_par_same_side Q4 Q5 Q6 Q7 T1 by blast have P13: "P \ Q0" using P10 os_distincts by auto { assume "B A OS P Q0" then have ?thesis using P10 P13 assms(2) in_angle_trans not_col_permutation_4 os2__inangle by blast } { assume V1: "\ B A OS P Q0" have "\ R. Bet P R Q0 \ Col P Q R \ Col B A R" proof cases assume V3: "Col B A Q0" have "Col P Q Q0" using Col_cases P10 by auto thus ?thesis using V3 between_trivial by auto next assume V4: "\ Col B A Q0" then have V5: "\ Col Q0 B A" using Col_perm by blast have "\ Col P B A" using Col_cases Q3 by blast then obtain R where V8: "Col R B A \ Bet P R Q0" using cop_nos__ts V1 V5 by (meson P10 TS_def ncoplanar_perm_2 os__coplanar) thus ?thesis by (metis Col_def P10 P13 col_transitivity_2) qed then obtain R where V9: "Bet P R Q0 \ Col P Q R \ Col B A R" by auto have V10: "P \ R" using Q3 V9 by blast have "R InAngle A B C" proof - have W1: "\ Col B P Q0" using P10 P13 T1 col2__eq by blast have "P Out Q0 R" using V10 V9 bet_out l6_6 by auto then have "B P OS Q0 R" using Q6 W1 out_one_side_1 by blast then have "B P OS A R" using P10 one_side_transitivity by blast then have "B Out A R" using V9 col_one_side_out by auto thus ?thesis by (simp add: P2 out321__inangle) qed then have ?thesis using V10 V9 by blast } thus ?thesis using \B A OS P Q0 \ \R. R InAngle A B C \ P \ R \ Col P Q R\ by blast next assume Z1: "\ Col B C P" then have Z6: "\ Col B P C" by (simp add: not_col_permutation_5) have Z3: "Col B P P" by (simp add: col_trivial_2) have Z4: "Col Q P P" by (simp add: col_trivial_2) have "Coplanar A B C P" using Q1 assms(2) inangle__coplanar ncoplanar_perm_18 by blast then have "Coplanar B P Q C" using Q1 assms(3) coplanar_trans_1 ncoplanar_perm_5 by blast then obtain Q0 where Z5: "Col Q P Q0 \ B P OS C Q0" using cop_not_par_same_side by (metis Z3 Z4 T1 Z6) thus ?thesis proof cases assume "B C OS P Q0" thus ?thesis proof - have "\p. p InAngle C B A \ \ p InAngle C B P" using assms(2) in_angle_trans l11_24 by blast then show ?thesis by (metis Col_perm Z5 \B C OS P Q0\ l11_24 os2__inangle os_distincts) qed next assume Z6: "\ B C OS P Q0" have Z7: "\ R. Bet P R Q0 \ Col P Q R \ Col B C R" proof cases assume "Col B C Q0" thus ?thesis using Col_def Col_perm Z5 between_trivial by blast next assume Z8: "\ Col B C Q0" have "\ R. Col R B C \ Bet P R Q0" proof - have Z10: "Coplanar B C P Q0" using Z5 ncoplanar_perm_2 os__coplanar by blast have Z11: "\ Col P B C" using Col_cases Z1 by blast have "\ Col Q0 B C" using Col_perm Z8 by blast thus ?thesis using cop_nos__ts Z6 Z10 Z11 by (simp add: TS_def) qed then obtain R where "Col R B C \ Bet P R Q0" by blast thus ?thesis by (smt Z5 bet_col col2__eq col_permutation_1 os_distincts) qed then obtain R where Z12: "Bet P R Q0 \ Col P Q R \ Col B C R" by blast have Z13: "P \ R" using Z1 Z12 by auto have Z14: "\ Col B P Q0" using Z5 one_side_not_col124 by blast have "P Out Q0 R" using Z12 Z13 bet_out l6_6 by auto then have "B P OS Q0 R" using Z14 Z3 out_one_side_1 by blast then have "B P OS C R" using Z5 one_side_transitivity by blast then have "B Out C R" using Z12 col_one_side_out by blast then have "R InAngle A B C" using P1 out341__inangle by auto thus ?thesis using Z12 Z13 by auto qed qed qed qed qed qed lemma col_inangle2__out: assumes "\ Bet A B C" and "P InAngle A B C" and "Q InAngle A B C" and "Col B P Q" shows "B Out P Q" proof cases assume "Col A B C" thus ?thesis by (meson assms(1) assms(2) assms(3) assms(4) bet_in_angle_bet bet_out__bet in_angle_out l6_6 not_col_permutation_4 or_bet_out) next assume P1: "\ Col A B C" thus ?thesis proof cases assume "Col B A P" thus ?thesis by (meson assms(1) assms(2) assms(3) assms(4) bet_in_angle_bet bet_out__bet l6_6 not_col_permutation_4 or_bet_out) next assume P2: "\ Col B A P" have "\ Col B A Q" using P2 assms(3) assms(4) col2__eq col_permutation_4 inangle_distincts by blast then have "B A OS P Q" using P1 P2 assms(2) assms(3) inangle_one_side invert_one_side not_col_permutation_4 by auto thus ?thesis using assms(4) col_one_side_out by auto qed qed lemma inangle2__lea: assumes "P InAngle A B C" and "Q InAngle A B C" shows "P B Q LeA A B C" proof - have P1: "P InAngle C B A" by (simp add: assms(1) l11_24) have P2: "Q InAngle C B A" by (simp add: assms(2) l11_24) have P3: "A \ B" using assms(1) inangle_distincts by auto have P4: "C \ B" using assms(1) inangle_distincts by blast have P5: "P \ B" using assms(1) inangle_distincts by auto have P6: "Q \ B" using assms(2) inangle_distincts by auto thus ?thesis proof cases assume P7: "Col A B C" thus ?thesis proof cases assume "Bet A B C" thus ?thesis by (simp add: P3 P4 P5 P6 l11_31_2) next assume "\ Bet A B C" then have "B Out A C" using P7 not_out_bet by blast then have "B Out P Q" using Out_cases assms(1) assms(2) in_angle_out l6_7 by blast thus ?thesis by (simp add: P3 P4 l11_31_1) qed next assume T1: "\ Col A B C" thus ?thesis proof cases assume T2: "Col B P Q" have "\ Bet A B C" using T1 bet_col by auto then have "B Out P Q" using T2 assms(1) assms(2) col_inangle2__out by auto thus ?thesis by (simp add: P3 P4 l11_31_1) next assume T3: "\ Col B P Q" thus ?thesis proof cases assume "Col B A P" then have "B Out A P" using Col_def T1 assms(1) col_in_angle_out by blast then have "P B Q CongA A B Q" using P6 out2__conga out_trivial by auto thus ?thesis using LeA_def assms(2) by blast next assume W0: "\ Col B A P" show ?thesis proof cases assume "Col B C P" then have "B Out C P" by (metis P1 P3 T1 bet_out_1 col_in_angle_out out_col) thus ?thesis by (smt P3 P4 P6 Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms assms(2) inangle__lea_1 out_trivial) next assume W0A: "\ Col B C P" show ?thesis proof cases assume "Col B A Q" then have "B Out A Q" using Col_def T1 assms(2) col_in_angle_out by blast thus ?thesis by (smt P3 P4 P5 Tarski_neutral_dimensionless.lea_left_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms assms(1) inangle__lea out_trivial) next assume W0AA: "\ Col B A Q" thus ?thesis proof cases assume "Col B C Q" then have "B Out C Q" using Bet_cases P2 T1 bet_col col_in_angle_out by blast thus ?thesis by (smt P1 P3 P4 P5 Tarski_neutral_dimensionless.lea_comm Tarski_neutral_dimensionless.lea_out4__lea Tarski_neutral_dimensionless_axioms inangle__lea out_trivial) next assume W0B: "\ Col B C Q" have W1: "Coplanar B P A Q" by (metis Col_perm T1 assms(1) assms(2) col__coplanar inangle_one_side ncoplanar_perm_13 os__coplanar) have W2: "\ Col A B P" by (simp add: W0 not_col_permutation_4) have W3: "\ Col Q B P" using Col_perm T3 by blast then have W4: "B P TS A Q \ B P OS A Q" using cop__one_or_two_sides by (simp add: W1 W2) { assume W4A: "B P TS A Q" have "Q InAngle P B C" proof - have W5: "P B OS C Q" using OS_def P1 W0 W0A W4A in_angle_two_sides invert_two_sides l9_2 by blast have "C B OS P Q" by (meson P1 P2 T1 W0A W0B inangle_one_side not_col_permutation_3 not_col_permutation_4) thus ?thesis by (simp add: W5 invert_one_side os2__inangle) qed then have "P B Q LeA A B C" by (meson assms(1) inangle__lea inangle__lea_1 lea_trans) } { assume W6: "B P OS A Q" have "B A OS P Q" using Col_perm T1 W2 W0AA assms(1) assms(2) inangle_one_side invert_one_side by blast then have "Q InAngle P B A" by (simp add: W6 os2__inangle) then have "P B Q LeA A B C" by (meson P1 inangle__lea inangle__lea_1 lea_right_comm lea_trans) } thus ?thesis using W4 \B P TS A Q \ P B Q LeA A B C\ by blast qed qed qed qed qed qed qed lemma conga_inangle_per__acute: assumes "Per A B C" and "P InAngle A B C" and "P B A CongA P B C" shows "Acute A B P" proof - have P1: "\ Col A B C" using assms(1) assms(3) conga_diff2 conga_diff56 l8_9 by blast have P2: "A B P LeA A B C" by (simp add: assms(2) inangle__lea) { assume "A B P CongA A B C" then have P3: "Per A B P" by (meson Tarski_neutral_dimensionless.l11_17 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms assms(1)) have P4: "Coplanar P C A B" using assms(2) inangle__coplanar ncoplanar_perm_3 by blast have P5: "P \ B" using assms(2) inangle_distincts by blast have "Per C B P" using P3 Per_cases assms(3) l11_17 by blast then have "False" using P1 P3 P4 P5 col_permutation_1 cop_per2__col by blast } then have "\ A B P CongA A B C" by auto then have "A B P LtA A B C" by (simp add: LtA_def P2) thus ?thesis using Acute_def assms(1) by blast qed lemma conga_inangle2_per__acute: assumes "Per A B C" and "P InAngle A B C" and "P B A CongA P B C" and "Q InAngle A B C" shows "Acute P B Q" proof - have P1: "P InAngle C B A" using assms(2) l11_24 by auto have P2: "Q InAngle C B A" using assms(4) l11_24 by blast have P3: "A \ B" using assms(3) conga_diff2 by auto have P5: "P \ B" using assms(2) inangle_distincts by blast have P7: "\ Col A B C" using assms(1) assms(3) conga_distinct l8_9 by blast have P8: "Acute A B P" using assms(1) assms(2) assms(3) conga_inangle_per__acute by auto { assume "Col P B A" then have "Col P B C" using assms(3) col_conga_col by blast then have "False" using Col_perm P5 P7 \Col P B A\ col_transitivity_2 by blast } then have P9: "\ Col P B A" by auto have P10: "\ Col P B C" using \Col P B A \ False\ assms(3) ncol_conga_ncol by blast have P11: "\ Bet A B C" using P7 bet_col by blast show ?thesis proof cases assume "Col B A Q" then have "B Out A Q" using P11 assms(4) col_in_angle_out by auto thus ?thesis using Out_cases P5 P8 acute_out2__acute acute_sym out_trivial by blast next assume S0: "\ Col B A Q" show ?thesis proof cases assume S1: "Col B C Q" then have "B Out C Q" using P11 P2 between_symmetry col_in_angle_out by blast then have S2: "B Out Q C" using l6_6 by blast have S3: "B Out P P" by (simp add: P5 out_trivial) have "B Out A A" by (simp add: P3 out_trivial) then have "A B P CongA P B Q" using S2 conga_left_comm l11_10 S3 assms(3) by blast thus ?thesis using P8 acute_conga__acute by blast next assume S4: "\ Col B C Q" show ?thesis proof cases assume "Col B P Q" thus ?thesis using out__acute col_inangle2__out P11 assms(2) assms(4) by blast next assume S5: "\ Col B P Q" have S6: "Coplanar B P A Q" by (metis Col_perm P7 assms(2) assms(4) coplanar_trans_1 inangle__coplanar ncoplanar_perm_12 ncoplanar_perm_21) have S7: "\ Col A B P" using Col_cases P9 by auto have "\ Col Q B P" using Col_perm S5 by blast then have S8: "B P TS A Q \ B P OS A Q" using cop__one_or_two_sides S6 S7 by blast { assume S9: "B P TS A Q" have S10: "Acute P B C" using P8 acute_conga__acute acute_sym assms(3) by blast have "Q InAngle P B C" proof - have S11: "P B OS C Q" by (metis Col_perm OS_def P1 P10 P9 S9 in_angle_two_sides invert_two_sides l9_2) have "C B OS P Q" by (meson P1 P10 P2 P7 S4 inangle_one_side not_col_permutation_3 not_col_permutation_4) thus ?thesis by (simp add: S11 invert_one_side os2__inangle) qed then have "P B Q LeA P B C" by (simp add: inangle__lea) then have "Acute P B Q" using S10 acute_lea_acute by blast } { assume S12: "B P OS A Q" have "B A OS P Q" using Col_perm P7 S7 S0 assms(2) assms(4) inangle_one_side invert_one_side by blast then have "Q InAngle P B A" by (simp add: S12 os2__inangle) then have "Q B P LeA P B A" by (simp add: P3 P5 inangle1123 inangle2__lea) then have "P B Q LeA A B P" by (simp add: lea_comm) then have "Acute P B Q" using P8 acute_lea_acute by blast } thus ?thesis using \B P TS A Q \ Acute P B Q\ S8 by blast qed qed qed qed lemma lta_os__ts: assumes (*"\ Col A O1 P" and*) "A O1 P LtA A O1 B" and "O1 A OS B P" shows "O1 P TS A B" proof - have "A O1 P LeA A O1 B" by (simp add: assms(1) lta__lea) then have "\ P0. P0 InAngle A O1 B \ A O1 P CongA A O1 P0" by (simp add: LeA_def) then obtain P' where P1: "P' InAngle A O1 B \ A O1 P CongA A O1 P'" by blast have P2: "\ Col A O1 B" using assms(2) col123__nos not_col_permutation_4 by blast obtain R where P3: "O1 A TS B R \ O1 A TS P R" using OS_def assms(2) by blast { assume "Col B O1 P" then have "Bet B O1 P" by (metis Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) between_trivial col_trivial_2 lta_not_conga one_side_chara or_bet_out out_trivial) then have "O1 A TS B P" using assms(2) col_trivial_1 one_side_chara by blast then have P6: "\ O1 A OS B P" using l9_9_bis by auto then have "False" using P6 assms(2) by auto } then have P4: "\ Col B O1 P" by auto thus ?thesis by (meson P3 assms(1) inangle__lta l9_8_1 not_and_lta not_col_permutation_4 os_ts__inangle two_sides_cases) qed lemma bet__suppa: assumes "A \ B" and "B \ C" and "B \ A'" and "Bet A B A'" shows "A B C SuppA C B A'" proof - have "C B A' CongA C B A'" using assms(2) assms(3) conga_refl by auto thus ?thesis using assms(4) assms(1) SuppA_def by auto qed lemma ex_suppa: assumes "A \ B" and "B \ C" shows "\ D E F. A B C SuppA D E F" proof - obtain A' where "Bet A B A' \ Cong B A' A B" using segment_construction by blast thus ?thesis by (meson assms(1) assms(2) bet__suppa point_construction_different) qed lemma suppa_distincts: assumes "A B C SuppA D E F" shows "A \ B \ B \ C \ D \ E \ E \ F" using CongA_def SuppA_def assms by auto lemma suppa_right_comm: assumes "A B C SuppA D E F" shows "A B C SuppA F E D" using SuppA_def assms conga_left_comm by auto lemma suppa_left_comm: assumes "A B C SuppA D E F" shows "C B A SuppA D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto obtain C' where P2: "Bet C B C' \ Cong B C' C B" using segment_construction by blast then have "C B A' CongA A B C'" by (metis Bet_cases P1 SuppA_def assms cong_diff_3 conga_diff45 conga_diff56 conga_left_comm l11_14) then have "D E F CongA A B C'" using P1 conga_trans by blast thus ?thesis by (metis CongA_def P1 P2 SuppA_def) qed lemma suppa_comm: assumes "A B C SuppA D E F" shows "C B A SuppA F E D" using assms suppa_left_comm suppa_right_comm by blast lemma suppa_sym: assumes "A B C SuppA D E F" shows "D E F SuppA A B C" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto obtain D' where P2: "Bet D E D' \ Cong E D' D E" using segment_construction by blast have "A' B C CongA D E F" using P1 conga_right_comm not_conga_sym by blast then have "A B C CongA F E D'" by (metis P1 P2 Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.l11_13 Tarski_neutral_dimensionless.suppa_distincts Tarski_neutral_dimensionless_axioms assms between_symmetry cong_diff_3) thus ?thesis by (metis CongA_def P1 P2 SuppA_def) qed lemma conga2_suppa__suppa: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "A B C SuppA D E F" shows "A' B' C' SuppA D' E' F'" proof - obtain A0 where P1: "Bet A B A0 \ D E F CongA C B A0" using SuppA_def assms(3) by auto then have "A B C SuppA D' E' F'" by (metis Tarski_neutral_dimensionless.SuppA_def Tarski_neutral_dimensionless_axioms assms(2) assms(3) conga_sym conga_trans) then have P2: "D' E' F' SuppA A B C" by (simp add: suppa_sym) then obtain D0 where P3: "Bet D' E' D0 \ A B C CongA F' E' D0" using P2 SuppA_def by auto have P5: "A' B' C' CongA F' E' D0" using P3 assms(1) not_conga not_conga_sym by blast then have "D' E' F' SuppA A' B' C'" using P2 P3 SuppA_def by auto thus ?thesis by (simp add: suppa_sym) qed lemma suppa2__conga456: assumes "A B C SuppA D E F" and "A B C SuppA D' E' F'" shows "D E F CongA D' E' F'" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(1) by auto obtain A'' where P2: "Bet A B A'' \ D' E' F' CongA C B A''" using SuppA_def assms(2) by auto have "C B A' CongA C B A''" proof - have P3: "B Out C C" using P1 by (simp add: CongA_def out_trivial) have "B Out A'' A'" using P1 P2 l6_2 by (metis assms(1) between_symmetry conga_distinct suppa_distincts) thus ?thesis by (simp add: P3 out2__conga) qed then have "C B A' CongA D' E' F'" using P2 not_conga not_conga_sym by blast thus ?thesis using P1 not_conga by blast qed lemma suppa2__conga123: assumes "A B C SuppA D E F" and "A' B' C' SuppA D E F" shows "A B C CongA A' B' C'" using assms(1) assms(2) suppa2__conga456 suppa_sym by blast lemma bet_out__suppa: assumes "A \ B" and "B \ C" and "Bet A B C" and "E Out D F" shows "A B C SuppA D E F" proof - have "D E F CongA C B C" using assms(2) assms(4) l11_21_b out_trivial by auto thus ?thesis using SuppA_def assms(1) assms(3) by blast qed lemma bet_suppa__out: assumes "Bet A B C" and "A B C SuppA D E F" shows "E Out D F" proof - have "A B C SuppA C B C" using assms(1) assms(2) bet__suppa suppa_distincts by auto then have "C B C CongA D E F" using assms(2) suppa2__conga456 by auto thus ?thesis using eq_conga_out by auto qed lemma out_suppa__bet: assumes "B Out A C" and "A B C SuppA D E F" shows "Bet D E F" proof - obtain B' where P1: "Bet A B B' \ Cong B B' A B" using segment_construction by blast have "A B C SuppA A B B'" by (metis P1 assms(1) assms(2) bet__suppa bet_cong_eq bet_out__bet suppa_distincts suppa_left_comm) then have "A B B' CongA D E F" using assms(2) suppa2__conga456 by auto thus ?thesis using P1 bet_conga__bet by blast qed lemma per_suppa__per: assumes "Per A B C" and "A B C SuppA D E F" shows "Per D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto have "Per C B A'" proof - have P2: "A \ B" using assms(2) suppa_distincts by auto have P3: "Per C B A" by (simp add: assms(1) l8_2) have "Col B A A'" using P1 Col_cases Col_def by blast thus ?thesis by (metis P2 P3 per_col) qed thus ?thesis using P1 l11_17 not_conga_sym by blast qed lemma per2__suppa: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "Per A B C" and "Per D E F" shows "A B C SuppA D E F" proof - obtain D' E' F' where P1: "A B C SuppA D' E' F'" using assms(1) assms(2) ex_suppa by blast have "D' E' F' CongA D E F" using P1 assms(3) assms(4) assms(5) assms(6) l11_16 per_suppa__per suppa_distincts by blast thus ?thesis by (meson P1 conga2_suppa__suppa suppa2__conga123) qed lemma suppa__per: assumes "A B C SuppA A B C" shows "Per A B C" proof - obtain A' where P1: "Bet A B A' \ A B C CongA C B A'" using SuppA_def assms by auto then have "C B A CongA C B A'" by (simp add: conga_left_comm) thus ?thesis using P1 Per_perm l11_18_2 by blast qed lemma acute_suppa__obtuse: assumes "Acute A B C" and "A B C SuppA D E F" shows "Obtuse D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto then have "Obtuse C B A'" by (metis Tarski_neutral_dimensionless.obtuse_sym Tarski_neutral_dimensionless_axioms acute_bet__obtuse assms(1) conga_distinct) thus ?thesis by (meson P1 Tarski_neutral_dimensionless.conga_obtuse__obtuse Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms) qed lemma obtuse_suppa__acute: assumes "Obtuse A B C" and "A B C SuppA D E F" shows "Acute D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(2) by auto then have "Acute C B A'" using acute_sym assms(1) bet_obtuse__acute conga_distinct by blast thus ?thesis using P1 acute_conga__acute not_conga_sym by blast qed lemma lea_suppa2__lea: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" "A B C LeA D E F" shows "D' E' F' LeA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto have "F E D0 LeA C B A0" proof - have P3: "D0 \ E" using CongA_def P2 by auto have P4: "A0 \ B" using CongA_def P1 by blast have P6: "Bet D0 E D" by (simp add: P2 between_symmetry) have "Bet A0 B A" by (simp add: P1 between_symmetry) thus ?thesis by (metis P3 P4 P6 assms(3) l11_36_aux2 lea_comm lea_distincts) qed thus ?thesis by (meson P1 P2 Tarski_neutral_dimensionless.l11_30 Tarski_neutral_dimensionless.not_conga_sym Tarski_neutral_dimensionless_axioms) qed lemma lta_suppa2__lta: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" and "A B C LtA D E F" shows "D' E' F' LtA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto have "F E D0 LtA C B A0" proof - have P5: "A0 \ B" using CongA_def P1 by blast have "D0 \ E" using CongA_def P2 by auto thus ?thesis using assms(3) P1 P5 P2 bet2_lta__lta lta_comm by blast qed thus ?thesis using P1 P2 conga_preserves_lta not_conga_sym by blast qed lemma suppa_dec: "A B C SuppA D E F \ \ A B C SuppA D E F" by simp lemma acute_one_side_aux: assumes "C A OS P B" and "Acute A C P" and "C A Perp B C" shows "C B OS A P" proof - obtain R where T1: "C A TS P R \ C A TS B R" using OS_def assms(1) by blast obtain A' B' C' where P1: "Per A' B' C' \ A C P LtA A' B' C'" using Acute_def assms(2) by auto have P2: "Per A C B" by (simp add: assms(3) perp_per_1) then have P3: "A' B' C' CongA A C B" using P1 assms(1) l11_16 lta_distincts os_distincts by blast have P4: "A C P LtA A C B" by (metis P2 acute_per__lta assms(1) assms(2) os_distincts) { assume P4A: "Col P C B" have "Per A C P" proof - have P4B: "C \ B" using assms(1) os_distincts by blast have P4C: "Per A C B" by (simp add: P2) have "Col C B P" using P4A Col_cases by auto thus ?thesis using per_col P4B P4C by blast qed then have "False" using acute_not_per assms(2) by auto } then have P5: "\ Col P C B" by auto have P6: "\ Col A C P" using assms(1) col123__nos not_col_permutation_4 by blast have P7: "C B TS A P \ C B OS A P" using P5 assms(1) not_col_permutation_4 os_ts1324__os two_sides_cases by blast { assume P8: "C B TS A P" then obtain T where P9: "Col T C B \ Bet A T P" using TS_def by blast then have P10: "C \ T" using Col_def P6 P9 by auto have "T InAngle A C P" by (meson P4 P5 P8 Tarski_neutral_dimensionless.inangle__lta Tarski_neutral_dimensionless_axioms assms(1) not_and_lta not_col_permutation_3 os_ts__inangle) then have "C A OS T P" by (metis P10 P9 T1 TS_def col123__nos in_angle_one_side invert_one_side l6_16_1 one_side_reflexivity) then have P13: "C A OS T B" using assms(1) one_side_transitivity by blast have "C B OS A P" by (meson P4 Tarski_neutral_dimensionless.lta_os__ts Tarski_neutral_dimensionless_axioms assms(1) one_side_symmetry os_ts1324__os) } thus ?thesis using P7 by blast qed lemma acute_one_side_aux0: assumes "Col A C P" and "Acute A C P" and "C A Perp B C" shows "C B OS A P" proof - have "Per A C B" by (simp add: assms(3) perp_per_1) then have P1: "A C P LtA A C B" using Tarski_neutral_dimensionless.acute_per__lta Tarski_neutral_dimensionless_axioms acute_distincts assms(2) assms(3) perp_not_eq_2 by fastforce have P2: "C Out A P" using acute_col__out assms(1) assms(2) by auto thus ?thesis using Perp_cases assms(3) out_one_side perp_not_col by blast qed lemma acute_cop_perp__one_side: assumes "Acute A C P" and "C A Perp B C" and "Coplanar A B C P" shows "C B OS A P" proof cases assume "Col A C P" thus ?thesis by (simp add: acute_one_side_aux0 assms(1) assms(2)) next assume P1: "\ Col A C P" have P2: "C A TS P B \ C A OS P B" using Col_cases P1 assms(2) assms(3) cop_nos__ts coplanar_perm_13 perp_not_col by blast { assume P3: "C A TS P B" obtain Bs where P4: "C Midpoint B Bs" using symmetric_point_construction by auto have "C A TS Bs B" by (metis P3 P4 assms(2) bet__ts l9_2 midpoint_bet midpoint_distinct_2 perp_not_col ts_distincts) then have P6: "C A OS P Bs" using P3 l9_8_1 by auto have "C Bs Perp A C" proof - have P6A: "C \ Bs" using P6 os_distincts by blast have "Col C B Bs" using Bet_cases Col_def P4 midpoint_bet by blast thus ?thesis using Perp_cases P6A assms(2) perp_col by blast qed then have "Bs C Perp C A" using Perp_perm by blast then have "C A Perp Bs C" using Perp_perm by blast then have "C B OS A P" using acute_one_side_aux by (metis P4 P6 assms(1) assms(2) col_one_side midpoint_col not_col_permutation_5 perp_distinct) } { assume "C A OS P B" then have "C B OS A P" using acute_one_side_aux using assms(1) assms(2) by blast } thus ?thesis using P2 \C A TS P B \ C B OS A P\ by auto qed lemma acute__not_obtuse: assumes "Acute A B C" shows "\ Obtuse A B C" using acute_obtuse__lta assms nlta by blast subsubsection "Sum of angles" lemma suma_distincts: assumes "A B C D E F SumA G H I" shows "A \ B \ B \C \ D \ E \ E \ F \ G \ H \ H \ I" proof - obtain J where "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms by auto thus ?thesis using CongA_def by blast qed lemma trisuma_distincts: assumes "A B C TriSumA D E F" shows "A \ B \ B \ C \ A \ C \ D \ E \ E \ F" proof - obtain G H I where "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms by auto thus ?thesis using suma_distincts by blast qed lemma ex_suma: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" shows "\ G H I. A B C D E F SumA G H I" proof - have "\ I. A B C D E F SumA A B I" proof cases assume P1: "Col A B C" obtain J where P2: "D E F CongA C B J \ Coplanar C B J A" using angle_construction_4 using assms(2) assms(3) assms(4) by presburger have P3: "J \ B" using CongA_def P2 by blast have "\ B C OS A J" by (metis P1 between_trivial2 one_side_chara) then have "A B C D E F SumA A B J" by (meson P2 P3 SumA_def assms(1) conga_refl ncoplanar_perm_15 not_conga_sym) thus ?thesis by blast next assume T1: "\ Col A B C" show ?thesis proof cases assume T2: "Col D E F" show ?thesis proof cases assume T3: "Bet D E F" obtain J where T4: "B Midpoint C J" using symmetric_point_construction by blast have "A B C D E F SumA A B J" proof - have "C B J CongA D E F" by (metis T3 T4 assms(2) assms(3) assms(4) conga_line midpoint_bet midpoint_distinct_2) moreover have "\ B C OS A J" by (simp add: T4 col124__nos midpoint_col) moreover have "Coplanar A B C J" using T3 bet__coplanar bet_conga__bet calculation(1) conga_sym ncoplanar_perm_15 by blast moreover have "A B J CongA A B J" using CongA_def assms(1) calculation(1) conga_refl by auto ultimately show ?thesis using SumA_def by blast qed then show ?thesis by auto next assume T5: "\ Bet D E F" have "A B C D E F SumA A B C" proof - have "E Out D F" using T2 T5 l6_4_2 by auto then have "C B C CongA D E F" using assms(2) l11_21_b out_trivial by auto moreover have "\ B C OS A C" using os_distincts by blast moreover have "Coplanar A B C C" using ncop_distincts by auto moreover have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto ultimately show ?thesis using SumA_def by blast qed then show ?thesis by auto qed next assume T6: "\ Col D E F" then obtain J where T7: "D E F CongA C B J \ C B TS J A" using T1 ex_conga_ts not_col_permutation_4 not_col_permutation_5 by presburger then show ?thesis proof - have "C B J CongA D E F" using T7 not_conga_sym by blast moreover have "\ B C OS A J" by (simp add: T7 invert_two_sides l9_2 l9_9) moreover have "Coplanar A B C J" using T7 ncoplanar_perm_15 ts__coplanar by blast moreover have "A B J CongA A B J" using T7 assms(1) conga_diff56 conga_refl by blast ultimately show ?thesis using SumA_def by blast qed qed qed then show ?thesis by auto qed lemma suma2__conga: assumes "A B C D E F SumA G H I" and "A B C D E F SumA G' H' I'" shows "G H I CongA G' H' I'" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast obtain J' where P2: "C B J' CongA D E F \ \ B C OS A J' \ Coplanar A B C J' \ A B J' CongA G' H' I'" using SumA_def assms(2) by blast have P3: "C B J CongA C B J'" proof - have "C B J CongA D E F" by (simp add: P1) moreover have "D E F CongA C B J'" by (simp add: P2 conga_sym) ultimately show ?thesis using not_conga by blast qed have P4: "A B J CongA A B J'" proof cases assume P5: "Col A B C" then show ?thesis proof cases assume P6: "Bet A B C" show ?thesis proof - have "C B J CongA C B J'" by (simp add: P3) moreover have "Bet C B A" by (simp add: P6 between_symmetry) moreover have "A \ B" using assms(1) suma_distincts by blast ultimately show ?thesis using l11_13 by blast qed next assume P7: "\ Bet A B C" moreover have "B Out A C" by (simp add: P5 calculation l6_4_2) moreover have "B \ J" using CongA_def P3 by blast then moreover have "B Out J J" using out_trivial by auto moreover have "B \ J'" using CongA_def P3 by blast then moreover have "B Out J' J'" using out_trivial by auto ultimately show ?thesis using P3 l11_10 by blast qed next assume P8: "\ Col A B C" show ?thesis proof cases assume P9: "Col D E F" have "B Out J' J" proof cases assume P10: "Bet D E F" show ?thesis proof - have "D E F CongA J' B C" using P2 conga_right_comm not_conga_sym by blast then have "Bet J' B C" using P10 bet_conga__bet by blast moreover have "D E F CongA J B C" by (simp add: P1 conga_right_comm conga_sym) then moreover have "Bet J B C" using P10 bet_conga__bet by blast ultimately show ?thesis by (metis CongA_def P3 l6_2) qed next assume P11: "\ Bet D E F" have P12: "E Out D F" by (simp add: P11 P9 l6_4_2) show ?thesis proof - have "B Out J' C" proof - have "D E F CongA J' B C" using P2 conga_right_comm conga_sym by blast then show ?thesis using l11_21_a P12 by blast qed moreover have "B Out C J" by (metis P3 P8 bet_conga__bet calculation col_conga_col col_out2_col l6_4_2 l6_6 not_col_distincts not_conga_sym out_bet_out_1 out_trivial) ultimately show ?thesis using l6_7 by blast qed qed then show ?thesis using P8 not_col_distincts out2__conga out_trivial by blast next assume P13: "\ Col D E F" show ?thesis proof - have "B C TS A J" proof - have "Coplanar B C A J" using P1 coplanar_perm_8 by blast moreover have "\ Col A B C" by (simp add: P8) moreover have "\ B C OS A J" using P1 by simp moreover have "\ Col J B C" proof - have "D E F CongA J B C" using P1 conga_right_comm not_conga_sym by blast then show ?thesis using P13 ncol_conga_ncol by blast qed ultimately show ?thesis using cop__one_or_two_sides by blast qed moreover have "B C TS A J'" proof - have "Coplanar B C A J'" using P2 coplanar_perm_8 by blast moreover have "\ Col A B C" by (simp add: P8) moreover have "\ B C OS A J'" using P2 by simp moreover have "\ Col J' B C" proof - have "D E F CongA J' B C" using P2 conga_right_comm not_conga_sym by blast then show ?thesis using P13 ncol_conga_ncol by blast qed ultimately show ?thesis using cop_nos__ts by blast qed moreover have "A B C CongA A B C" by (metis P8 conga_pseudo_refl conga_right_comm not_col_distincts) moreover have "C B J CongA C B J'" by (simp add: P3) ultimately show ?thesis using l11_22a by blast qed qed qed then show ?thesis by (meson P1 P2 not_conga not_conga_sym) qed lemma suma_sym: assumes "A B C D E F SumA G H I" shows "D E F A B C SumA G H I" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast show ?thesis proof cases assume P2: "Col A B C" then show ?thesis proof cases assume P3: "Bet A B C" obtain K where P4: "Bet F E K \ Cong F E E K" using Cong_perm segment_construction by blast show ?thesis proof - have P5: "F E K CongA A B C" by (metis CongA_def P1 P3 P4 cong_diff conga_line) moreover have "\ E F OS D K" using P4 bet_col col124__nos invert_one_side by blast moreover have "Coplanar D E F K" using P4 bet__coplanar ncoplanar_perm_15 by blast moreover have "D E K CongA G H I" proof - have "D E K CongA A B J" proof - have "F E D CongA C B J" by (simp add: P1 conga_left_comm conga_sym) moreover have "Bet F E K" by (simp add: P4) moreover have "K \ E" using P4 calculation(1) cong_identity conga_diff1 by blast moreover have "Bet C B A" by (simp add: Bet_perm P3) moreover have "A \ B" using CongA_def P5 by blast ultimately show ?thesis using conga_right_comm l11_13 not_conga_sym by blast qed then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed next assume T1: "\ Bet A B C" then have T2: "B Out A C" by (simp add: P2 l6_4_2) show ?thesis proof - have "F E F CongA A B C" by (metis T2 assms l11_21_b out_trivial suma_distincts) moreover have "\ E F OS D F" using os_distincts by auto moreover have "Coplanar D E F F" using ncop_distincts by auto moreover have "D E F CongA G H I" proof - have "A B J CongA D E F" proof - have "C B J CongA D E F" by (simp add: P1) moreover have "B Out A C" by (simp add: T2) moreover have "J \ B" using calculation(1) conga_distinct by auto moreover have "D \ E" using calculation(1) conga_distinct by blast moreover have "F \ E" using calculation(1) conga_distinct by blast ultimately show ?thesis by (meson Out_cases not_conga out2__conga out_trivial) qed then have "D E F CongA A B J" using not_conga_sym by blast then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed qed next assume Q1: "\ Col A B C" show ?thesis proof cases assume Q2: "Col D E F" obtain K where Q3: "A B C CongA F E K" using P1 angle_construction_3 conga_diff1 conga_diff56 by fastforce show ?thesis proof - have "F E K CongA A B C" by (simp add: Q3 conga_sym) moreover have "\ E F OS D K" using Col_cases Q2 one_side_not_col123 by blast moreover have "Coplanar D E F K" by (simp add: Q2 col__coplanar) moreover have "D E K CongA G H I" proof - have "D E K CongA A B J" proof cases assume "Bet D E F" then have "J B A CongA D E K" by (metis P1 bet_conga__bet calculation(1) conga_diff45 conga_right_comm l11_13 not_conga_sym) then show ?thesis using conga_right_comm not_conga_sym by blast next assume "\ Bet D E F" then have W2: "E Out D F" using Q2 or_bet_out by blast have "A B J CongA D E K" proof - have "A B C CongA F E K" by (simp add: Q3) moreover have "A \ B" using Q1 col_trivial_1 by auto moreover have "E Out D F" by (simp add: W2) moreover have "B Out J C" proof - have "D E F CongA J B C" by (simp add: P1 conga_left_comm conga_sym) then show ?thesis using W2 out_conga_out by blast qed moreover have "K \ E" using CongA_def Q3 by blast ultimately show ?thesis using l11_10 out_trivial by blast qed then show ?thesis using not_conga_sym by blast qed then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed next assume W3: "\ Col D E F" then obtain K where W4: "A B C CongA F E K \ F E TS K D" using Q1 ex_conga_ts not_col_permutation_3 by blast show ?thesis proof - have "F E K CongA A B C" using W4 not_conga_sym by blast moreover have "\ E F OS D K" proof - have "E F TS D K" using W4 invert_two_sides l9_2 by blast then show ?thesis using l9_9 by auto qed moreover have "Coplanar D E F K" proof - have "E F TS D K" using W4 invert_two_sides l9_2 by blast then show ?thesis using ncoplanar_perm_8 ts__coplanar by blast qed moreover have "D E K CongA G H I" proof - have "A B J CongA K E D" proof - have "B C TS A J" proof - have "Coplanar B C A J" using P1 ncoplanar_perm_12 by blast moreover have "\ Col A B C" by (simp add: Q1) moreover have "\ B C OS A J" using P1 by simp moreover have "\ Col J B C" proof - { assume "Col J B C" have "Col D E F" proof - have "Col C B J" using Col_perm \Col J B C\ by blast moreover have "C B J CongA D E F" by (simp add: P1) ultimately show ?thesis using col_conga_col by blast qed then have "False" by (simp add: W3) } then show ?thesis by blast qed ultimately show ?thesis using cop_nos__ts by blast qed moreover have "E F TS K D" using W4 invert_two_sides by blast moreover have "A B C CongA K E F" by (simp add: W4 conga_right_comm) moreover have "C B J CongA F E D" by (simp add: P1 conga_right_comm) ultimately show ?thesis using l11_22a by auto qed then have "D E K CongA A B J" using conga_left_comm not_conga_sym by blast then show ?thesis using P1 not_conga by blast qed ultimately show ?thesis using SumA_def by blast qed qed qed qed lemma conga3_suma__suma: assumes "A B C D E F SumA G H I" and "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "G H I CongA G' H' I'" shows "A' B' C' D' E' F' SumA G' H' I'" proof - have "D' E' F' A B C SumA G' H' I'" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast have "A B C D' E' F' SumA G' H' I'" proof - have "C B J CongA D' E' F'" using P1 assms(3) not_conga by blast moreover have "\ B C OS A J" using P1 by simp moreover have "Coplanar A B C J" using P1 by simp moreover have "A B J CongA G' H' I'" using P1 assms(4) not_conga by blast ultimately show ?thesis using SumA_def by blast qed then show ?thesis by (simp add: suma_sym) qed then obtain J where P2: "F' E' J CongA A B C \ \ E' F' OS D' J \ Coplanar D' E' F' J \ D' E' J CongA G' H' I'" using SumA_def by blast have "D' E' F' A' B' C' SumA G' H' I'" proof - have "F' E' J CongA A' B' C'" proof - have "F' E' J CongA A B C" by (simp add: P2) moreover have "A B C CongA A' B' C'" by (simp add: assms(2)) ultimately show ?thesis using not_conga by blast qed moreover have "\ E' F' OS D' J" using P2 by simp moreover have "Coplanar D' E' F' J" using P2 by simp moreover have "D' E' J CongA G' H' I'" by (simp add: P2) ultimately show ?thesis using SumA_def by blast qed then show ?thesis by (simp add: suma_sym) qed lemma out6_suma__suma: assumes "A B C D E F SumA G H I" and "B Out A A'" and "B Out C C'" and "E Out D D'" and "E Out F F'" and "H Out G G'" and "H Out I I'" shows "A' B C' D' E F' SumA G' H I'" proof - have "A B C CongA A' B C'" using Out_cases assms(2) assms(3) out2__conga by blast moreover have "D E F CongA D' E F'" using Out_cases assms(4) assms(5) out2__conga by blast moreover have "G H I CongA G' H I'" by (simp add: assms(6) assms(7) l6_6 out2__conga) ultimately show ?thesis using assms(1) conga3_suma__suma by blast qed lemma out546_suma__conga: assumes "A B C D E F SumA G H I" and "E Out D F" shows "A B C CongA G H I" proof - have "A B C D E F SumA A B C" proof - have "C B C CongA D E F" by (metis assms(1) assms(2) l11_21_b out_trivial suma_distincts) moreover have "\ B C OS A C" using os_distincts by auto moreover have "Coplanar A B C C" using ncop_distincts by auto moreover have "A B C CongA A B C" by (metis Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms assms(1) conga_pseudo_refl conga_right_comm) ultimately show ?thesis using SumA_def by blast qed then show ?thesis using suma2__conga assms(1) by blast qed lemma out546__suma: assumes "A \ B" and "B \ C" and "E Out D F" shows "A B C D E F SumA A B C" proof - have P1: "D \ E" using assms(3) out_diff1 by auto have P2: "F \ E" using Out_def assms(3) by auto then obtain G H I where P3: "A B C D E F SumA G H I" using P1 assms(1) assms(2) ex_suma by presburger then have "G H I CongA A B C" by (meson Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.out546_suma__conga Tarski_neutral_dimensionless_axioms assms(3)) then show ?thesis using P1 P2 P3 assms(1) assms(2) assms(3) conga3_suma__suma conga_refl out_diff1 by auto qed lemma out213_suma__conga: assumes "A B C D E F SumA G H I" and "B Out A C" shows "D E F CongA G H I" using assms(1) assms(2) out546_suma__conga suma_sym by blast lemma out213__suma: assumes "D \ E" and "E \ F" and "B Out A C" shows "A B C D E F SumA D E F" by (simp add: assms(1) assms(2) assms(3) out546__suma suma_sym) lemma suma_left_comm: assumes "A B C D E F SumA G H I" shows "C B A D E F SumA G H I" proof - have "A B C CongA C B A" using assms conga_pseudo_refl suma_distincts by fastforce moreover have "D E F CongA D E F" by (metis assms conga_refl suma_distincts) moreover have "G H I CongA G H I" by (metis assms conga_refl suma_distincts) ultimately show ?thesis using assms conga3_suma__suma by blast qed lemma suma_middle_comm: assumes "A B C D E F SumA G H I" shows "A B C F E D SumA G H I" using assms suma_left_comm suma_sym by blast lemma suma_right_comm: assumes "A B C D E F SumA G H I" shows "A B C D E F SumA I H G" proof - have "A B C CongA A B C" using assms conga_refl suma_distincts by fastforce moreover have "D E F CongA D E F" by (metis assms conga_refl suma_distincts) moreover have "G H I CongA I H G" by (meson Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms assms) ultimately show ?thesis using assms conga3_suma__suma by blast qed lemma suma_comm: assumes "A B C D E F SumA G H I" shows "C B A F E D SumA I H G" by (simp add: assms suma_left_comm suma_middle_comm suma_right_comm) lemma ts__suma: assumes "A B TS C D" shows "C B A A B D SumA C B D" proof - have "A B D CongA A B D" by (metis Tarski_neutral_dimensionless.conga_right_comm Tarski_neutral_dimensionless_axioms assms conga_pseudo_refl ts_distincts) moreover have "\ B A OS C D" using assms invert_one_side l9_9 by blast moreover have "Coplanar C B A D" using assms ncoplanar_perm_14 ts__coplanar by blast moreover have "C B D CongA C B D" by (metis assms conga_refl ts_distincts) ultimately show ?thesis using SumA_def by blast qed lemma ts__suma_1: assumes "A B TS C D" shows "C A B B A D SumA C A D" by (simp add: assms invert_two_sides ts__suma) lemma inangle__suma: assumes "P InAngle A B C" shows "A B P P B C SumA A B C" proof - have "Coplanar A B P C" by (simp add: assms coplanar_perm_8 inangle__coplanar) moreover have "\ B P OS A C" by (meson assms col123__nos col124__nos in_angle_two_sides invert_two_sides l9_9_bis not_col_permutation_5) ultimately show ?thesis using SumA_def assms conga_refl inangle_distincts by blast qed lemma bet__suma: assumes "A \ B" and "B \ C" and "P \ B" and "Bet A B C" shows "A B P P B C SumA A B C" proof - have "P InAngle A B C" using assms(1) assms(2) assms(3) assms(4) in_angle_line by auto then show ?thesis by (simp add: inangle__suma) qed lemma sams_chara: assumes "A \ B" and "A' \ B" and "Bet A B A'" shows "SAMS A B C D E F \ D E F LeA C B A'" proof - { assume T1: "SAMS A B C D E F" obtain J where T2: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def T1 by auto have T3: "A \ A'" using assms(2) assms(3) between_identity by blast have T4: "C \ B" using T2 conga_distinct by blast have T5: "J \ B" using T2 conga_diff2 by blast have T6: "D \ E" using CongA_def T2 by auto have T7: "F \ E" using CongA_def T2 by blast { assume "E Out D F" then have "D E F LeA C B A'" by (simp add: T4 assms(2) l11_31_1) } { assume T8: "\ Bet A B C" have "D E F LeA C B A'" proof cases assume "Col A B C" then have "Bet C B A'" using T8 assms(1) assms(3) between_exchange3 outer_transitivity_between2 third_point by blast then show ?thesis by (simp add: T4 T6 T7 assms(2) l11_31_2) next assume T9: "\ Col A B C" show ?thesis proof cases assume T10: "Col D E F" show ?thesis proof cases assume T11: "Bet D E F" have "D E F CongA C B J" by (simp add: T2 conga_sym) then have T12: "Bet C B J" using T11 bet_conga__bet by blast have "A B TS C J" proof - have "\ Col J A B" using T5 T9 T12 bet_col col2__eq col_permutation_1 by blast moreover have "\ T. Col T A B \ Bet C T J" using T12 col_trivial_3 by blast ultimately show ?thesis using T9 TS_def col_permutation_1 by blast qed then have "False" using T2 by simp then show ?thesis by simp next assume "\ Bet D E F" then show ?thesis using T10 \E Out D F \ D E F LeA C B A'\ or_bet_out by auto qed next assume T13: "\ Col D E F" show ?thesis proof - have "C B J LeA C B A'" proof - have "J InAngle C B A'" proof - have "A' \ B" by (simp add: assms(2)) moreover have "Bet A B A'" by (simp add: assms(3)) moreover have "C InAngle A B J" proof - have "\ Col J B C" proof - have "\ Col D E F" by (simp add: T13) moreover have "D E F CongA J B C" using T2 conga_left_comm not_conga_sym by blast ultimately show ?thesis using ncol_conga_ncol by blast qed then have "B C TS A J" by (simp add: T2 T9 cop_nos__ts coplanar_perm_8) then obtain X where T14: "Col X B C \ Bet A X J" using TS_def by blast { assume T15: "X \ B" have "B Out X C" proof - have "Col B X C" by (simp add: Col_perm T14) moreover have "B A OS X C" proof - have "A B OS X C" proof - have "A B OS X J" by (smt T14 T9 T15 bet_out calculation col_transitivity_2 col_trivial_2 l6_21 out_one_side) moreover have "A B OS J C" by (metis T14 T2 T9 calculation cop_nts__os l5_2 not_col_permutation_2 one_side_chara one_side_symmetry) ultimately show ?thesis using one_side_transitivity by blast qed then show ?thesis by (simp add: invert_one_side) qed ultimately show ?thesis using col_one_side_out by auto qed } then have "Bet A X J \ (X = B \ B Out X C)" using T14 by blast then show ?thesis using InAngle_def T4 T5 assms(1) by auto qed ultimately show ?thesis using in_angle_reverse l11_24 by blast qed moreover have "C B J CongA C B J" by (simp add: T4 T5 conga_refl) ultimately show ?thesis by (simp add: inangle__lea) qed moreover have "D E F LeA C B J" by (simp add: T2 conga__lea456123) ultimately show ?thesis using lea_trans by blast qed qed qed } then have "D E F LeA C B A'" using SAMS_def T1 \E Out D F \ D E F LeA C B A'\ by blast } { assume P1: "D E F LeA C B A'" have P2: "A \ A'" using assms(2) assms(3) between_identity by blast have P3: "C \ B" using P1 lea_distincts by auto have P4: "D \ E" using P1 lea_distincts by auto have P5: "F \ E" using P1 lea_distincts by auto have "SAMS A B C D E F" proof cases assume P6: "Col A B C" show ?thesis proof cases assume P7: "Bet A B C" have "E Out D F" proof - have "B Out C A'" by (meson Bet_perm P3 P7 assms(1) assms(2) assms(3) l6_2) moreover have "C B A' CongA D E F" using P1 calculation l11_21_b out_lea__out by blast ultimately show ?thesis using out_conga_out by blast qed moreover have "C B C CongA D E F" using P3 calculation l11_21_b out_trivial by auto moreover have "\ B C OS A C" using os_distincts by auto moreover have "\ A B TS C C" by (simp add: not_two_sides_id) moreover have "Coplanar A B C C" using ncop_distincts by auto ultimately show ?thesis using SAMS_def assms(1) by blast next assume P8: "\ Bet A B C" have P9: "B Out A C" by (simp add: P6 P8 l6_4_2) obtain J where P10: "D E F CongA C B J" using P3 P4 P5 angle_construction_3 by blast show ?thesis proof - have "C B J CongA D E F" using P10 not_conga_sym by blast moreover have "\ B C OS A J" using Col_cases P6 one_side_not_col123 by blast moreover have "\ A B TS C J" using Col_cases P6 TS_def by blast moreover have "Coplanar A B C J" using P6 col__coplanar by auto ultimately show ?thesis using P8 SAMS_def assms(1) by blast qed qed next assume P11: "\ Col A B C" have P12: "\ Col A' B C" using P11 assms(2) assms(3) bet_col bet_col1 colx by blast show ?thesis proof cases assume P13: "Col D E F" have P14: "E Out D F" proof - { assume P14: "Bet D E F" have "D E F LeA C B A'" by (simp add: P1) then have "Bet C B A'" using P14 bet_lea__bet by blast then have "Col A' B C" using Col_def Col_perm by blast then have "False" by (simp add: P12) } then have "\ Bet D E F" by auto then show ?thesis by (simp add: P13 l6_4_2) qed show ?thesis proof - have "C B C CongA D E F" by (simp add: P3 P14 l11_21_b out_trivial) moreover have "\ B C OS A C" using os_distincts by auto moreover have "\ A B TS C C" by (simp add: not_two_sides_id) moreover have "Coplanar A B C C" using ncop_distincts by auto ultimately show ?thesis using P14 SAMS_def assms(1) by blast qed next assume P15: "\ Col D E F" obtain J where P16: "D E F CongA C B J \ C B TS J A" using P11 P15 ex_conga_ts not_col_permutation_3 by presburger show ?thesis proof - have "C B J CongA D E F" by (simp add: P16 conga_sym) moreover have "\ B C OS A J" proof - have "C B TS A J" using P16 by (simp add: l9_2) then show ?thesis using invert_one_side l9_9 by blast qed moreover have "\ A B TS C J \ Coplanar A B C J" proof cases assume "Col A B J" then show ?thesis using TS_def ncop__ncols not_col_permutation_1 by blast next assume P17: "\ Col A B J" have "\ A B TS C J" proof - have "A' B OS J C" proof - have "\ Col A' B C" by (simp add: P12) moreover have "\ Col B A' J" proof - { assume "Col B A' J" then have "False" by (metis P17 assms(2) assms(3) bet_col col_trivial_2 colx) } then show ?thesis by auto qed moreover have "J InAngle A' B C" proof - obtain K where P20: "K InAngle C B A' \ D E F CongA C B K" using LeA_def P1 by blast have "J InAngle C B A'" proof - have "C B A' CongA C B A'" by (simp add: P3 assms(2) conga_pseudo_refl conga_right_comm) moreover have "C B K CongA C B J" proof - have "C B K CongA D E F" using P20 not_conga_sym by blast moreover have "D E F CongA C B J" by (simp add: P16) ultimately show ?thesis using not_conga by blast qed moreover have "K InAngle C B A'" using P20 by simp moreover have "C B OS J A'" proof - have "C B TS J A" using P16 by simp moreover have "C B TS A' A" using Col_perm P12 assms(3) bet__ts between_symmetry calculation invert_two_sides ts_distincts by blast ultimately show ?thesis using OS_def by auto qed ultimately show ?thesis using conga_preserves_in_angle by blast qed then show ?thesis by (simp add: l11_24) qed ultimately show ?thesis by (simp add: in_angle_one_side) qed then have "A' B OS C J" by (simp add: one_side_symmetry) then have "\ A' B TS C J" by (simp add: l9_9_bis) then show ?thesis using assms(2) assms(3) bet_col bet_col1 col_preserves_two_sides by blast qed moreover have "Coplanar A B C J" proof - have "C B TS J A" using P16 by simp then show ?thesis by (simp add: coplanar_perm_20 ts__coplanar) qed ultimately show ?thesis by auto qed ultimately show ?thesis using P11 SAMS_def assms(1) bet_col by auto qed qed qed } then show ?thesis using \SAMS A B C D E F \ D E F LeA C B A'\ by blast qed lemma sams_distincts: assumes "SAMS A B C D E F" shows "A \ B \ B \ C \ D \ E \ E \ F" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms by auto then show ?thesis by (metis SAMS_def assms conga_distinct) qed lemma sams_sym: assumes "SAMS A B C D E F" shows "SAMS D E F A B C" proof - have P1: "A \ B" using assms sams_distincts by blast have P3: "D \ E" using assms sams_distincts by blast obtain D' where P5: "E Midpoint D D'" using symmetric_point_construction by blast obtain A' where P6: "B Midpoint A A'" using symmetric_point_construction by blast have P8: "E \ D'" using P3 P5 is_midpoint_id_2 by blast have P9: "A \ A'" using P1 P6 l7_3 by auto then have P10: "B \ A'" using P6 P9 midpoint_not_midpoint by auto then have "D E F LeA C B A'" using P1 P6 assms midpoint_bet sams_chara by fastforce then have "D E F LeA A' B C" by (simp add: lea_right_comm) then have "A B C LeA D' E F" by (metis Mid_cases P1 P10 P3 P5 P6 P8 l11_36 midpoint_bet) then have "A B C LeA F E D'" by (simp add: lea_right_comm) moreover have "D \ E" by (simp add: P3) moreover have "D' \ E" using P8 by auto moreover have "Bet D E D'" by (simp add: P5 midpoint_bet) then show ?thesis using P3 P8 calculation(1) sams_chara by fastforce qed lemma sams_right_comm: assumes "SAMS A B C D E F" shows "SAMS A B C F E D" proof - have P1: "E Out D F \ \ Bet A B C" using SAMS_def assms by blast obtain J where P2: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms by auto { assume "E Out D F" then have "E Out F D \ \ Bet A B C" by (simp add: l6_6) } { assume "\ Bet A B C" then have "E Out F D \ \ Bet A B C" by auto } then have "E Out F D \ \ Bet A B C" using \E Out D F \ E Out F D \ \ Bet A B C\ P1 by auto moreover have "C B J CongA F E D" proof - have "C B J CongA D E F" by (simp add: P2) then show ?thesis by (simp add: conga_right_comm) qed ultimately show ?thesis using P2 SAMS_def assms by auto qed lemma sams_left_comm: assumes "SAMS A B C D E F" shows "SAMS C B A D E F" proof - have "SAMS D E F A B C" by (simp add: assms sams_sym) then have "SAMS D E F C B A" using sams_right_comm by blast then show ?thesis using sams_sym by blast qed lemma sams_comm: assumes "SAMS A B C D E F" shows "SAMS C B A F E D" using assms sams_left_comm sams_right_comm by blast lemma conga2_sams__sams: assumes "A B C CongA A' B' C'" and "D E F CongA D' E' F'" and "SAMS A B C D E F" shows "SAMS A' B' C' D' E' F'" proof - obtain A0 where "B Midpoint A A0" using symmetric_point_construction by auto obtain A'0 where "B' Midpoint A' A'0" using symmetric_point_construction by blast have "D' E' F' LeA C' B' A'0" proof - have "D E F LeA C B A0" by (metis \B Midpoint A A0\ assms(1) assms(3) conga_distinct midpoint_bet midpoint_distinct_2 sams_chara) moreover have "D E F CongA D' E' F'" by (simp add: assms(2)) moreover have "C B A0 CongA C' B' A'0" proof - have "A0 B C CongA A'0 B' C'" by (metis \B Midpoint A A0\ \B' Midpoint A' A'0\ assms(1) calculation(1) conga_diff45 l11_13 lea_distincts midpoint_bet midpoint_not_midpoint) then show ?thesis using conga_comm by blast qed ultimately show ?thesis using l11_30 by blast qed then show ?thesis by (metis \B' Midpoint A' A'0\ assms(1) conga_distinct lea_distincts midpoint_bet sams_chara) qed lemma out546__sams: assumes "A \ B" and "B \ C" and "E Out D F" shows "SAMS A B C D E F" proof - obtain A' where "Bet A B A' \ Cong B A' A B" using segment_construction by blast then have "D E F LeA C B A'" using assms(1) assms(2) assms(3) cong_diff_3 l11_31_1 by fastforce then show ?thesis using \Bet A B A' \ Cong B A' A B\ assms(1) lea_distincts sams_chara by blast qed lemma out213__sams: assumes "D \ E" and "E \ F" and "B Out A C" shows "SAMS A B C D E F" by (simp add: Tarski_neutral_dimensionless.sams_sym Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) out546__sams) lemma bet_suma__sams: assumes "A B C D E F SumA G H I" and "Bet G H I" shows "SAMS A B C D E F" proof - obtain A' where P1: "C B A' CongA D E F \ \ B C OS A A' \ Coplanar A B C A' \ A B A' CongA G H I" using SumA_def assms(1) by auto then have "G H I CongA A B A'" using not_conga_sym by blast then have "Bet A B A'" using assms(2) bet_conga__bet by blast show ?thesis proof - have "E Out D F \ \ Bet A B C" proof - { assume "Bet A B C" then have "E Out D F" proof - have "B Out C A'" proof - have "C \ B" using assms(1) suma_distincts by blast moreover have "A' \ B" using CongA_def \G H I CongA A B A'\ by blast moreover have "A \ B" using CongA_def \G H I CongA A B A'\ by blast moreover have "Bet C B A" by (simp add: Bet_perm \Bet A B C\) ultimately show ?thesis using Out_def \Bet A B A'\ \Bet A B C\ l5_2 by auto qed moreover have "C B A' CongA D E F" using P1 by simp ultimately show ?thesis using l11_21_a by blast qed } then show ?thesis by blast qed moreover have "\ J. (C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "C B A' CongA D E F" by (simp add: P1) moreover have "\ B C OS A A'" by (simp add: P1) moreover have "\ A B TS C A'" using Col_def TS_def \Bet A B A'\ by blast moreover have "Coplanar A B C A'" by (simp add: P1) ultimately show ?thesis by blast qed ultimately show ?thesis using CongA_def SAMS_def \C B A' CongA D E F \ \ B C OS A A' \ Coplanar A B C A' \ A B A' CongA G H I\ by auto qed qed lemma bet__sams: assumes "A \ B" and "B \ C" and "P \ B" and "Bet A B C" shows "SAMS A B P P B C" by (meson assms(1) assms(2) assms(3) assms(4) bet__suma bet_suma__sams) lemma suppa__sams: assumes "A B C SuppA D E F" shows "SAMS A B C D E F" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms by auto then have "SAMS A B C C B A'" by (metis assms bet__sams conga_diff45 conga_diff56 suppa2__conga123) thus ?thesis by (meson P1 assms conga2_sams__sams not_conga_sym suppa2__conga123) qed lemma os_ts__sams: assumes "B P TS A C" and "A B OS P C" shows "SAMS A B P P B C" proof - have "B Out P C \ \ Bet A B P" using assms(2) bet_col col123__nos by blast moreover have "\ J. (P B J CongA P B C \ \ B P OS A J \ \ A B TS P J \ Coplanar A B P J)" by (metis assms(1) assms(2) conga_refl l9_9 os__coplanar os_distincts) ultimately show ?thesis using SAMS_def assms(2) os_distincts by auto qed lemma os2__sams: assumes "A B OS P C" and "C B OS P A" shows "SAMS A B P P B C" by (simp add: Tarski_neutral_dimensionless.os_ts__sams Tarski_neutral_dimensionless_axioms assms(1) assms(2) invert_one_side l9_31) lemma inangle__sams: assumes "P InAngle A B C" shows "SAMS A B P P B C" proof - have "Bet A B C \ B Out A C \ \ Col A B C" using l6_4_2 by blast { assume "Bet A B C" then have "SAMS A B P P B C" using assms bet__sams inangle_distincts by fastforce } { assume "B Out A C" then have "SAMS A B P P B C" by (metis assms in_angle_out inangle_distincts out213__sams) } { assume "\ Col A B C" then have "\ Bet A B C" using Col_def by auto { assume "Col B A P" have "SAMS A B P P B C" by (metis \Col B A P\ \\ Bet A B C\ assms col_in_angle_out inangle_distincts out213__sams) } { assume "\ Col B A P" { assume "Col B C P" have "SAMS A B P P B C" by (metis Tarski_neutral_dimensionless.sams_comm Tarski_neutral_dimensionless_axioms \Col B C P\ \\ Bet A B C\ assms between_symmetry col_in_angle_out inangle_distincts l11_24 out546__sams) } { assume "\ Col B C P" have "SAMS A B P P B C" proof - have "B P TS A C" by (simp add: \\ Col B A P\ \\ Col B C P\ assms in_angle_two_sides invert_two_sides) moreover have "A B OS P C" by (simp add: \\ Col A B C\ \\ Col B A P\ assms in_angle_one_side) ultimately show ?thesis by (simp add: os_ts__sams) qed } then have "SAMS A B P P B C" using \Col B C P \ SAMS A B P P B C\ by blast } then have "SAMS A B P P B C" using \Col B A P \ SAMS A B P P B C\ by blast } thus ?thesis using \B Out A C \ SAMS A B P P B C\ \Bet A B C \ SAMS A B P P B C\ \Bet A B C \ B Out A C \ \ Col A B C\ by blast qed lemma sams_suma__lea123789: assumes "A B C D E F SumA G H I" and "SAMS A B C D E F" shows "A B C LeA G H I" proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A B C" have P1: "(A \ B \ (E Out D F \ \ Bet A B C)) \ (\ J. (C B J CongA D E F \ \ (B C OS A J) \ \ (A B TS C J) \ Coplanar A B C J))" using SAMS_def assms(2) by auto { assume "E Out D F" then have "A B C CongA G H I" using assms(1) out546_suma__conga by auto then have "A B C LeA G H I" by (simp add: conga__lea) } thus ?thesis using P1 \Bet A B C\ by blast next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ or_bet_out by auto thus ?thesis by (metis assms(1) l11_31_1 suma_distincts) qed next assume "\ Col A B C" show ?thesis proof cases assume "Col D E F" show ?thesis proof cases assume "Bet D E F" have "SAMS D E F A B C" using assms(2) sams_sym by auto then have "B Out A C" using SAMS_def \Bet D E F\ by blast thus ?thesis using l11_31_1 by (metis assms(1) suma_distincts) next assume "\ Bet D E F" have "A B C CongA G H I" proof - have "A B C D E F SumA G H I" by (simp add: assms(1)) moreover have "E Out D F" using \Col D E F\ \\ Bet D E F\ l6_4_2 by auto ultimately show ?thesis using out546_suma__conga by auto qed show ?thesis by (simp add: \A B C CongA G H I\ conga__lea) qed next assume "\ Col D E F" show ?thesis proof cases assume "Col G H I" show ?thesis proof cases assume "Bet G H I" thus ?thesis by (metis assms(1) l11_31_2 suma_distincts) next assume "\ Bet G H I" then have "H Out G I" by (simp add: \Col G H I\ l6_4_2) have "E Out D F \ \ Bet A B C" using \\ Col A B C\ bet_col by auto { assume "\ Bet A B C" then obtain J where P2: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(1) by blast have "G H I CongA A B J" using P2 not_conga_sym by blast then have "B Out A J" using \H Out G I\ out_conga_out by blast then have "B C OS A J" using Col_perm \\ Col A B C\ out_one_side by blast then have "False" using \C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I\ by linarith } then have "False" using Col_def \\ Col A B C\ by blast thus ?thesis by blast qed next assume "\ Col G H I" obtain J where P4: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using SAMS_def assms(2) by auto { assume "Col J B C" have "J B C CongA D E F" by (simp add: P4 conga_left_comm) then have "Col D E F" using col_conga_col \Col J B C\ by blast then have "False" using \\ Col D E F\ by blast } then have "\ Col J B C" by blast have "A B J CongA G H I" proof - have "A B C D E F SumA A B J" proof - have "C B J CongA D E F" using P4 by simp moreover have "\ B C OS A J" by (simp add: P4) moreover have "Coplanar A B C J" by (simp add: P4) moreover have "A B J CongA A B J" by (metis \\ Col A B C\ \\ Col J B C\ col_trivial_1 conga_refl) ultimately show ?thesis using SumA_def by blast qed then show ?thesis using assms(1) suma2__conga by auto qed have "\ Col J B A" using \A B J CongA G H I\ \\ Col G H I\ col_conga_col not_col_permutation_3 by blast have "A B C LeA A B J" proof - have "C InAngle A B J" by (metis Col_perm P4 \\ Col A B C\ \\ Col J B A\ \\ Col J B C\ cop_nos__ts coplanar_perm_7 coplanar_perm_8 invert_two_sides l9_2 os_ts__inangle) moreover have "A B C CongA A B C" using calculation in_angle_asym inangle3123 inangle_distincts by auto ultimately show ?thesis using inangle__lea by blast qed thus ?thesis using \A B J CongA G H I\ conga__lea lea_trans by blast qed qed qed lemma sams_suma__lea456789: assumes "A B C D E F SumA G H I" and "SAMS A B C D E F" shows "D E F LeA G H I" proof - have "D E F A B C SumA G H I" by (simp add: assms(1) suma_sym) moreover have "SAMS D E F A B C" using assms(2) sams_sym by blast ultimately show ?thesis using sams_suma__lea123789 by auto qed lemma sams_lea2__sams: assumes "SAMS A' B' C' D' E' F'" and "A B C LeA A' B' C'" and "D E F LeA D' E' F'" shows "SAMS A B C D E F" proof - obtain A0 where "B Midpoint A A0" using symmetric_point_construction by auto obtain A'0 where "B' Midpoint A' A'0" using symmetric_point_construction by auto have "D E F LeA C B A0" proof - have "D' E' F' LeA C B A0" proof - have "D' E' F' LeA C' B' A'0" by (metis \B' Midpoint A' A'0\ assms(1) assms(2) lea_distincts midpoint_bet midpoint_distinct_2 sams_chara) moreover have "C' B' A'0 LeA C B A0" by (metis Mid_cases \B Midpoint A A0\ \B' Midpoint A' A'0\ assms(2) l11_36_aux2 l7_3_2 lea_comm lea_distincts midpoint_bet sym_preserve_diff) ultimately show ?thesis using lea_trans by blast qed moreover have "D E F LeA D' E' F'" using assms(3) by auto ultimately show ?thesis using \D' E' F' LeA C B A0\ assms(3) lea_trans by blast qed then show ?thesis by (metis \B Midpoint A A0\ assms(2) lea_distincts midpoint_bet sams_chara) qed lemma sams_lea456_suma2__lea: assumes "D E F LeA D' E' F'" and "SAMS A B C D' E' F'" and "A B C D E F SumA G H I" and "A B C D' E' F' SumA G' H' I'" shows "G H I LeA G' H' I'" proof cases assume "E' Out D' F'" have "G H I CongA G' H' I'" proof - have "G H I CongA A B C" proof - have "A B C D E F SumA G H I" by (simp add: assms(3)) moreover have "E Out D F" using \E' Out D' F'\ assms(1) out_lea__out by blast ultimately show ?thesis using conga_sym out546_suma__conga by blast qed moreover have "A B C CongA G' H' I'" using \E' Out D' F'\ assms(4) out546_suma__conga by blast ultimately show ?thesis using conga_trans by blast qed thus ?thesis by (simp add: conga__lea) next assume T1: "\ E' Out D' F'" show ?thesis proof cases assume T2: "Col A B C" have "E' Out D' F' \ \ Bet A B C" using assms(2) SAMS_def by simp { assume "\ Bet A B C" then have "B Out A C" by (simp add: T2 l6_4_2) have "G H I LeA G' H' I'" proof - have "D E F LeA D' E' F'" by (simp add: assms(1)) moreover have "D E F CongA G H I" using \B Out A C\ assms(3) out213_suma__conga by auto moreover have "D' E' F' CongA G' H' I'" using \B Out A C\ assms(4) out213_suma__conga by auto ultimately show ?thesis using l11_30 by blast qed } thus ?thesis using T1 \E' Out D' F' \ \ Bet A B C\ by auto next assume "\ Col A B C" show ?thesis proof cases assume "Col D' E' F'" have "SAMS D' E' F' A B C" by (simp add: assms(2) sams_sym) { assume "\ Bet D' E' F'" then have "G H I LeA G' H' I'" using not_bet_out T1 \Col D' E' F'\ by auto } thus ?thesis by (metis assms(2) assms(3) assms(4) bet_lea__bet l11_31_2 sams_suma__lea456789 suma_distincts) next assume "\ Col D' E' F'" show ?thesis proof cases assume "Col D E F" have "\ Bet D E F" using bet_lea__bet Col_def \\ Col D' E' F'\ assms(1) by blast thus ?thesis proof - have "A B C LeA G' H' I'" using assms(2) assms(4) sams_suma__lea123789 by auto moreover have "A B C CongA G H I" by (meson \Col D E F\ \\ Bet D E F\ assms(3) or_bet_out out213_suma__conga suma_sym) moreover have "G' H' I' CongA G' H' I'" proof - have "G' \ H'" using calculation(1) lea_distincts by blast moreover have "H' \ I'" using \A B C LeA G' H' I'\ lea_distincts by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using l11_30 by blast qed next assume "\ Col D E F" show ?thesis proof cases assume "Col G' H' I'" show ?thesis proof cases assume "Bet G' H' I'" show ?thesis proof - have "G \ H" using assms(3) suma_distincts by auto moreover have "I \ H" using assms(3) suma_distincts by blast moreover have "G' \ H'" using assms(4) suma_distincts by auto moreover have "I' \ H'" using assms(4) suma_distincts by blast ultimately show ?thesis by (simp add: \Bet G' H' I'\ l11_31_2) qed next assume "\ Bet G' H' I'" have "B Out A C" proof - have "H' Out G' I'" using \Col G' H' I'\ l6_4_2 by (simp add: \\ Bet G' H' I'\) moreover have "A B C LeA G' H' I'" using sams_suma__lea123789 using assms(2) assms(4) by auto ultimately show ?thesis using out_lea__out by blast qed then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto thus ?thesis by blast qed next assume "\ Col G' H' I'" obtain F'1 where P5: "C B F'1 CongA D' E' F' \ \ B C OS A F'1 \ \ A B TS C F'1 \ Coplanar A B C F'1" using SAMS_def assms(2) by auto have P6: "D E F LeA C B F'1" proof - have "D E F CongA D E F" using \\ Col D E F\ conga_refl not_col_distincts by fastforce moreover have "D' E' F' CongA C B F'1" by (simp add: P5 conga_sym) ultimately show ?thesis using assms(1) l11_30 by blast qed then obtain F1 where P6: "F1 InAngle C B F'1 \ D E F CongA C B F1" using LeA_def by auto have "A B F'1 CongA G' H' I'" proof - have "A B C D' E' F' SumA A B F'1" proof - have "C B F'1 CongA D' E' F'" using P5 by blast moreover have "\ B C OS A F'1" using P5 by auto moreover have "Coplanar A B C F'1" by (simp add: P5) moreover have "A B F'1 CongA A B F'1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F'1" using P6 inangle_distincts by auto ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed moreover have "A B C D' E' F' SumA G' H' I'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed have "\ Col A B F'1" using \A B F'1 CongA G' H' I'\ \\ Col G' H' I'\ col_conga_col by blast have "\ Col C B F'1" proof - have "\ Col D' E' F'" by (simp add: \\ Col D' E' F'\) moreover have "D' E' F' CongA C B F'1" using P5 not_conga_sym by blast ultimately show ?thesis using ncol_conga_ncol by blast qed show ?thesis proof - have "A B F1 LeA A B F'1" proof - have "F1 InAngle A B F'1" proof - have "F1 InAngle F'1 B A" proof - have "F1 InAngle F'1 B C" by (simp add: P6 l11_24) moreover have "C InAngle F'1 B A" proof - have "B C TS A F'1" using Col_perm P5 \\ Col A B C\ \\ Col C B F'1\ cop_nos__ts ncoplanar_perm_12 by blast obtain X where "Col X B C \ Bet A X F'1" using TS_def \B C TS A F'1\ by auto have "Bet F'1 X A" using Bet_perm \Col X B C \ Bet A X F'1\ by blast moreover have "(X = B) \ (B Out X C)" proof - have "B A OS X C" proof - have "A B OS X F'1" by (metis \Col X B C \ Bet A X F'1\ \\ Col A B C\ \\ Col A B F'1\ bet_out_1 calculation out_one_side) moreover have "A B OS F'1 C" using Col_perm P5 \\ Col A B C\ \\ Col A B F'1\ cop_nos__ts one_side_symmetry by blast ultimately show ?thesis using invert_one_side one_side_transitivity by blast qed thus ?thesis using Col_cases \Col X B C \ Bet A X F'1\ col_one_side_out by blast qed ultimately show ?thesis by (metis InAngle_def \\ Col A B C\ \\ Col A B F'1\ not_col_distincts) qed ultimately show ?thesis using in_angle_trans by blast qed then show ?thesis using l11_24 by blast qed moreover have "A B F1 CongA A B F1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F1" using P6 conga_diff56 by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis by (simp add: inangle__lea) qed moreover have "A B F1 CongA G H I" proof - have "A B C D E F SumA A B F1" proof - have "B C TS F1 A" proof - have "B C TS F'1 A" proof - have "B C TS A F'1" using Col_perm P5 \\ Col A B C\ \\ Col C B F'1\ cop_nos__ts ncoplanar_perm_12 by blast thus ?thesis using l9_2 by blast qed moreover have "B C OS F'1 F1" proof - have "\ Col C B F'1" by (simp add: \\ Col C B F'1\) moreover have "\ Col B C F1" proof - have "\ Col D E F" using \\ Col D E F\ by auto moreover have "D E F CongA C B F1" by (simp add: P6) ultimately show ?thesis using ncol_conga_ncol not_col_permutation_4 by blast qed moreover have "F1 InAngle C B F'1" using P6 by blast ultimately show ?thesis using in_angle_one_side invert_one_side one_side_symmetry by blast qed ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis proof - have "C B F1 CongA D E F" using P6 not_conga_sym by blast moreover have "\ B C OS A F1" using \B C TS F1 A\ l9_9 one_side_symmetry by blast moreover have "Coplanar A B C F1" using \B C TS F1 A\ ncoplanar_perm_9 ts__coplanar by blast moreover have "A B F1 CongA A B F1" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by blast moreover have "B \ F1" using P6 conga_diff56 by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed qed moreover have "A B C D E F SumA G H I" by (simp add: assms(3)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using \A B F'1 CongA G' H' I'\ l11_30 by blast qed qed qed qed qed qed lemma sams_lea123_suma2__lea: assumes "A B C LeA A' B' C'" and "SAMS A' B' C' D E F" and "A B C D E F SumA G H I" and "A' B' C' D E F SumA G' H' I'" shows "G H I LeA G' H' I'" by (meson assms(1) assms(2) assms(3) assms(4) sams_lea456_suma2__lea sams_sym suma_sym) lemma sams_lea2_suma2__lea: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LeA G' H' I'" proof - obtain G'' H'' I'' where "A B C D' E' F' SumA G'' H'' I''" using assms(4) assms(5) ex_suma suma_distincts by presburger have "G H I LeA G'' H'' I''" proof - have "D E F LeA D' E' F'" by (simp add: assms(2)) moreover have "SAMS A B C D' E' F'" proof - have "SAMS A' B' C' D' E' F'" by (simp add: assms(3)) moreover have "A B C LeA A' B' C'" using assms(1) by auto moreover have "D' E' F' LeA D' E' F'" using assms(2) lea_distincts lea_refl by blast ultimately show ?thesis using sams_lea2__sams by blast qed moreover have "A B C D E F SumA G H I" by (simp add: assms(4)) moreover have "A B C D' E' F' SumA G'' H'' I''" by (simp add: \A B C D' E' F' SumA G'' H'' I''\) ultimately show ?thesis using sams_lea456_suma2__lea by blast qed moreover have "G'' H'' I'' LeA G' H' I'" using sams_lea123_suma2__lea assms(3) assms(5) \A B C D' E' F' SumA G'' H'' I''\ assms(1) by blast ultimately show ?thesis using lea_trans by blast qed lemma sams2_suma2__conga456: assumes "SAMS A B C D E F" and "SAMS A B C D' E' F'" and "A B C D E F SumA G H I" and "A B C D' E' F' SumA G H I" shows "D E F CongA D' E' F'" proof cases assume "Col A B C" show ?thesis proof cases assume P2: "Bet A B C" then have "E Out D F" using assms(1) SAMS_def by blast moreover have "E' Out D' F'" using P2 assms(2) SAMS_def by blast ultimately show ?thesis by (simp add: l11_21_b) next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ or_bet_out by blast show ?thesis proof - have "D E F CongA G H I" proof - have "A B C D E F SumA G H I" by (simp add: assms(3)) thus ?thesis using \B Out A C\ out213_suma__conga by auto qed moreover have "G H I CongA D' E' F'" proof - have "A B C D' E' F' SumA G H I" by (simp add: assms(4)) then have "D' E' F' CongA G H I" using \B Out A C\ out213_suma__conga by auto thus ?thesis using not_conga_sym by blast qed ultimately show ?thesis using not_conga by blast qed qed next assume "\ Col A B C" obtain J where T1: "C B J CongA D E F \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J" using assms(1) SAMS_def by blast have T1A: "C B J CongA D E F" using T1 by simp have T1B: "\ B C OS A J" using T1 by simp have T1C: "\ A B TS C J" using T1 by simp have T1D: "Coplanar A B C J" using T1 by simp obtain J' where T2: "C B J' CongA D' E' F' \ \ B C OS A J' \ \ A B TS C J' \ Coplanar A B C J'" using assms(2) SAMS_def by blast have T2A: "C B J' CongA D' E' F'" using T2 by simp have T2B: "\ B C OS A J'" using T2 by simp have T2C: "\ A B TS C J'" using T2 by simp have T2D: "Coplanar A B C J'" using T2 by simp have T3: "C B J CongA C B J'" proof - have "A B J CongA A B J'" proof - have "A B J CongA G H I" proof - have "A B C D E F SumA A B J" using SumA_def T1A T1B T1D \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto thus ?thesis using assms(3) suma2__conga by blast qed moreover have "G H I CongA A B J'" proof - have "A B C D' E' F' SumA G H I" by (simp add: assms(4)) moreover have "A B C D' E' F' SumA A B J'" using SumA_def T2A T2B T2D \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed have "B Out J J' \ A B TS J J'" proof - have "Coplanar A B J J'" using T1D T2D \\ Col A B C\ coplanar_trans_1 ncoplanar_perm_8 not_col_permutation_2 by blast moreover have "A B J CongA A B J'" by (simp add: \A B J CongA A B J'\) ultimately show ?thesis by (simp add: conga_cop__or_out_ts) qed { assume "B Out J J'" then have "C B J CongA C B J'" by (metis Out_cases \\ Col A B C\ bet_out between_trivial not_col_distincts out2__conga) } { assume "A B TS J J'" then have "A B OS J C" by (meson T1C T1D TS_def \\ Col A B C\ cop_nts__os not_col_permutation_2 one_side_symmetry) then have "A B TS C J'" using \A B TS J J'\ l9_8_2 by blast then have "False" by (simp add: T2C) then have "C B J CongA C B J'" by blast } thus ?thesis using \B Out J J' \ C B J CongA C B J'\ \B Out J J' \ A B TS J J'\ by blast qed then have "C B J CongA D' E' F'" using T2A not_conga by blast thus ?thesis using T1A not_conga not_conga_sym by blast qed lemma sams2_suma2__conga123: assumes "SAMS A B C D E F" and "SAMS A' B' C' D E F" and "A B C D E F SumA G H I" and "A' B' C' D E F SumA G H I" shows "A B C CongA A' B' C'" proof - have "SAMS D E F A B C" by (simp add: assms(1) sams_sym) moreover have "SAMS D E F A' B' C'" by (simp add: assms(2) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(3) suma_sym) moreover have "D E F A' B' C' SumA G H I" using assms(4) suma_sym by blast ultimately show ?thesis using sams2_suma2__conga456 by auto qed lemma suma_assoc_1: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "A' B' C' G H I SumA K L M" shows "A B C D' E' F' SumA K L M" proof - obtain A0 where P1: "Bet A B A0 \ Cong A B B A0" using Cong_perm segment_construction by blast obtain D0 where P2: "Bet D E D0 \ Cong D E E D0" using Cong_perm segment_construction by blast show ?thesis proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A B C" then have "E Out D F" using SAMS_def assms(1) by simp show ?thesis proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using assms(4) \E Out D F\ out213_suma__conga by auto ultimately show ?thesis by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms assms(5)) qed next assume "\ Bet A B C" then have "B Out A C" using \Col A B C\ l6_4_2 by auto have "A \ B" using \B Out A C\ out_distinct by auto have "B \ C" using \\ Bet A B C\ between_trivial by auto have "D' \ E'" using assms(4) suma_distincts by blast have "E' \ F'" using assms(4) suma_distincts by auto show ?thesis proof - obtain K0 L0 M0 where P3:"A B C D' E' F' SumA K0 L0 M0" using ex_suma \A \ B\ \B \ C\ \D' \ E'\ \E' \ F'\ by presburger moreover have "A B C CongA A B C" using \A \ B\ \B \ C\ conga_refl by auto moreover have "D' E' F' CongA D' E' F'" using \D' \ E'\ \E' \ F'\ conga_refl by auto moreover have "K0 L0 M0 CongA K L M" proof - have "K0 L0 M0 CongA D' E' F'" using P3 \B Out A C\ conga_sym out213_suma__conga by blast moreover have "D' E' F' CongA K L M" proof - have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) moreover have "D E F G H I SumA K L M" by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.out213_suma__conga Tarski_neutral_dimensionless.sams2_suma2__conga456 Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms \B Out A C\ assms(2) assms(3) assms(5) calculation not_conga_sym) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using conga3_suma__suma by blast qed qed next assume T1: "\ Col A B C" have "\ Col C B A0" by (metis Col_def P1 \\ Col A B C\ cong_diff l6_16_1) show ?thesis proof cases assume "Col D E F" show ?thesis proof cases assume "Bet D E F" have "H Out G I" using SAMS_def assms(2) \Bet D E F\ by blast have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "A B C CongA A B C" by (metis \\ Col A B C\ conga_pseudo_refl conga_right_comm not_col_distincts) moreover have "D E F CongA D' E' F'" using \H Out G I\ assms(4) out546_suma__conga by auto moreover have "A' B' C' CongA K L M" using \H Out G I\ assms(5) out546_suma__conga by auto ultimately show ?thesis using conga3_suma__suma by blast next assume "\ Bet D E F" then have "E Out D F" using not_bet_out by (simp add: \Col D E F\) show ?thesis proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using out546_suma__conga \E Out D F\ assms(4) out213_suma__conga by auto moreover have "K L M CongA K L M" proof - have "K \ L \ L \ M" using assms(5) suma_distincts by blast thus ?thesis using conga_refl by auto qed ultimately show ?thesis using assms(5) conga3_suma__suma by blast qed qed next assume "\ Col D E F" then have "\ Col F E D0" by (metis Col_def P2 cong_diff l6_16_1 not_col_distincts) show ?thesis proof cases assume "Col G H I" show ?thesis proof cases assume "Bet G H I" have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) then have "E Out D F" using SAMS_def \Bet G H I\ by blast then have "Col D E F" using Col_perm out_col by blast then have "False" using \\ Col D E F\ by auto thus ?thesis by simp next assume "\ Bet G H I" then have "H Out G I" using SAMS_def by (simp add: \Col G H I\ l6_4_2) show ?thesis proof - have "A B C CongA A B C" by (metis \\ Col A B C\ conga_refl not_col_distincts) moreover have "D E F CongA D' E' F'" using assms(4) out546_suma__conga \H Out G I\ by auto moreover have "A' B' C' CongA K L M" using \H Out G I\ assms(5) out546_suma__conga by auto ultimately show ?thesis using assms(3) conga3_suma__suma by blast qed qed next assume "\ Col G H I" have "\ B C OS A A0" using P1 col_trivial_1 one_side_chara by blast have "E F TS D D0" by (metis P2 \\ Col D E F\ \\ Col F E D0\ bet__ts bet_col between_trivial not_col_permutation_1) show ?thesis proof cases assume "Col A' B' C'" show ?thesis proof cases assume "Bet A' B' C'" show ?thesis proof cases assume "Col D' E' F'" show ?thesis proof cases assume "Bet D' E' F'" have "A B C CongA G H I" proof - have "A B C CongA D0 E F" proof - have "SAMS A B C D E F" by (simp add: assms(1)) moreover have "SAMS D0 E F D E F" by (metis P2 \\ Col D E F\ \\ Col F E D0\ bet__sams between_symmetry not_col_distincts sams_right_comm) moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "D0 E F D E F SumA A' B' C'" proof - have "D E F D0 E F SumA A' B' C'" proof - have "F E D0 CongA D0 E F" by (metis \\ Col F E D0\ col_trivial_1 col_trivial_2 conga_pseudo_refl) moreover have "\ E F OS D D0" using P2 col_trivial_1 one_side_chara by blast moreover have "Coplanar D E F D0" by (meson P2 bet__coplanar ncoplanar_perm_1) moreover have "D E D0 CongA A' B' C'" using assms(3) P2 \Bet A' B' C'\ \\ Col F E D0\ conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using SumA_def by blast qed thus ?thesis by (simp add: \D E F D0 E F SumA A' B' C'\ suma_sym) qed ultimately show ?thesis using sams2_suma2__conga123 by blast qed moreover have "D0 E F CongA G H I" proof - have "SAMS D E F D0 E F" using P2 \\ Col D E F\ \\ Col F E D0\ bet__sams not_col_distincts sams_right_comm by auto moreover have "D E F D0 E F SumA D' E' F'" proof - have "F E D0 CongA D0 E F" by (metis (no_types) \\ Col F E D0\ col_trivial_1 col_trivial_2 conga_pseudo_refl) moreover have "\ E F OS D D0" using P2 col_trivial_1 one_side_chara by blast moreover have "Coplanar D E F D0" using P2 bet__coplanar ncoplanar_perm_1 by blast moreover have "D E D0 CongA D' E' F'" using assms(3) P2 \Bet D' E' F'\ \\ Col F E D0\ assms(4) conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using SumA_def by blast qed ultimately show ?thesis using assms(2) assms(4) sams2_suma2__conga456 by auto qed ultimately show ?thesis using conga_trans by blast qed then have "G H I CongA A B C" using not_conga_sym by blast have "D' E' F' A B C SumA K L M" proof - have "A' B' C' CongA D' E' F'" by (metis Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms \Bet A' B' C'\ \Bet D' E' F'\ assms(4) assms(5) conga_line) then show ?thesis by (meson Tarski_neutral_dimensionless.conga3_suma__suma Tarski_neutral_dimensionless.suma2__conga Tarski_neutral_dimensionless_axioms \G H I CongA A B C\ assms(5)) qed thus ?thesis by (simp add: suma_sym) next assume "\ Bet D' E' F'" then have "E' Out D' F'" by (simp add: \Col D' E' F'\ l6_4_2) have "D E F LeA D' E' F'" using assms(2) assms(4) sams_suma__lea123789 by blast then have "E Out D F" using \E' Out D' F'\ out_lea__out by blast then have "Col D E F" using Col_perm out_col by blast then have "False" using \\ Col D E F\ by auto thus ?thesis by simp qed next assume "\ Col D' E' F'" have "D E F CongA C B A0" proof - have "SAMS A B C D E F" by (simp add: assms(1)) moreover have "SAMS A B C C B A0" using P1 \\ Col A B C\ \\ Col C B A0\ bet__sams not_col_distincts by auto moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) moreover have "A B C C B A0 SumA A' B' C'" proof - have "A B C C B A0 SumA A B A0" by (metis P1 \\ Col A B C\ \\ Col C B A0\ bet__suma col_trivial_1 col_trivial_2) moreover have "A B C CongA A B C" using \SAMS A B C C B A0\ calculation sams2_suma2__conga123 by auto moreover have "C B A0 CongA C B A0" using \SAMS A B C C B A0\ calculation(1) sams2_suma2__conga456 by auto moreover have "A B A0 CongA A' B' C'" using P1 \Bet A' B' C'\ \\ Col C B A0\ assms(3) conga_line not_col_distincts suma_distincts by auto ultimately show ?thesis using conga3_suma__suma by blast qed ultimately show ?thesis using sams2_suma2__conga456 by blast qed have "SAMS C B A0 G H I" proof - have "D E F CongA C B A0" by (simp add: \D E F CongA C B A0\) moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "SAMS D E F G H I" by (simp add: assms(2)) ultimately show ?thesis using conga2_sams__sams by blast qed then obtain J where P3: "A0 B J CongA G H I \ \ B A0 OS C J \ \ C B TS A0 J \ Coplanar C B A0 J" using SAMS_def by blast obtain F1 where P4: "F E F1 CongA G H I \ \ E F OS D F1 \ \ D E TS F F1 \ Coplanar D E F F1" using SAMS_def assms(2) by auto have "C B J CongA D' E' F'" proof - have "C B J CongA D E F1" proof - have "(B A0 TS C J \ E F TS D F1) \ (B A0 OS C J \ E F OS D F1)" proof - have "B A0 TS C J" proof - have "Coplanar B A0 C J" using P3 ncoplanar_perm_12 by blast moreover have "\ Col C B A0" by (simp add: \\ Col C B A0\) moreover have "\ Col J B A0" using P3 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast moreover have "\ B A0 OS C J" using P3 by simp ultimately show ?thesis by (simp add: cop_nos__ts) qed moreover have "E F TS D F1" proof - have "Coplanar E F D F1" using P4 ncoplanar_perm_12 by blast moreover have "\ Col D E F" by (simp add: \\ Col D E F\) moreover have "\ Col F1 E F" using P4 \\ Col G H I\ col_conga_col col_permutation_3 by blast moreover have "\ E F OS D F1" using P4 by auto ultimately show ?thesis by (simp add: cop_nos__ts) qed ultimately show ?thesis by simp qed moreover have "C B A0 CongA D E F" using \D E F CongA C B A0\ not_conga_sym by blast moreover have "A0 B J CongA F E F1" proof - have "A0 B J CongA G H I" by (simp add: P3) moreover have "G H I CongA F E F1" using P4 not_conga_sym by blast ultimately show ?thesis using conga_trans by blast qed ultimately show ?thesis using l11_22 by auto qed moreover have "D E F1 CongA D' E' F'" proof - have "D E F G H I SumA D E F1" using P4 SumA_def \\ Col D E F\ conga_distinct conga_refl not_col_distincts by auto moreover have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed show ?thesis proof - have "A B C D' E' F' SumA A B J" proof - have "C B TS J A" proof - have "C B TS A0 A" proof - have "B \ A0" using \\ Col C B A0\ not_col_distincts by blast moreover have "\ Col B C A" using Col_cases \\ Col A B C\ by auto moreover have "Bet A B A0" by (simp add: P1) ultimately show ?thesis by (metis Bet_cases Col_cases \\ Col C B A0\ bet__ts invert_two_sides not_col_distincts) qed moreover have "C B OS A0 J" proof - have "\ Col J C B" using \C B J CongA D' E' F'\ \\ Col D' E' F'\ col_conga_col not_col_permutation_2 by blast moreover have "\ Col A0 C B" using Col_cases \\ Col C B A0\ by blast ultimately show ?thesis using P3 cop_nos__ts by blast qed ultimately show ?thesis using l9_8_2 by blast qed moreover have "C B J CongA D' E' F'" by (simp add: \C B J CongA D' E' F'\) moreover have "\ B C OS A J" using calculation(1) invert_one_side l9_9_bis one_side_symmetry by blast moreover have "Coplanar A B C J" using calculation(1) ncoplanar_perm_15 ts__coplanar by blast moreover have "A B J CongA A B J" proof - have "A \ B" using \\ Col A B C\ col_trivial_1 by auto moreover have "B \ J" using \C B TS J A\ ts_distincts by blast ultimately show ?thesis using conga_refl by auto qed ultimately show ?thesis using SumA_def by blast qed moreover have "A B J CongA K L M" proof - have "A' B' C' G H I SumA A B J" proof - have "A B A0 CongA A' B' C'" using P1 \Bet A' B' C'\ \\ Col A B C\ \\ Col C B A0\ assms(5) conga_line not_col_distincts suma_distincts by auto moreover have "A0 B J CongA G H I" by (simp add: P3) moreover have "A B A0 A0 B J SumA A B J" proof - have "A0 B J CongA A0 B J" proof - have "A0 \ B" using \\ Col C B A0\ col_trivial_2 by auto moreover have "B \ J" using CongA_def \A0 B J CongA G H I\ by blast ultimately show ?thesis using conga_refl by auto qed moreover have "\ B A0 OS A J" by (simp add: Col_def P1 col123__nos) moreover have "Coplanar A B A0 J" using P1 bet__coplanar by auto moreover have "A B J CongA A B J" using P1 \\ Col A B C\ between_symmetry calculation(1) l11_13 not_col_distincts by blast ultimately show ?thesis using SumA_def by blast qed ultimately show ?thesis by (meson conga3_suma__suma suma2__conga) qed moreover have "A' B' C' G H I SumA K L M" by (simp add: assms(5)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis proof - have "A B C CongA A B C \ D' E' F' CongA D' E' F'" using CongA_def \A B J CongA K L M\ \C B J CongA D' E' F'\ conga_refl by presburger then show ?thesis using \A B C D' E' F' SumA A B J\ \A B J CongA K L M\ conga3_suma__suma by blast qed qed qed next assume "\ Bet A' B' C'" have "B Out A C" proof - have "A B C LeA A' B' C'" using assms(1) assms(3) sams_suma__lea123789 by auto moreover have "B' Out A' C'" using \Col A' B' C'\ \\ Bet A' B' C'\ or_bet_out by blast ultimately show ?thesis using out_lea__out by blast qed then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto thus ?thesis by simp qed next assume "\ Col A' B' C'" obtain C1 where P6: "C B C1 CongA D E F \ \ B C OS A C1 \ \ A B TS C C1 \ Coplanar A B C C1" using SAMS_def assms(1) by auto have P6A: "C B C1 CongA D E F" using P6 by simp have P6B: "\ B C OS A C1" using P6 by simp have P6C: "\ A B TS C C1" using P6 by simp have P6D: "Coplanar A B C C1" using P6 by simp have "A' B' C' CongA A B C1" proof - have "A B C D E F SumA A B C1" using P6A P6B P6D SumA_def \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto moreover have "A B C D E F SumA A' B' C'" by (simp add: assms(3)) ultimately show ?thesis using suma2__conga by auto qed have "B C1 OS C A" proof - have "B A OS C C1" proof - have "A B OS C C1" proof - have "\ Col C A B" using Col_perm \\ Col A B C\ by blast moreover have "\ Col C1 A B" using \\ Col A' B' C'\ \A' B' C' CongA A B C1\ col_permutation_1 ncol_conga_ncol by blast ultimately show ?thesis using P6C P6D cop_nos__ts by blast qed thus ?thesis by (simp add: invert_one_side) qed moreover have "B C TS A C1" proof - have "Coplanar B C A C1" using P6D ncoplanar_perm_12 by blast moreover have "\ Col C1 B C" proof - have "D E F CongA C1 B C" using P6A conga_left_comm not_conga_sym by blast thus ?thesis using \\ Col D E F\ ncol_conga_ncol by blast qed ultimately show ?thesis using T1 P6B cop_nos__ts by blast qed ultimately show ?thesis using os_ts1324__os one_side_symmetry by blast qed show ?thesis proof cases assume "Col D' E' F'" show ?thesis proof cases assume "Bet D' E' F'" obtain C0 where P7: "Bet C B C0 \ Cong C B B C0" using Cong_perm segment_construction by blast have "B C1 TS C C0" by (metis P7 \B C1 OS C A\ bet__ts cong_diff_2 not_col_distincts one_side_not_col123) show ?thesis proof - have "A B C C B C0 SumA A B C0" proof - have "C B C0 CongA C B C0" by (metis P7 T1 cong_diff conga_line not_col_distincts) moreover have "\ B C OS A C0" using P7 bet_col col124__nos invert_one_side by blast moreover have "Coplanar A B C C0" using P7 bet__coplanar ncoplanar_perm_15 by blast moreover have "A B C0 CongA A B C0" by (metis P7 T1 cong_diff conga_refl not_col_distincts) ultimately show ?thesis using SumA_def by blast qed moreover have "A B C0 CongA K L M" proof - have "A' B' C' G H I SumA A B C0" proof - have "A B C1 C1 B C0 SumA A B C0" using \B C1 TS C C0\ \B C1 OS C A\ l9_8_2 ts__suma_1 by blast moreover have "A B C1 CongA A' B' C'" by (simp add: P6 \A' B' C' CongA A B C1\ conga_sym) moreover have "C1 B C0 CongA G H I" proof - have "SAMS C B C1 C1 B C0" by (metis P7 \B C1 TS C C0\ bet__sams ts_distincts) moreover have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "SAMS D E F G H I" by (simp add: assms(2)) ultimately show ?thesis using conga2_sams__sams by blast qed moreover have "C B C1 C1 B C0 SumA C B C0" by (simp add: \B C1 TS C C0\ ts__suma_1) moreover have "C B C1 G H I SumA C B C0" proof - have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) moreover have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce moreover have "D' E' F' CongA C B C0" using P7 assms(4) by (metis P6A Tarski_neutral_dimensionless.suma_distincts Tarski_neutral_dimensionless_axioms \Bet D' E' F'\ cong_diff conga_diff1 conga_line) ultimately show ?thesis using conga3_suma__suma by blast qed ultimately show ?thesis using sams2_suma2__conga456 by auto qed moreover have "A B C0 CongA A B C0" by (metis P7 T1 cong_diff conga_refl not_col_distincts) ultimately show ?thesis using conga3_suma__suma by blast qed thus ?thesis using assms(5) suma2__conga by auto qed moreover have "A B C CongA A B C" proof - have "A \ B \ B \ C" using T1 col_trivial_1 col_trivial_2 by auto thus ?thesis using conga_refl by auto qed moreover have "C B C0 CongA D' E' F'" proof - have "C \ B" using T1 col_trivial_2 by blast moreover have "B \ C0" using \B C1 TS C C0\ ts_distincts by blast moreover have "D' \ E'" using assms(4) suma_distincts by blast moreover have "E' \ F'" using assms(4) suma_distincts by auto ultimately show ?thesis by (simp add: P7 \Bet D' E' F'\ conga_line) qed ultimately show ?thesis using conga3_suma__suma by blast qed next assume "\ Bet D' E' F'" then have "E' Out D' F'" by (simp add: \Col D' E' F'\ l6_4_2) have "D E F LeA D' E' F'" using sams_suma__lea123789 assms(2) assms(4) by auto then have "E Out D F" using \E' Out D' F'\ out_lea__out by blast then have "False" using Col_cases \\ Col D E F\ out_col by blast thus ?thesis by simp qed next assume "\ Col D' E' F'" have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P6A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce ultimately show ?thesis using assms(2) conga2_sams__sams by blast qed then obtain J where P7: "C1 B J CongA G H I \ \ B C1 OS C J \ \ C B TS C1 J \ Coplanar C B C1 J" using SAMS_def by blast have P7A: "C1 B J CongA G H I" using P7 by simp have P7B: "\ B C1 OS C J" using P7 by simp have P7C: "\ C B TS C1 J" using P7 by simp have P7D: "Coplanar C B C1 J" using P7 by simp obtain F1 where P8: "F E F1 CongA G H I \ \ E F OS D F1 \ \ D E TS F F1 \ Coplanar D E F F1" using SAMS_def assms(2) by auto have P8A: "F E F1 CongA G H I" using P8 by simp have P8B: "\ E F OS D F1" using P8 by simp have P8C: "\ D E TS F F1" using P8 by simp have P8D: "Coplanar D E F F1" using P8 by simp have "C B J CongA D' E' F'" proof - have "C B J CongA D E F1" proof - have "B C1 TS C J" proof - have "Coplanar B C1 C J" using P7D ncoplanar_perm_12 by blast moreover have "\ Col C B C1" using \B C1 OS C A\ not_col_permutation_2 one_side_not_col123 by blast moreover have "\ Col J B C1" using P7 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast moreover have "\ B C1 OS C J" by (simp add: P7B) ultimately show ?thesis by (simp add: cop_nos__ts) qed moreover have "E F TS D F1" proof - have "Coplanar E F D F1" using P8D ncoplanar_perm_12 by blast moreover have "\ Col F1 E F" using P8 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast ultimately show ?thesis using P8B \\ Col D E F\ cop_nos__ts by blast qed moreover have "C B C1 CongA D E F" using P6A by blast moreover have "C1 B J CongA F E F1" using P8 by (meson P7A not_conga not_conga_sym) ultimately show ?thesis using l11_22a by blast qed moreover have "D E F1 CongA D' E' F'" proof - have "D E F G H I SumA D E F1" using P8A P8B P8D SumA_def \\ Col D E F\ conga_distinct conga_refl not_col_distincts by auto moreover have "D E F G H I SumA D' E' F'" by (simp add: assms(4)) ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using conga_trans by blast qed have "\ Col C B C1" using \B C1 OS C A\ col123__nos col_permutation_1 by blast show ?thesis proof - have "A B C C B J SumA A B J" proof - have "B C TS J A" proof - have "B C TS C1 A" using cop_nos__ts using P6B P6D T1 \\ Col C B C1\ l9_2 ncoplanar_perm_12 not_col_permutation_3 by blast moreover have "B C OS C1 J" proof - have "\ Col C1 C B" using Col_perm \\ Col C B C1\ by blast moreover have "\ Col J C B" using \C B J CongA D' E' F'\ \\ Col D' E' F'\ col_conga_col col_permutation_1 by blast ultimately show ?thesis using P7C P7D cop_nos__ts invert_one_side by blast qed ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis by (simp add: l9_2 ts__suma_1) qed moreover have "A B C CongA A B C" using T1 conga_refl not_col_distincts by fastforce moreover have "A B J CongA K L M" proof - have "A' B' C' G H I SumA A B J" proof - have "A B C1 C1 B J SumA A B J" proof - have "B C1 TS A J" proof - have "B C1 TS C J" proof - have "Coplanar B C1 C J" using P7D ncoplanar_perm_12 by blast moreover have "\ Col J B C1" using P7 \\ Col G H I\ col_conga_col not_col_permutation_3 by blast ultimately show ?thesis by (simp add: \\ Col C B C1\ P7B cop_nos__ts) qed moreover have "B C1 OS C A" by (simp add: \B C1 OS C A\) ultimately show ?thesis using l9_8_2 by blast qed thus ?thesis by (simp add: ts__suma_1) qed moreover have "A B C1 CongA A' B' C'" using \A' B' C' CongA A B C1\ not_conga_sym by blast moreover have "C1 B J CongA G H I" by (simp add: P7A) moreover have "A B J CongA A B J" using \A B C C B J SumA A B J\ suma2__conga by auto ultimately show ?thesis using conga3_suma__suma by blast qed moreover have "A' B' C' G H I SumA K L M" using assms(5) by simp ultimately show ?thesis using suma2__conga by auto qed ultimately show ?thesis using \C B J CongA D' E' F'\ conga3_suma__suma by blast qed qed qed qed qed qed qed lemma suma_assoc_2: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "A B C D' E' F' SumA K L M" shows "A' B' C' G H I SumA K L M" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) sams_sym suma_assoc_1 suma_sym) lemma suma_assoc: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" shows "A' B' C' G H I SumA K L M \ A B C D' E' F' SumA K L M" by (meson assms(1) assms(2) assms(3) assms(4) suma_assoc_1 suma_assoc_2) lemma sams_assoc_1: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "SAMS A' B' C' G H I" shows "SAMS A B C D' E' F'" proof cases assume "Col A B C" { assume "E Out D F" have "SAMS A B C D' E' F'" proof - have "A' B' C' CongA A B C" using assms(3) \E Out D F\ conga_sym out546_suma__conga by blast moreover have "G H I CongA D' E' F'" using \E Out D F\ assms(4) out213_suma__conga by blast ultimately show ?thesis using assms(5) conga2_sams__sams by blast qed } { assume "\ Bet A B C" then have P1: "B Out A C" using \Col A B C\ or_bet_out by blast have "SAMS D' E' F' A B C" proof - have "D' \ E'" using assms(4) suma_distincts by auto moreover have "F' E' F' CongA A B C" proof - have "E' \ F'" using assms(4) suma_distincts by blast then have "E' Out F' F'" using out_trivial by auto thus ?thesis using P1 l11_21_b by blast qed moreover have "\ E' F' OS D' F'" using os_distincts by blast moreover have "\ D' E' TS F' F'" by (simp add: not_two_sides_id) moreover have "Coplanar D' E' F' F'" using ncop_distincts by blast ultimately show ?thesis using SAMS_def P1 by blast qed then have "SAMS A B C D' E' F'" using sams_sym by blast } thus ?thesis using SAMS_def assms(1) \E Out D F \ SAMS A B C D' E' F'\ by blast next assume "\ Col A B C" show ?thesis proof cases assume "Col D E F" have "H Out G I \ \ Bet D E F" using SAMS_def assms(2) by blast { assume "H Out G I" have "SAMS A B C D' E' F'" proof - have "A B C CongA A B C" using \\ Col A B C\ conga_refl not_col_distincts by fastforce moreover have "D E F CongA D' E' F'" using \H Out G I\ assms(4) out546_suma__conga by blast ultimately show ?thesis using assms(1) conga2_sams__sams by blast qed } { assume "\ Bet D E F" then have "E Out D F" using \Col D E F\ l6_4_2 by blast have "SAMS A B C D' E' F'" proof - have "A' B' C' CongA A B C" using out546_suma__conga \E Out D F\ assms(3) not_conga_sym by blast moreover have "G H I CongA D' E' F'" using out213_suma__conga \E Out D F\ assms(4) by auto ultimately show ?thesis using assms(5) conga2_sams__sams by auto qed } thus ?thesis using \H Out G I \ SAMS A B C D' E' F'\ \H Out G I \ \ Bet D E F\ by blast next assume "\ Col D E F" show ?thesis proof cases assume "Col G H I" have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) then have "E Out D F \ \ Bet G H I" using SAMS_def by blast { assume "E Out D F" then have "False" using Col_cases \\ Col D E F\ out_col by blast then have "SAMS A B C D' E' F'" by simp } { assume "\ Bet G H I" then have "H Out G I" by (simp add: \Col G H I\ l6_4_2) have "SAMS A B C D' E' F'" proof - have "A B C CongA A B C" by (metis \\ Col A B C\ col_trivial_1 col_trivial_2 conga_refl) moreover have "D E F CongA D' E' F'" using out546_suma__conga \H Out G I\ assms(4) by blast moreover have "SAMS A B C D E F" using assms(1) by auto ultimately show ?thesis using conga2_sams__sams by auto qed } thus ?thesis using \E Out D F \ \ Bet G H I\ \E Out D F \ SAMS A B C D' E' F'\ by blast next assume "\ Col G H I" show ?thesis proof - have "\ Bet A B C" using Col_def \\ Col A B C\ by auto moreover have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "\ Col A' B' C'" proof - { assume "Col A' B' C'" have "H Out G I \ \ Bet A' B' C'" using SAMS_def assms(5) by blast { assume "H Out G I" then have "False" using Col_cases \\ Col G H I\ out_col by blast } { assume "\ Bet A' B' C'" then have "B' Out A' C'" using not_bet_out \Col A' B' C'\ by blast have "A B C LeA A' B' C'" using assms(1) assms(3) sams_suma__lea123789 by auto then have "B Out A C" using \B' Out A' C'\ out_lea__out by blast then have "Col A B C" using Col_perm out_col by blast then have "False" using \\ Col A B C\ by auto } then have "False" using \H Out G I \ False\ \H Out G I \ \ Bet A' B' C'\ by blast } thus ?thesis by blast qed obtain C1 where P1: "C B C1 CongA D E F \ \ B C OS A C1 \ \ A B TS C C1 \ Coplanar A B C C1" using SAMS_def assms(1) by auto have P1A: "C B C1 CongA D E F" using P1 by simp have P1B: "\ B C OS A C1" using P1 by simp have P1C: "\ A B TS C C1" using P1 by simp have P1D: "Coplanar A B C C1" using P1 by simp have "A B C1 CongA A' B' C'" proof - have "A B C D E F SumA A B C1" using P1A P1B P1D SumA_def \\ Col A B C\ conga_distinct conga_refl not_col_distincts by auto thus ?thesis using assms(3) suma2__conga by auto qed have "SAMS C B C1 G H I" proof - have "D E F CongA C B C1" using P1A not_conga_sym by blast moreover have "G H I CongA G H I" using \\ Col G H I\ conga_refl not_col_distincts by fastforce ultimately show ?thesis using conga2_sams__sams using assms(2) by blast qed then obtain J where T1: "C1 B J CongA G H I \ \ B C1 OS C J \ \ C B TS C1 J \ Coplanar C B C1 J" using SAMS_def by auto have T1A: "C1 B J CongA G H I" using T1 by simp have T1B: "\ B C1 OS C J" using T1 by simp have T1C: "\ C B TS C1 J" using T1 by simp have T1D: "Coplanar C B C1 J" using T1 by simp have "SAMS A B C1 C1 B J" proof - have "A' B' C' CongA A B C1" using \A B C1 CongA A' B' C'\ not_conga_sym by blast moreover have "G H I CongA C1 B J" using T1A not_conga_sym by blast ultimately show ?thesis using assms(5) conga2_sams__sams by auto qed then obtain J' where T2: "C1 B J' CongA C1 B J \ \ B C1 OS A J' \ \ A B TS C1 J' \ Coplanar A B C1 J'" using SAMS_def by auto have T2A: "C1 B J' CongA C1 B J" using T2 by simp have T2B: "\ B C1 OS A J'" using T2 by simp have T2C: "\ A B TS C1 J'" using T2 by simp have T2D: "Coplanar A B C1 J'" using T2 by simp have "A' B' C' CongA A B C1" using \A B C1 CongA A' B' C'\ not_conga_sym by blast then have "\ Col A B C1" using ncol_conga_ncol \\ Col A' B' C'\ by blast have "D E F CongA C B C1" using P1A not_conga_sym by blast then have "\ Col C B C1" using ncol_conga_ncol \\ Col D E F\ by blast then have "Coplanar C1 B A J" using coplanar_trans_1 P1D T1D coplanar_perm_15 coplanar_perm_6 by blast have "Coplanar C1 B J' J" using T2D \Coplanar C1 B A J\ \\ Col A B C1\ coplanar_perm_14 coplanar_perm_6 coplanar_trans_1 by blast have "B Out J' J \ C1 B TS J' J" by (meson T2 \Coplanar C1 B A J\ \\ Col A B C1\ conga_cop__or_out_ts coplanar_trans_1 ncoplanar_perm_14 ncoplanar_perm_6) { assume "B Out J' J" have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" proof - have "C B C1 C1 B J SumA C B J" proof - have "C1 B J CongA C1 B J" using CongA_def T2A conga_refl by auto moreover have "C B J CongA C B J" using \\ Col C B C1\ calculation conga_diff56 conga_pseudo_refl conga_right_comm not_col_distincts by blast ultimately show ?thesis using T1D T1B SumA_def by blast qed then have "D E F G H I SumA C B J" using conga3_suma__suma by (meson P1A T1A suma2__conga) then have "C B J CongA D' E' F'" using assms(4) suma2__conga by auto moreover have "\ B C OS A J" by (metis (no_types, lifting) Col_perm P1B P1D T1C \\ Col A B C\ \\ Col C B C1\ cop_nos__ts coplanar_perm_8 invert_two_sides l9_2 l9_8_2) moreover have "\ A B TS C J" proof cases assume "Col A B J" thus ?thesis using TS_def invert_two_sides not_col_permutation_3 by blast next assume "\ Col A B J" have "A B OS C J" proof - have "A B OS C C1" by (simp add: P1C P1D \\ Col A B C1\ \\ Col A B C\ cop_nts__os not_col_permutation_2) moreover have "A B OS C1 J" proof - have "A B OS C1 J'" by (metis T2C T2D \B Out J' J\ \\ Col A B C1\ \\ Col A B J\ col_trivial_2 colx cop_nts__os not_col_permutation_2 out_col out_distinct) thus ?thesis using \B Out J' J\ invert_one_side out_out_one_side by blast qed ultimately show ?thesis using one_side_transitivity by blast qed thus ?thesis using l9_9 by blast qed moreover have "Coplanar A B C J" by (meson P1D \Coplanar C1 B A J\ \\ Col A B C1\ coplanar_perm_18 coplanar_perm_2 coplanar_trans_1 not_col_permutation_2) ultimately show ?thesis by blast qed } { assume "C1 B TS J' J" have "B C1 OS C J" proof - have "B C1 TS C J'" proof - have "B C1 TS A J'" by (meson T2B T2D TS_def \C1 B TS J' J\ \\ Col A B C1\ cop_nts__os invert_two_sides ncoplanar_perm_12) moreover have "B C1 OS A C" by (meson P1B P1C P1D \\ Col A B C1\ \\ Col A B C\ \\ Col C B C1\ cop_nts__os invert_one_side ncoplanar_perm_12 not_col_permutation_2 not_col_permutation_3 os_ts1324__os) ultimately show ?thesis using l9_8_2 by blast qed moreover have "B C1 TS J J'" using \C1 B TS J' J\ invert_two_sides l9_2 by blast ultimately show ?thesis using OS_def by blast qed then have "False" by (simp add: T1B) then have "\ J. (C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J)" by auto } thus ?thesis using \B Out J' J \ \J. C B J CongA D' E' F' \ \ B C OS A J \ \ A B TS C J \ Coplanar A B C J\ \B Out J' J \ C1 B TS J' J\ by blast qed ultimately show ?thesis using SAMS_def not_bet_distincts by auto qed qed qed qed lemma sams_assoc_2: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" and "SAMS A B C D' E' F'" shows "SAMS A' B' C' G H I" proof - have "SAMS G H I A' B' C'" proof - have "SAMS G H I D E F" by (simp add: assms(2) sams_sym) moreover have "SAMS D E F A B C" by (simp add: assms(1) sams_sym) moreover have "G H I D E F SumA D' E' F'" by (simp add: assms(4) suma_sym) moreover have "D E F A B C SumA A' B' C'" by (simp add: assms(3) suma_sym) moreover have "SAMS D' E' F' A B C" by (simp add: assms(5) sams_sym) ultimately show ?thesis using sams_assoc_1 by blast qed thus ?thesis using sams_sym by blast qed lemma sams_assoc: assumes "SAMS A B C D E F" and "SAMS D E F G H I" and "A B C D E F SumA A' B' C'" and "D E F G H I SumA D' E' F'" shows "(SAMS A' B' C' G H I) \ (SAMS A B C D' E' F')" using sams_assoc_1 sams_assoc_2 by (meson assms(1) assms(2) assms(3) assms(4)) lemma sams_nos__nts: assumes "SAMS A B C C B J" and "\ B C OS A J" shows "\ A B TS C J" proof - obtain J' where P1: "C B J' CongA C B J \ \ B C OS A J' \ \ A B TS C J' \ Coplanar A B C J'" using SAMS_def assms(1) by blast have P1A: "C B J' CongA C B J" using P1 by simp have P1B: "\ B C OS A J'" using P1 by simp have P1C: "\ A B TS C J'" using P1 by simp have P1D: "Coplanar A B C J'" using P1 by simp have P2: "B Out C J \ \ Bet A B C" using SAMS_def assms(1) by blast { assume "A B TS C J" have "Coplanar C B J J'" proof - have "\ Col A C B" using TS_def \A B TS C J\ not_col_permutation_4 by blast moreover have "Coplanar A C B J" using \A B TS C J\ ncoplanar_perm_2 ts__coplanar by blast moreover have "Coplanar A C B J'" using P1D ncoplanar_perm_2 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed have "B Out J J' \ C B TS J J'" by (metis P1 \A B TS C J\ \Coplanar C B J J'\ bet_conga__bet bet_out col_conga_col col_two_sides_bet conga_distinct conga_os__out conga_sym cop_nts__os invert_two_sides l5_2 l6_6 not_col_permutation_3 not_col_permutation_4) { assume "B Out J J'" have "\ Col B A J \ \ Col B A J'" using TS_def \A B TS C J\ not_col_permutation_3 by blast then have "A B OS C J'" by (metis (full_types) \B Out J J'\ Col_cases P1C P1D TS_def \A B TS C J\ col2__eq cop_nts__os l6_3_1 out_col) then have "A B TS C J'" by (meson \A B TS C J\ \B Out J J'\ l6_6 l9_2 not_col_distincts out_two_sides_two_sides) then have "False" by (simp add: P1C) } { assume "C B TS J J'" then have "\ Col C A B \ \ Col J A B" using TS_def \A B TS C J\ by blast then have "False" by (metis P1B P1D TS_def \C B TS J J'\ assms(2) cop_nts__os invert_two_sides l9_8_1 ncoplanar_perm_12 not_col_permutation_1) } then have "False" using \B Out J J' \ False\ \B Out J J' \ C B TS J J'\ by blast } thus ?thesis by auto qed lemma conga_sams_nos__nts: assumes "SAMS A B C D E F" and "C B J CongA D E F" and "\ B C OS A J" shows "\ A B TS C J" proof - have "SAMS A B C C B J" proof - have "A B C CongA A B C" using assms(1) conga_refl sams_distincts by fastforce moreover have "D E F CongA C B J" using assms(2) not_conga_sym by blast ultimately show ?thesis using assms(1) conga2_sams__sams by auto qed thus ?thesis by (simp add: assms(3) sams_nos__nts) qed lemma sams_lea2_suma2__conga123: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G H I" shows "A B C CongA A' B' C'" proof - have "SAMS A B C D E F" using assms(1) assms(2) assms(3) sams_lea2__sams by blast moreover have "SAMS A' B' C' D E F" by (metis assms(2) assms(3) lea_refl sams_distincts sams_lea2__sams) moreover have "A' B' C' D E F SumA G H I" proof - obtain G' H' I' where P1: "A' B' C' D E F SumA G' H' I'" using calculation(2) ex_suma sams_distincts by blast show ?thesis proof - have "A' \ B' \ B' \ C'" using assms(1) lea_distincts by blast then have "A' B' C' CongA A' B' C'" using conga_refl by auto moreover have "D \ E \ E \ F" using \SAMS A B C D E F\ sams_distincts by blast then have "D E F CongA D E F" using conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' H' I' LeA G H I" using P1 assms(2) assms(3) assms(5) sams_lea456_suma2__lea by blast moreover have "G H I LeA G' H' I'" proof - have "SAMS A' B' C' D E F" using \SAMS A' B' C' D E F\ by auto thus ?thesis using P1 assms(1) assms(4) sams_lea123_suma2__lea by blast qed ultimately show ?thesis by (simp add: lea_asym) qed ultimately show ?thesis using P1 conga3_suma__suma by blast qed qed ultimately show ?thesis using assms(4) sams2_suma2__conga123 by blast qed lemma sams_lea2_suma2__conga456: assumes "A B C LeA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G H I" shows "D E F CongA D' E' F'" proof - have "SAMS D' E' F' A' B' C'" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G H I" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__conga123 by auto qed lemma sams_suma__out213: assumes "A B C D E F SumA D E F" and "SAMS A B C D E F" shows "B Out A C" proof - have "E \ D" using assms(2) sams_distincts by blast then have "E Out D D" using out_trivial by auto moreover have "D E D CongA A B C" proof - have "D E D LeA A B C" using assms(1) lea121345 suma_distincts by auto moreover have "E \ D \ E \ F" using assms(2) sams_distincts by blast then have "D E F LeA D E F" using lea_refl by auto moreover have "D E D D E F SumA D E F" proof - have "\ E D OS D F" using os_distincts by auto moreover have "Coplanar D E D F" using ncop_distincts by auto ultimately show ?thesis using SumA_def \D E F LeA D E F\ lea_asym by blast qed ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__conga123 by auto qed ultimately show ?thesis using eq_conga_out by blast qed lemma sams_suma__out546: assumes "A B C D E F SumA A B C" and "SAMS A B C D E F" shows "E Out D F" proof - have "D E F A B C SumA A B C" using assms(1) suma_sym by blast moreover have "SAMS D E F A B C" using assms(2) sams_sym by blast ultimately show ?thesis using sams_suma__out213 by blast qed lemma sams_lea_lta123_suma2__lta: assumes "A B C LtA A' B' C'" and "D E F LeA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" proof - have "G H I LeA G' H' I'" proof - have "A B C LeA A' B' C'" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea2_suma2__lea by blast qed moreover have "\ G H I CongA G' H' I'" proof - { assume "G H I CongA G' H' I'" have "A B C CongA A' B' C'" proof - have "A B C LeA A' B' C'" by (simp add: assms(1) lta__lea) moreover have "A' B' C' D' E' F' SumA G H I" by (meson \G H I CongA G' H' I'\ assms(3) assms(5) conga3_suma__suma conga_sym sams2_suma2__conga123 sams2_suma2__conga456) ultimately show ?thesis using assms(2) assms(3) assms(4) sams_lea2_suma2__conga123 by blast qed then have "False" using assms(1) lta_not_conga by auto } thus ?thesis by auto qed ultimately show ?thesis using LtA_def by blast qed lemma sams_lea_lta456_suma2__lta: assumes "A B C LeA A' B' C'" and "D E F LtA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" using sams_lea_lta123_suma2__lta by (meson assms(1) assms(2) assms(3) assms(4) assms(5) sams_sym suma_sym) lemma sams_lta2_suma2__lta: assumes "A B C LtA A' B' C'" and "D E F LtA D' E' F'" and "SAMS A' B' C' D' E' F'" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "G H I LtA G' H' I'" using sams_lea_lta123_suma2__lta by (meson LtA_def assms(1) assms(2) assms(3) assms(4) assms(5)) lemma sams_lea2_suma2__lea123: assumes "D' E' F' LeA D E F" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LeA A' B' C'" proof cases assume "A' B' C' LtA A B C" then have "G' H' I' LtA G H I" using sams_lea_lta123_suma2__lta using assms(1) assms(3) assms(4) assms(5) by blast then have "\ G H I LeA G' H' I'" using lea__nlta by blast then have "False" using assms(2) by auto thus ?thesis by auto next assume "\ A' B' C' LtA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis by (simp add: \\ A' B' C' LtA A B C\ nlta__lea) qed lemma sams_lea2_suma2__lea456: assumes "A' B' C' LeA A B C" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LeA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea2_suma2__lea123 by blast qed lemma sams_lea_lta456_suma2__lta123: assumes "D' E' F' LtA D E F" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof cases assume "A' B' C' LeA A B C" then have "G' H' I' LtA G H I" using sams_lea_lta456_suma2__lta assms(1) assms(3) assms(4) assms(5) by blast then have "\ G H I LeA G' H' I'" using lea__nlta by blast then have "False" using assms(2) by blast thus ?thesis by blast next assume "\ A' B' C' LeA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis using nlea__lta by (simp add: \\ A' B' C' LeA A B C\) qed lemma sams_lea_lta123_suma2__lta456: assumes "A' B' C' LtA A B C" and "G H I LeA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" by (simp add: assms(5) suma_sym) ultimately show ?thesis using assms(1) assms(2) sams_lea_lta456_suma2__lta123 by blast qed lemma sams_lea_lta789_suma2__lta123: assumes "D' E' F' LeA D E F" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof cases assume "A' B' C' LeA A B C" then have "G' H' I' LeA G H I" using assms(1) assms(3) assms(4) assms(5) sams_lea2_suma2__lea by blast then have "\ G H I LtA G' H' I'" by (simp add: lea__nlta) then have "False" using assms(2) by blast thus ?thesis by auto next assume "\ A' B' C' LeA A B C" have "A' \ B' \ B' \ C' \ A \ B \ B \ C" using assms(4) assms(5) suma_distincts by auto thus ?thesis using nlea__lta by (simp add: \\ A' B' C' LeA A B C\) qed lemma sams_lea_lta789_suma2__lta456: assumes "A' B' C' LeA A B C" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "SAMS D E F A B C" by (simp add: assms(3) sams_sym) moreover have "D E F A B C SumA G H I" by (simp add: assms(4) suma_sym) moreover have "D' E' F' A' B' C' SumA G' H' I'" using assms(5) suma_sym by blast ultimately show ?thesis using assms(1) assms(2) sams_lea_lta789_suma2__lta123 by blast qed lemma sams_lta2_suma2__lta123: assumes "D' E' F' LtA D E F" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "A B C LtA A' B' C'" proof - have "D' E' F' LeA D E F" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea_lta789_suma2__lta123 by blast qed lemma sams_lta2_suma2__lta456: assumes "A' B' C' LtA A B C" and "G H I LtA G' H' I'" and "SAMS A B C D E F" and "A B C D E F SumA G H I" and "A' B' C' D' E' F' SumA G' H' I'" shows "D E F LtA D' E' F'" proof - have "A' B' C' LeA A B C" by (simp add: assms(1) lta__lea) thus ?thesis using assms(2) assms(3) assms(4) assms(5) sams_lea_lta789_suma2__lta456 by blast qed lemma sams123231: assumes "A \ B" and "A \ C" and "B \ C" shows "SAMS A B C B C A" proof - obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto then have "A' \ B" using assms(1) midpoint_not_midpoint by blast moreover have "Bet A B A'" by (simp add: \B Midpoint A A'\ midpoint_bet) moreover have "B C A LeA C B A'" proof cases assume "Col A B C" show ?thesis proof cases assume "Bet A C B" thus ?thesis by (metis assms(2) assms(3) between_exchange3 calculation(1) calculation(2) l11_31_2) next assume "\ Bet A C B" then have "C Out B A" using Col_cases \Col A B C\ l6_6 or_bet_out by blast thus ?thesis using assms(3) calculation(1) l11_31_1 by auto qed next assume "\ Col A B C" thus ?thesis using l11_41_aux \B Midpoint A A'\ calculation(1) lta__lea midpoint_bet not_col_permutation_4 by blast qed ultimately show ?thesis using assms(1) sams_chara by blast qed lemma col_suma__col: assumes "Col D E F" and "A B C B C A SumA D E F" shows "Col A B C" proof - { assume "\ Col A B C" have "False" proof cases assume "Bet D E F" obtain P where P1: "Bet A B P \ Cong A B B P" using Cong_perm segment_construction by blast have "D E F LtA A B P" proof - have "A B C LeA A B C" using \\ Col A B C\ lea_total not_col_distincts by blast moreover have "B C TS A P" by (metis Cong_perm P1 \\ Col A B C\ bet__ts bet_col between_trivial2 cong_reverse_identity not_col_permutation_1) then have "B C A LtA C B P" using Col_perm P1 \\ Col A B C\ calculation l11_41_aux ts_distincts by blast moreover have "A B C C B P SumA A B P" by (simp add: \B C TS A P\ ts__suma_1) ultimately show ?thesis by (meson P1 Tarski_neutral_dimensionless.sams_lea_lta456_suma2__lta Tarski_neutral_dimensionless_axioms assms(2) bet_suma__sams) qed thus ?thesis using P1 \Bet D E F\ bet2_lta__lta lta_distincts by blast next assume "\ Bet D E F" have "C Out B A" proof - have "E Out D F" by (simp add: \\ Bet D E F\ assms(1) l6_4_2) moreover have "B C A LeA D E F" using sams_suma__lea456789 by (metis assms(2) sams123231 suma_distincts) ultimately show ?thesis using out_lea__out by blast qed thus ?thesis using Col_cases \\ Col A B C\ out_col by blast qed } thus ?thesis by auto qed lemma ncol_suma__ncol: assumes "\ Col A B C" and "A B C B C A SumA D E F" shows "\ Col D E F" using col_suma__col assms(1) assms(2) by blast lemma per2_suma__bet: assumes "Per A B C" and "Per D E F" and "A B C D E F SumA G H I" shows "Bet G H I" proof - obtain A1 where P1: "C B A1 CongA D E F \ \ B C OS A A1 \ Coplanar A B C A1 \ A B A1 CongA G H I" using SumA_def assms(3) by blast then have "D E F CongA A1 B C" using conga_right_comm conga_sym by blast then have "Per A1 B C" using assms(2) l11_17 by blast have "Bet A B A1" proof - have "Col B A A1" proof - have "Coplanar C A A1 B" using P1 ncoplanar_perm_10 by blast moreover have "C \ B" using \D E F CongA A1 B C\ conga_distinct by auto ultimately show ?thesis using assms(1) \Per A1 B C\ col_permutation_2 cop_per2__col by blast qed moreover have "B C TS A A1" proof - have "Coplanar B C A A1" using calculation ncop__ncols by blast moreover have "A \ B \ B \ C" using CongA_def P1 by blast then have "\ Col A B C" by (simp add: assms(1) per_not_col) moreover have "A1 \ B \ B \ C" using \D E F CongA A1 B C\ conga_distinct by blast then have "\ Col A1 B C" using \Per A1 B C\ by (simp add: per_not_col) ultimately show ?thesis by (simp add: P1 cop_nos__ts) qed ultimately show ?thesis using col_two_sides_bet by blast qed thus ?thesis using P1 bet_conga__bet by blast qed lemma bet_per2__suma: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "G \ H" and "H \ I" and "Per A B C" and "Per D E F" and "Bet G H I" shows "A B C D E F SumA G H I" proof - obtain G' H' I' where "A B C D E F SumA G' H' I'" using assms(1) assms(2) assms(3) assms(4) ex_suma by blast moreover have "A B C CongA A B C" using assms(1) assms(2) conga_refl by auto moreover have "D E F CongA D E F" using assms(3) assms(4) conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' \ H'" using calculation(1) suma_distincts by auto moreover have "H' \ I'" using \A B C D E F SumA G' H' I'\ suma_distincts by blast moreover have "Bet G' H' I'" using \A B C D E F SumA G' H' I'\ assms(7) assms(8) per2_suma__bet by auto ultimately show ?thesis using conga_line by (simp add: assms(5) assms(6) assms(9)) qed ultimately show ?thesis using conga3_suma__suma by blast qed lemma per2__sams: assumes "A \ B" and "B \ C" and "D \ E" and "E \ F" and "Per A B C" and "Per D E F" shows "SAMS A B C D E F" proof - obtain G H I where "A B C D E F SumA G H I" using assms(1) assms(2) assms(3) assms(4) ex_suma by blast moreover then have "Bet G H I" using assms(5) assms(6) per2_suma__bet by auto ultimately show ?thesis using bet_suma__sams by blast qed lemma bet_per_suma__per456: assumes "Per A B C" and "Bet G H I" and "A B C D E F SumA G H I" shows "Per D E F" proof - obtain A1 where "B Midpoint A A1" using symmetric_point_construction by auto have "\ Col A B C" using assms(1) assms(3) per_col_eq suma_distincts by blast have "A B C CongA D E F" proof - have "SAMS A B C A B C" using \\ Col A B C\ assms(1) not_col_distincts per2__sams by auto moreover have "SAMS A B C D E F" using assms(2) assms(3) bet_suma__sams by blast moreover have "A B C A B C SumA G H I" using assms(1) assms(2) assms(3) bet_per2__suma suma_distincts by blast ultimately show ?thesis using assms(3) sams2_suma2__conga456 by auto qed thus ?thesis using assms(1) l11_17 by blast qed lemma bet_per_suma__per123: assumes "Per D E F" and "Bet G H I" and "A B C D E F SumA G H I" shows "Per A B C" using bet_per_suma__per456 by (meson assms(1) assms(2) assms(3) suma_sym) lemma bet_suma__per: assumes "Bet D E F" and "A B C A B C SumA D E F" shows "Per A B C" proof - obtain A' where "C B A' CongA A B C \ A B A' CongA D E F" using SumA_def assms(2) by blast have "Per C B A" proof - have "Bet A B A'" proof - have "D E F CongA A B A'" using \C B A' CongA A B C \ A B A' CongA D E F\ not_conga_sym by blast thus ?thesis using assms(1) bet_conga__bet by blast qed moreover have "C B A CongA C B A'" using conga_left_comm not_conga_sym \C B A' CongA A B C \ A B A' CongA D E F\ by blast ultimately show ?thesis using l11_18_2 by auto qed thus ?thesis using Per_cases by auto qed lemma acute__sams: assumes "Acute A B C" shows "SAMS A B C A B C" proof - obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto show ?thesis proof - have "A \ B \ A' \ B" using \B Midpoint A A'\ acute_distincts assms is_midpoint_id_2 by blast moreover have "Bet A B A'" by (simp add: \B Midpoint A A'\ midpoint_bet) moreover have "Obtuse C B A'" using acute_bet__obtuse assms calculation(1) calculation(2) obtuse_sym by blast then have "A B C LeA C B A'" by (metis acute__not_obtuse assms calculation(1) lea_obtuse_obtuse lea_total obtuse_distincts) ultimately show ?thesis using sams_chara by blast qed qed lemma acute_suma__nbet: assumes "Acute A B C" and "A B C A B C SumA D E F" shows "\ Bet D E F" proof - { assume "Bet D E F" then have "Per A B C" using assms(2) bet_suma__per by auto then have "A B C LtA A B C" using acute_not_per assms(1) by auto then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma acute2__sams: assumes "Acute A B C" and "Acute D E F" shows "SAMS A B C D E F" by (metis acute__sams acute_distincts assms(1) assms(2) lea_total sams_lea2__sams) lemma acute2_suma__nbet_a: assumes "Acute A B C" and "D E F LeA A B C" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - { assume "Bet G H I" obtain A' B' C' where "A B C A B C SumA A' B' C'" using acute_distincts assms(1) ex_suma by moura have "G H I LeA A' B' C'" proof - have "A B C LeA A B C" using acute_distincts assms(1) lea_refl by blast moreover have "SAMS A B C A B C" by (simp add: acute__sams assms(1)) ultimately show ?thesis using \A B C A B C SumA A' B' C'\ assms(1) assms(2) assms(3) sams_lea456_suma2__lea by blast qed then have "Bet A' B' C'" using \Bet G H I\ bet_lea__bet by blast then have "False" using acute_suma__nbet assms(1) assms(3) \A B C A B C SumA A' B' C'\ by blast } thus ?thesis by blast qed lemma acute2_suma__nbet: assumes "Acute A B C" and "Acute D E F" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - have "A \ B \ B \ C \ D \ E \ E \ F" using assms(3) suma_distincts by auto then have "A B C LeA D E F \ D E F LeA A B C" by (simp add: lea_total) moreover { assume P3: "A B C LeA D E F" have "D E F A B C SumA G H I" by (simp add: assms(3) suma_sym) then have "\ Bet G H I" using P3 assms(2) acute2_suma__nbet_a by auto } { assume "D E F LeA A B C" then have "\ Bet G H I" using acute2_suma__nbet_a assms(1) assms(3) by auto } thus ?thesis using \A B C LeA D E F \ \ Bet G H I\ calculation by blast qed lemma acute_per__sams: assumes "A \ B" and "B \ C" and "Per A B C" and "Acute D E F" shows "SAMS A B C D E F" proof - have "SAMS A B C A B C" by (simp add: assms(1) assms(2) assms(3) per2__sams) moreover have "A B C LeA A B C" using assms(1) assms(2) lea_refl by auto moreover have "D E F LeA A B C" by (metis acute_distincts acute_lea_acute acute_not_per assms(1) assms(2) assms(3) assms(4) lea_total) ultimately show ?thesis using sams_lea2__sams by blast qed lemma acute_per_suma__nbet: assumes "A \ B" and "B \ C" and "Per A B C" and "Acute D E F" and "A B C D E F SumA G H I" shows "\ Bet G H I" proof - { assume "Bet G H I" have "G H I LtA G H I" proof - have "A B C LeA A B C" using assms(1) assms(2) lea_refl by auto moreover have "D E F LtA A B C" by (simp add: acute_per__lta assms(1) assms(2) assms(3) assms(4)) moreover have "SAMS A B C A B C" by (simp add: assms(1) assms(2) assms(3) per2__sams) moreover have "A B C D E F SumA G H I" by (simp add: assms(5)) moreover have "A B C A B C SumA G H I" by (meson Tarski_neutral_dimensionless.bet_per_suma__per456 Tarski_neutral_dimensionless_axioms \Bet G H I\ acute_not_per assms(3) assms(4) calculation(4)) ultimately show ?thesis using sams_lea_lta456_suma2__lta by blast qed then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma obtuse__nsams: assumes "Obtuse A B C" shows "\ SAMS A B C A B C" proof - { assume "SAMS A B C A B C" obtain A' where "B Midpoint A A'" using symmetric_point_construction by auto have "A B C LtA A B C" proof - have "A B C LeA A' B C" by (metis \B Midpoint A A'\ \SAMS A B C A B C\ lea_right_comm midpoint_bet midpoint_distinct_2 sams_chara sams_distincts) moreover have "A' B C LtA A B C" using \B Midpoint A A'\ assms calculation lea_distincts midpoint_bet obtuse_chara by blast ultimately show ?thesis using lea__nlta by blast qed then have "False" by (simp add: nlta) } thus ?thesis by blast qed lemma nbet_sams_suma__acute: assumes "\ Bet D E F" and "SAMS A B C A B C" and "A B C A B C SumA D E F" shows "Acute A B C" proof - have "Acute A B C \ Per A B C \ Obtuse A B C" by (metis angle_partition l8_20_1_R1 l8_5) { assume "Per A B C" then have "Bet D E F" using assms(3) per2_suma__bet by auto then have "False" using assms(1) by auto } { assume "Obtuse A B C" then have "\ SAMS A B C A B C" by (simp add: obtuse__nsams) then have "False" using assms(2) by auto } thus ?thesis using \Acute A B C \ Per A B C \ Obtuse A B C\ \Per A B C \ False\ by auto qed lemma nsams__obtuse: assumes "A \ B" and "B \ C" and "\ SAMS A B C A B C" shows "Obtuse A B C" proof - { assume "Per A B C" obtain A' where "B Midpoint A A'" using symmetric_point_construction by blast have "\ Col A B C" using \Per A B C\ assms(1) assms(2) per_col_eq by blast have "SAMS A B C A B C" proof - have "C B A' CongA A B C" using \Per A B C\ assms(1) assms(2) assms(3) per2__sams by blast moreover have "\ B C OS A A'" by (meson Col_cases \B Midpoint A A'\ col_one_side_out l6_4_1 midpoint_bet midpoint_col) moreover have "\ A B TS C A'" using Col_def Midpoint_def TS_def \B Midpoint A A'\ by blast moreover have "Coplanar A B C A'" using \Per A B C\ assms(1) assms(2) assms(3) per2__sams by blast ultimately show ?thesis using SAMS_def \\ Col A B C\ assms(1) bet_col by auto qed then have "False" using assms(3) by auto } { assume "Acute A B C" then have "SAMS A B C A B C" by (simp add: acute__sams) then have "False" using assms(3) by auto } thus ?thesis using \Per A B C \ False\ angle_partition assms(1) assms(2) by auto qed lemma sams2_suma2__conga: assumes "SAMS A B C A B C" and "A B C A B C SumA D E F" and "SAMS A' B' C' A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - have "A B C LeA A' B' C' \ A' B' C' LeA A B C" using assms(1) assms(3) lea_total sams_distincts by auto moreover have "A B C LeA A' B' C' \ A B C CongA A' B' C'" using assms(2) assms(3) assms(4) sams_lea2_suma2__conga123 by auto ultimately show ?thesis by (meson Tarski_neutral_dimensionless.conga_sym Tarski_neutral_dimensionless.sams_lea2_suma2__conga123 Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4)) qed lemma acute2_suma2__conga: assumes "Acute A B C" and "A B C A B C SumA D E F" and "Acute A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - have "SAMS A B C A B C" by (simp add: acute__sams assms(1)) moreover have "SAMS A' B' C' A' B' C'" by (simp add: acute__sams assms(3)) ultimately show ?thesis using assms(2) assms(4) sams2_suma2__conga by auto qed lemma bet2_suma__out: assumes "Bet A B C" and "Bet D E F" and "A B C D E F SumA G H I" shows "H Out G I" proof - have "A B C D E F SumA A B A" proof - have "C B A CongA D E F" by (metis Bet_cases assms(1) assms(2) assms(3) conga_line suma_distincts) moreover have "\ B C OS A A" by (simp add: Col_def assms(1) col124__nos) moreover have "Coplanar A B C A" using ncop_distincts by blast moreover have "A B A CongA A B A" using calculation(1) conga_diff2 conga_trivial_1 by auto ultimately show ?thesis using SumA_def by blast qed then have "A B A CongA G H I" using assms(3) suma2__conga by auto thus ?thesis using eq_conga_out by auto qed lemma col2_suma__col: assumes "Col A B C" and "Col D E F" and "A B C D E F SumA G H I" shows "Col G H I" proof cases assume "Bet A B C" show ?thesis proof cases assume "Bet D E F" thus ?thesis using bet2_suma__out by (meson \Bet A B C\ assms(3) not_col_permutation_4 out_col) next assume "\ Bet D E F" show ?thesis proof - have "E Out D F" using \\ Bet D E F\ assms(2) or_bet_out by auto then have "A B C CongA G H I" using assms(3) out546_suma__conga by auto thus ?thesis using assms(1) col_conga_col by blast qed qed next assume "\ Bet A B C" have "D E F CongA G H I" proof - have "B Out A C" by (simp add: \\ Bet A B C\ assms(1) l6_4_2) thus ?thesis using assms(3) out213_suma__conga by auto qed thus ?thesis using assms(2) col_conga_col by blast qed lemma suma_suppa__bet: assumes "A B C SuppA D E F" and "A B C D E F SumA G H I" shows "Bet G H I" proof - obtain A' where P1: "Bet A B A' \ D E F CongA C B A'" using SuppA_def assms(1) by auto obtain A'' where P2: "C B A'' CongA D E F \ \ B C OS A A'' \ Coplanar A B C A'' \ A B A'' CongA G H I" using SumA_def assms(2) by auto have "B Out A' A'' \ C B TS A' A''" proof - have "Coplanar C B A' A''" proof - have "Coplanar C A'' B A" using P2 coplanar_perm_17 by blast moreover have "B \ A" using SuppA_def assms(1) by auto moreover have "Col B A A'" using P1 by (simp add: bet_col col_permutation_4) ultimately show ?thesis using col_cop__cop coplanar_perm_3 by blast qed moreover have "C B A' CongA C B A''" proof - have "C B A' CongA D E F" using P1 not_conga_sym by blast moreover have "D E F CongA C B A''" using P2 not_conga_sym by blast ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using conga_cop__or_out_ts by simp qed have "Bet A B A''" proof - have "\ C B TS A' A''" proof - { assume "C B TS A' A''" have "B C TS A A'" proof - { assume "Col A B C" then have "Col A' C B" using P1 assms(1) bet_col bet_col1 col3 suppa_distincts by blast then have "False" using TS_def \C B TS A' A''\ by auto } then have "\ Col A B C" by auto moreover have "\ Col A' B C" using TS_def \C B TS A' A''\ not_col_permutation_5 by blast moreover have "\ T. (Col T B C \ Bet A T A')" using P1 not_col_distincts by blast ultimately show ?thesis by (simp add: TS_def) qed then have "B C OS A A''" using OS_def \C B TS A' A''\ invert_two_sides l9_2 by blast then have "False" using P2 by simp } thus ?thesis by blast qed then have "B Out A' A''" using \B Out A' A'' \ C B TS A' A''\ by auto moreover have "A' \ B \ A'' \ B \ A \ B" using P2 calculation conga_diff1 out_diff1 out_diff2 by blast moreover have "Bet A' B A" using P1 Bet_perm by blast ultimately show ?thesis using bet_out__bet between_symmetry by blast qed moreover have "A B A'' CongA G H I" using P2 by blast ultimately show ?thesis using bet_conga__bet by blast qed lemma bet_suppa__suma: assumes "G \ H" and "H \ I" and "A B C SuppA D E F" and "Bet G H I" shows "A B C D E F SumA G H I" proof - obtain G' H' I' where "A B C D E F SumA G' H' I'" using assms(3) ex_suma suppa_distincts by blast moreover have "A B C CongA A B C" using calculation conga_refl suma_distincts by fastforce moreover have "D \ E \ E \ F" using assms(3) suppa_distincts by auto then have "D E F CongA D E F" using conga_refl by auto moreover have "G' H' I' CongA G H I" proof - have "G' \ H' \ H' \ I'" using calculation(1) suma_distincts by auto moreover have "Bet G' H' I'" using \A B C D E F SumA G' H' I'\ assms(3) suma_suppa__bet by blast ultimately show ?thesis using assms(1) assms(2) assms(4) conga_line by auto qed ultimately show ?thesis using conga3_suma__suma by blast qed lemma bet_suma__suppa: assumes "A B C D E F SumA G H I" and "Bet G H I" shows "A B C SuppA D E F" proof - obtain A' where "C B A' CongA D E F \ A B A' CongA G H I" using SumA_def assms(1) by blast moreover have "G H I CongA A B A'" using calculation not_conga_sym by blast then have "Bet A B A'" using assms(2) bet_conga__bet by blast moreover have "D E F CongA C B A'" using calculation(1) not_conga_sym by blast ultimately show ?thesis by (metis SuppA_def conga_diff1) qed lemma bet2_suma__suma: assumes "A' \ B" and "D' \ E" and "Bet A B A'" and "Bet D E D'" and "A B C D E F SumA G H I" shows "A' B C D' E F SumA G H I" proof - obtain J where P1: "C B J CongA D E F \ \ B C OS A J \ Coplanar A B C J \ A B J CongA G H I" using SumA_def assms(5) by auto moreover obtain C' where P2: "Bet C B C' \ Cong B C' B C" using segment_construction by blast moreover have "A B C' D' E F SumA G H I" proof - have "C' B J CongA D' E F" by (metis assms(2) assms(4) calculation(1) calculation(2) cong_diff_3 conga_diff1 l11_13) moreover have "\ B C' OS A J" by (metis Col_cases P1 P2 bet_col col_one_side cong_diff) moreover have "Coplanar A B C' J" by (smt P1 P2 bet_col bet_col1 col2_cop__cop cong_diff ncoplanar_perm_5) ultimately show ?thesis using P1 SumA_def by blast qed moreover have "A B C' CongA A' B C" using assms(1) assms(3) assms(5) between_symmetry calculation(2) calculation(3) l11_14 suma_distincts by auto moreover have "D' \ E \ E \ F" using assms(2) calculation(1) conga_distinct by blast then have "D' E F CongA D' E F" using conga_refl by auto moreover have "G \ H \ H \ I" using assms(5) suma_distincts by blast then have "G H I CongA G H I" using conga_refl by auto ultimately show ?thesis using conga3_suma__suma by blast qed lemma suma_suppa2__suma: assumes "A B C SuppA A' B' C'" and "D E F SuppA D' E' F'" and "A B C D E F SumA G H I" shows "A' B' C' D' E' F' SumA G H I" proof - obtain A0 where P1: "Bet A B A0 \ A' B' C' CongA C B A0" using SuppA_def assms(1) by auto obtain D0 where P2: "Bet D E D0 \ D' E' F' CongA F E D0" using SuppA_def assms(2) by auto show ?thesis proof - have "A0 B C D0 E F SumA G H I" proof - have "A0 \ B" using CongA_def P1 by auto moreover have "D0 \ E" using CongA_def P2 by blast ultimately show ?thesis using P1 P2 assms(3) bet2_suma__suma by auto qed moreover have "A0 B C CongA A' B' C'" using P1 conga_left_comm not_conga_sym by blast moreover have "D0 E F CongA D' E' F'" using P2 conga_left_comm not_conga_sym by blast moreover have "G \ H \ H \ I" using assms(3) suma_distincts by blast then have "G H I CongA G H I" using conga_refl by auto ultimately show ?thesis using conga3_suma__suma by blast qed qed lemma suma2_obtuse2__conga: assumes "Obtuse A B C" and "A B C A B C SumA D E F" and "Obtuse A' B' C'" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast moreover obtain A0' where P2: "Bet A' B' A0' \ Cong B' A0' A' B'" using segment_construction by blast moreover have "A0 B C CongA A0' B' C'" proof - have "Acute A0 B C" using assms(1) bet_obtuse__acute P1 cong_diff_3 obtuse_distincts by blast moreover have "A0 B C A0 B C SumA D E F" using P1 acute_distincts assms(2) bet2_suma__suma calculation by blast moreover have "Acute A0' B' C'" using P2 assms(3) bet_obtuse__acute cong_diff_3 obtuse_distincts by blast moreover have "A0' B' C' A0' B' C' SumA D E F" by (metis P2 assms(4) bet2_suma__suma cong_diff_3) ultimately show ?thesis using acute2_suma2__conga by blast qed moreover have "Bet A0 B A" using Bet_perm calculation(1) by blast moreover have "Bet A0' B' A'" using Bet_perm calculation(2) by blast moreover have "A \ B" using assms(1) obtuse_distincts by blast moreover have "A' \ B'" using assms(3) obtuse_distincts by blast ultimately show ?thesis using l11_13 by blast qed lemma bet_suma2__or_conga: assumes "A0 \ B" and "Bet A B A0" and "A B C A B C SumA D E F" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" proof - { fix A' B' C' assume"Acute A' B' C' \ A' B' C' A' B' C' SumA D E F" have "Per A B C \ Acute A B C \ Obtuse A B C" by (metis angle_partition l8_20_1_R1 l8_5) { assume "Per A B C" then have "Bet D E F" using per2_suma__bet assms(3) by auto then have "False" using \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\ acute_suma__nbet by blast then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } { assume "Acute A B C" have "Acute A' B' C'" by (simp add: \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\) moreover have "A' B' C' A' B' C' SumA D E F" by (simp add: \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\) ultimately have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using assms(3) \Acute A B C\ acute2_suma2__conga by auto } { assume "Obtuse A B C" have "Acute A0 B C" using \Obtuse A B C\ assms(1) assms(2) bet_obtuse__acute by auto moreover have "A0 B C A0 B C SumA D E F" using assms(1) assms(2) assms(3) bet2_suma__suma by auto ultimately have "A0 B C CongA A' B' C'" using \Acute A' B' C' \ A' B' C' A' B' C' SumA D E F\ acute2_suma2__conga by auto then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using \Acute A B C \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A B C \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A B C \ Acute A B C \ Obtuse A B C\ by blast } then have P1: "\ A' B' C'. (Acute A' B' C' \ A' B' C' A' B' C' SumA D E F) \ (A B C CongA A' B' C' \ A0 B C CongA A' B' C')" by blast have "Per A' B' C' \ Acute A' B' C' \ Obtuse A' B' C'" by (metis angle_partition l8_20_1_R1 l8_5) { assume P2: "Per A' B' C'" have "A B C CongA A' B' C'" proof - have "A \ B \ B \ C" using assms(3) suma_distincts by blast moreover have "A' \ B' \ B' \ C'" using assms(4) suma_distincts by auto moreover have "Per A B C" proof - have "Bet D E F" using P2 assms(4) per2_suma__bet by blast moreover have "A B C A B C SumA D E F" by (simp add: assms(3)) ultimately show ?thesis using assms(3) bet_suma__per by blast qed ultimately show ?thesis using P2 l11_16 by blast qed then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" by blast } { assume "Acute A' B' C'" then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using P1 assms(4) by blast } { assume "Obtuse A' B' C'" obtain A0' where "Bet A' B' A0' \ Cong B' A0' A' B'" using segment_construction by blast moreover have "Acute A0' B' C'" using \Obtuse A' B' C'\ bet_obtuse__acute calculation cong_diff_3 obtuse_distincts by blast moreover have "A0' B' C' A0' B' C' SumA D E F" using acute_distincts assms(4) bet2_suma__suma calculation(1) calculation(2) by blast ultimately have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using P1 by (metis assms(1) assms(2) assms(3) assms(4) between_symmetry l11_13 suma_distincts) } thus ?thesis using \Acute A' B' C' \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A' B' C' \ A B C CongA A' B' C' \ A0 B C CongA A' B' C'\ \Per A' B' C' \ Acute A' B' C' \ Obtuse A' B' C'\ by blast qed lemma suma2__or_conga_suppa: assumes "A B C A B C SumA D E F" and "A' B' C' A' B' C' SumA D E F" shows "A B C CongA A' B' C' \ A B C SuppA A' B' C'" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast then have "A0 \ B" using assms(1) bet_cong_eq suma_distincts by blast then have "A B C CongA A' B' C' \ A0 B C CongA A' B' C'" using assms(1) assms(2) P1 bet_suma2__or_conga by blast thus ?thesis by (metis P1 SuppA_def cong_diff conga_right_comm conga_sym) qed lemma ex_trisuma: assumes "A \ B" and "B \ C" and "A \ C" shows "\ D E F. A B C TriSumA D E F" proof - obtain G H I where "A B C B C A SumA G H I" using assms(1) assms(2) assms(3) ex_suma by presburger moreover then obtain D E F where "G H I C A B SumA D E F" using ex_suma suma_distincts by presburger ultimately show ?thesis using TriSumA_def by blast qed lemma trisuma_perm_231: assumes "A B C TriSumA D E F" shows "B C A TriSumA D E F" proof - obtain A1 B1 C1 where P1: "A B C B C A SumA A1 B1 C1 \ A1 B1 C1 C A B SumA D E F" using TriSumA_def assms(1) by auto obtain A2 B2 C2 where P2: "B C A C A B SumA B2 C2 A2" using P1 ex_suma suma_distincts by fastforce have "A B C B2 C2 A2 SumA D E F" proof - have "SAMS A B C B C A" using assms sams123231 trisuma_distincts by auto moreover have "SAMS B C A C A B" using P2 sams123231 suma_distincts by fastforce ultimately show ?thesis using P1 P2 suma_assoc by blast qed thus ?thesis using P2 TriSumA_def suma_sym by blast qed lemma trisuma_perm_312: assumes "A B C TriSumA D E F" shows "C A B TriSumA D E F" by (simp add: assms trisuma_perm_231) lemma trisuma_perm_321: assumes "A B C TriSumA D E F" shows "C B A TriSumA D E F" proof - obtain G H I where "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto thus ?thesis by (meson TriSumA_def suma_comm suma_right_comm suma_sym trisuma_perm_231) qed lemma trisuma_perm_213: assumes "A B C TriSumA D E F" shows "B A C TriSumA D E F" using assms trisuma_perm_231 trisuma_perm_321 by blast lemma trisuma_perm_132: assumes "A B C TriSumA D E F" shows "A C B TriSumA D E F" using assms trisuma_perm_312 trisuma_perm_321 by blast lemma conga_trisuma__trisuma: assumes "A B C TriSumA D E F" and "D E F CongA D' E' F'" shows "A B C TriSumA D' E' F'" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto have "G H I C A B SumA D' E' F'" proof - have f1: "B \ A" by (metis P1 suma_distincts) have f2: "C \ A" using P1 suma_distincts by blast have "G H I CongA G H I" by (metis (full_types) P1 conga_refl suma_distincts) then show ?thesis using f2 f1 by (meson P1 assms(2) conga3_suma__suma conga_refl) qed thus ?thesis using P1 TriSumA_def by blast qed lemma trisuma2__conga: assumes "A B C TriSumA D E F" and "A B C TriSumA D' E' F'" shows "D E F CongA D' E' F'" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto then have P1A: "G H I C A B SumA D E F" by simp obtain G' H' I' where P2: "A B C B C A SumA G' H' I' \ G' H' I' C A B SumA D' E' F'" using TriSumA_def assms(2) by auto have "G' H' I' C A B SumA D E F" proof - have "G H I CongA G' H' I'" using P1 P2 suma2__conga by blast moreover have "D E F CongA D E F \ C A B CongA C A B" by (metis assms(1) conga_refl trisuma_distincts) ultimately show ?thesis by (meson P1 conga3_suma__suma) qed thus ?thesis using P2 suma2__conga by auto qed lemma conga3_trisuma__trisuma: assumes "A B C TriSumA D E F" and "A B C CongA A' B' C'" and "B C A CongA B' C' A'" and "C A B CongA C' A' B'" shows "A' B' C' TriSumA D E F" proof - obtain G H I where P1: "A B C B C A SumA G H I \ G H I C A B SumA D E F" using TriSumA_def assms(1) by auto thus ?thesis proof - have "A' B' C' B' C' A' SumA G H I" using conga3_suma__suma P1 by (meson assms(2) assms(3) suma2__conga) moreover have "G H I C' A' B' SumA D E F" using conga3_suma__suma P1 by (meson P1 assms(4) suma2__conga) ultimately show ?thesis using TriSumA_def by blast qed qed lemma col_trisuma__bet: assumes "Col A B C" and "A B C TriSumA P Q R" shows "Bet P Q R" proof - obtain D E F where P1: "A B C B C A SumA D E F \ D E F C A B SumA P Q R" using TriSumA_def assms(2) by auto { assume "Bet A B C" have "A B C CongA P Q R" proof - have "A B C CongA D E F" proof - have "C \ A \ C \ B" using assms(2) trisuma_distincts by blast then have "C Out B A" using \ Bet A B C\ bet_out_1 by fastforce thus ?thesis using P1 out546_suma__conga by auto qed moreover have "D E F CongA P Q R" proof - have "A \ C \ A \ B" using assms(2) trisuma_distincts by blast then have "A Out C B" using Out_def \Bet A B C\ by auto thus ?thesis using P1 out546_suma__conga by auto qed ultimately show ?thesis using conga_trans by blast qed then have "Bet P Q R" using \Bet A B C\ bet_conga__bet by blast } { assume "Bet B C A" have "B C A CongA P Q R" proof - have "B C A CongA D E F" proof - have "B \ A \ B \ C" using assms(2) trisuma_distincts by blast then have "B Out A C" using Out_def \Bet B C A\ by auto thus ?thesis using P1 out213_suma__conga by blast qed moreover have "D E F CongA P Q R" proof - have "A \ C \ A \ B" using assms(2) trisuma_distincts by auto then have "A Out C B" using \Bet B C A\ bet_out_1 by auto thus ?thesis using P1 out546_suma__conga by blast qed ultimately show ?thesis using not_conga by blast qed then have "Bet P Q R" using \Bet B C A\ bet_conga__bet by blast } { assume "Bet C A B" have "E Out D F" proof - have "C Out B A" using \Bet C A B\ assms(2) bet_out l6_6 trisuma_distincts by blast moreover have "B C A CongA D E F" proof - have "B \ A \ B \ C" using P1 suma_distincts by blast then have "B Out A C" using \Bet C A B\ bet_out_1 by auto thus ?thesis using out213_suma__conga P1 by blast qed ultimately show ?thesis using l11_21_a by blast qed then have "C A B CongA P Q R" using P1 out213_suma__conga by blast then have "Bet P Q R" using \Bet C A B\ bet_conga__bet by blast } thus ?thesis using Col_def \Bet A B C \ Bet P Q R\ \Bet B C A \ Bet P Q R\ assms(1) by blast qed lemma suma_dec: "A B C D E F SumA G H I \ \ A B C D E F SumA G H I" by simp lemma sams_dec: "SAMS A B C D E F \ \ SAMS A B C D E F" by simp lemma trisuma_dec: "A B C TriSumA P Q R \ \ A B C TriSumA P Q R" by simp subsection "Parallelism" lemma par_reflexivity: assumes "A \ B" shows "A B Par A B" using Par_def assms not_col_distincts by blast lemma par_strict_irreflexivity: "\ A B ParStrict A B" using ParStrict_def col_trivial_3 by blast lemma not_par_strict_id: "\ A B ParStrict A C" using ParStrict_def col_trivial_1 by blast lemma par_id: assumes "A B Par A C" shows "Col A B C" using Col_cases Par_def assms not_par_strict_id by auto lemma par_strict_not_col_1: assumes "A B ParStrict C D" shows "\ Col A B C" using Col_perm ParStrict_def assms col_trivial_1 by blast lemma par_strict_not_col_2: assumes "A B ParStrict C D" shows "\ Col B C D" using ParStrict_def assms col_trivial_3 by auto lemma par_strict_not_col_3: assumes "A B ParStrict C D" shows "\ Col C D A" using Col_perm ParStrict_def assms col_trivial_1 by blast lemma par_strict_not_col_4: assumes "A B ParStrict C D" shows "\ Col A B D" using Col_perm ParStrict_def assms col_trivial_3 by blast lemma par_id_1: assumes "A B Par A C" shows "Col B A C" using Par_def assms not_par_strict_id by auto lemma par_id_2: assumes "A B Par A C" shows "Col B C A" using Col_perm assms par_id_1 by blast lemma par_id_3: assumes "A B Par A C" shows "Col A C B" using Col_perm assms par_id_2 by blast lemma par_id_4: assumes "A B Par A C" shows "Col C B A" using Col_perm assms par_id_2 by blast lemma par_id_5: assumes "A B Par A C" shows "Col C A B" using assms col_permutation_2 par_id by blast lemma par_strict_symmetry: assumes "A B ParStrict C D" shows "C D ParStrict A B" using ParStrict_def assms coplanar_perm_16 by blast lemma par_symmetry: assumes "A B Par C D" shows "C D Par A B" by (smt NCol_perm Par_def assms l6_16_1 par_strict_symmetry) lemma par_left_comm: assumes "A B Par C D" shows "B A Par C D" by (metis (mono_tags, lifting) ParStrict_def Par_def assms ncoplanar_perm_6 not_col_permutation_5) lemma par_right_comm: assumes "A B Par C D" shows "A B Par D C" using assms par_left_comm par_symmetry by blast lemma par_comm: assumes "A B Par C D" shows "B A Par D C" using assms par_left_comm par_right_comm by blast lemma par_strict_left_comm: assumes "A B ParStrict C D" shows "B A ParStrict C D" using ParStrict_def assms ncoplanar_perm_6 not_col_permutation_5 by blast lemma par_strict_right_comm: assumes "A B ParStrict C D" shows "A B ParStrict D C" using assms par_strict_left_comm par_strict_symmetry by blast lemma par_strict_comm: assumes "A B ParStrict C D" shows "B A ParStrict D C" by (simp add: assms par_strict_left_comm par_strict_right_comm) lemma par_strict_neq1: assumes "A B ParStrict C D" shows "A \ B" using assms col_trivial_1 par_strict_not_col_4 by blast lemma par_strict_neq2: assumes "A B ParStrict C D" shows "C \ D" using assms col_trivial_2 par_strict_not_col_2 by blast lemma par_neq1: assumes "A B Par C D" shows "A \ B" using Par_def assms par_strict_neq1 by blast lemma par_neq2: assumes "A B Par C D" shows "C \ D" using assms par_neq1 par_symmetry by blast lemma Par_cases: assumes "A B Par C D \ B A Par C D \ A B Par D C \ B A Par D C \ C D Par A B \ C D Par B A \ D C Par A B \ D C Par B A" shows "A B Par C D" using assms par_right_comm par_symmetry by blast lemma Par_perm: assumes "A B Par C D" shows "A B Par C D \ B A Par C D \ A B Par D C \ B A Par D C \ C D Par A B \ C D Par B A \ D C Par A B \ D C Par B A" using Par_cases assms by blast lemma Par_strict_cases: assumes "A B ParStrict C D \ B A ParStrict C D \ A B ParStrict D C \ B A ParStrict D C \ C D ParStrict A B \ C D ParStrict B A \ D C ParStrict A B \ D C ParStrict B A" shows "A B ParStrict C D" using assms par_strict_right_comm par_strict_symmetry by blast lemma Par_strict_perm: assumes "A B ParStrict C D" shows "A B ParStrict C D \ B A ParStrict C D \ A B ParStrict D C \ B A ParStrict D C \ C D ParStrict A B \ C D ParStrict B A \ D C ParStrict A B \ D C ParStrict B A" using Par_strict_cases assms by blast lemma l12_6: assumes "A B ParStrict C D" shows "A B OS C D" by (metis Col_def ParStrict_def Par_strict_perm TS_def assms cop_nts__os par_strict_not_col_2) lemma pars__os3412: assumes "A B ParStrict C D" shows "C D OS A B" by (simp add: assms l12_6 par_strict_symmetry) lemma perp_dec: "A B Perp C D \ \ A B Perp C D" by simp lemma col_cop2_perp2__col: assumes "X1 X2 Perp A B" and "Y1 Y2 Perp A B" and "Col X1 Y1 Y2" and "Coplanar A B X2 Y1" and "Coplanar A B X2 Y2" shows "Col X2 Y1 Y2" proof cases assume "X1 = Y2" thus ?thesis using assms(1) assms(2) assms(4) cop_perp2__col not_col_permutation_2 perp_left_comm by blast next assume "X1 \ Y2" then have "Y2 X1 Perp A B" by (metis Col_cases assms(2) assms(3) perp_col perp_comm perp_right_comm) then have P1: "X1 Y2 Perp A B" using Perp_perm by blast thus ?thesis proof cases assume "X1 = Y1" thus ?thesis using assms(1) assms(2) assms(5) cop_perp2__col not_col_permutation_4 by blast next assume "X1 \ Y1" then have "X1 Y1 Perp A B" using Col_cases P1 assms(3) perp_col by blast thus ?thesis using P1 assms(1) assms(4) assms(5) col_transitivity_2 cop_perp2__col perp_not_eq_1 by blast qed qed lemma col_perp2_ncol_col: assumes "X1 X2 Perp A B" and "Y1 Y2 Perp A B" and "Col X1 Y1 Y2" and "\ Col X1 A B" shows "Col X2 Y1 Y2" proof - have "Coplanar A B X2 Y1" proof cases assume "X1 = Y1" thus ?thesis using assms(1) ncoplanar_perm_22 perp__coplanar by blast next assume "X1 \ Y1" then have "Y1 X1 Perp A B" by (metis Col_cases assms(2) assms(3) perp_col) thus ?thesis by (meson assms(1) assms(4) coplanar_trans_1 ncoplanar_perm_18 ncoplanar_perm_4 perp__coplanar) qed then moreover have "Coplanar A B X2 Y2" by (smt assms(1) assms(2) assms(3) assms(4) col_cop2__cop coplanar_perm_17 coplanar_perm_18 coplanar_trans_1 perp__coplanar) ultimately show ?thesis using assms(1) assms(2) assms(3) col_cop2_perp2__col by blast qed lemma l12_9: assumes "Coplanar C1 C2 A1 B1" and "Coplanar C1 C2 A1 B2" and "Coplanar C1 C2 A2 B1" and "Coplanar C1 C2 A2 B2" and "A1 A2 Perp C1 C2" and "B1 B2 Perp C1 C2" shows "A1 A2 Par B1 B2" proof - have P1: "A1 \ A2 \ C1 \ C2" using assms(5) perp_distinct by auto have P2: "B1 \ B2" using assms(6) perp_distinct by auto show ?thesis proof cases assume "Col A1 B1 B2" then show ?thesis using P1 P2 Par_def assms(3) assms(4) assms(5) assms(6) col_cop2_perp2__col by blast next assume P3: "\ Col A1 B1 B2" { assume "\ Col C1 C2 A1" then have "Coplanar A1 A2 B1 B2" by (smt assms(1) assms(2) assms(5) coplanar_perm_22 coplanar_perm_8 coplanar_pseudo_trans ncop_distincts perp__coplanar) } have "C1 C2 Perp A1 A2" using Perp_cases assms(5) by blast then have "Coplanar A1 A2 B1 B2" by (smt \\ Col C1 C2 A1 \ Coplanar A1 A2 B1 B2\ assms(3) assms(4) coplanar_perm_1 coplanar_pseudo_trans ncop_distincts perp__coplanar perp_not_col2) { assume "\ X. Col X A1 A2 \ Col X B1 B2" then obtain AB where P4: "Col AB A1 A2 \ Col AB B1 B2" by auto then have "False" proof cases assume "AB = A1" thus ?thesis using P3 P4 by blast next assume "AB \ A1" then have "A1 AB Perp C1 C2" by (metis P4 assms(5) not_col_permutation_2 perp_col) then have "AB A1 Perp C1 C2" by (simp add: perp_left_comm) thus ?thesis using P3 P4 assms(1) assms(2) assms(6) col_cop2_perp2__col by blast qed } then show ?thesis using ParStrict_def Par_def \Coplanar A1 A2 B1 B2\ by blast qed qed lemma parallel_existence: assumes "A \ B" shows "\ C D. C \ D \ A B Par C D \ Col P C D" proof cases assume "Col A B P" then show ?thesis using Col_perm assms par_reflexivity by blast next assume P1: "\ Col A B P" then obtain P' where P2: "Col A B P' \ A B Perp P P'" using l8_18_existence by blast then have P3: "P \ P'" using P1 by blast show ?thesis proof cases assume P4: "P' = A" have "\ Q. Per Q P A \ Cong Q P A B \ A P OS Q B" proof - have "Col A P P" using not_col_distincts by auto moreover have "\ Col A P B" by (simp add: P1 not_col_permutation_5) ultimately show ?thesis using P3 P4 assms ex_per_cong by simp qed then obtain Q where T1: "Per Q P A \ Cong Q P A B \ A P OS Q B" by auto then have T2: "P \ Q" using os_distincts by auto have T3: "A B Par P Q" proof - have "P Q Perp P A" proof - have "P \ A" using P3 P4 by auto moreover have "Col P P Q" by (simp add: col_trivial_1) ultimately show ?thesis by (metis T1 T2 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms per_perp) qed moreover have "Coplanar P A A P" using ncop_distincts by auto moreover have "Coplanar P A B P" using ncop_distincts by auto moreover have "Coplanar P A B Q" by (metis (no_types) T1 ncoplanar_perm_7 os__coplanar) moreover have "A B Perp P A" using P2 P4 by auto ultimately show ?thesis using l12_9 ncop_distincts by blast qed thus ?thesis using T2 col_trivial_1 by auto next assume T4: "P' \ A" have "\ Q. Per Q P P' \ Cong Q P A B \ P' P OS Q A" proof - have "P' \ P" using P3 by auto moreover have "A \ B" by (simp add: assms) moreover have "Col P' P P" using not_col_distincts by blast moreover have "\ Col P' P A" by (metis P1 P2 T4 col2__eq col_permutation_1) ultimately show ?thesis using ex_per_cong by blast qed then obtain Q where T5: "Per Q P P' \ Cong Q P A B \ P' P OS Q A" by blast then have T6: "P \ Q" using os_distincts by blast moreover have "A B Par P Q" proof - have "Coplanar P P' A P" using ncop_distincts by auto moreover have "Coplanar P P' A Q" by (meson T5 ncoplanar_perm_7 os__coplanar) then moreover have "Coplanar P P' B Q" by (smt P2 T4 col2_cop__cop col_permutation_5 col_transitivity_1 coplanar_perm_5) moreover have "Coplanar P P' B P" using ncop_distincts by auto moreover have "A B Perp P P'" by (simp add: P2) moreover have "P Q Perp P P'" by (metis P3 T5 T6 Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms per_perp) ultimately show ?thesis using l12_9 by blast qed moreover have "Col P P Q" by (simp add: col_trivial_1) ultimately show ?thesis by blast qed qed lemma par_col_par: assumes "C \ D'" and "A B Par C D" and "Col C D D'" shows "A B Par C D'" proof - { assume P1: "A B ParStrict C D" have "Coplanar A B C D'" using assms(2) assms(3) col2__eq col2_cop__cop par__coplanar par_neq2 by blast then have "A B Par C D'" by (smt ParStrict_def Par_def P1 assms(1) assms(3) colx not_col_distincts not_col_permutation_5) } { assume "A \ B \ C \ D \ Col A C D \ Col B C D" then have "A B Par C D'" using Par_def assms(1) assms(3) col2__eq col_permutation_2 by blast } thus ?thesis using Par_def \A B ParStrict C D \ A B Par C D'\ assms(2) by auto qed lemma parallel_existence1: assumes "A \ B" shows "\ Q. A B Par P Q" proof - obtain C D where "C \ D \ A B Par C D \ Col P C D" using assms parallel_existence by blast then show ?thesis by (metis Col_cases Par_cases par_col_par) qed lemma par_not_col: assumes "A B ParStrict C D" and "Col X A B" shows "\ Col X C D" using ParStrict_def assms(1) assms(2) by blast lemma not_strict_par1: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B C" by (smt Par_def assms(1) assms(2) assms(3) col2__eq col_permutation_2 par_not_col) lemma not_strict_par2: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B D" using Par_cases assms(1) assms(2) assms(3) not_col_permutation_4 not_strict_par1 by blast lemma not_strict_par: assumes "A B Par C D" and "Col A B X" and "Col C D X" shows "Col A B C \ Col A B D" using assms(1) assms(2) assms(3) not_strict_par1 not_strict_par2 by blast lemma not_par_not_col: assumes "A \ B" and "A \ C" and "\ A B Par A C" shows "\ Col A B C" using Par_def assms(1) assms(2) assms(3) not_col_distincts not_col_permutation_4 by blast lemma not_par_inter_uniqueness: assumes "A \ B" and "C \ D" and "\ A B Par C D" and "Col A B X" and "Col C D X" and "Col A B Y" and "Col C D Y" shows "X = Y" proof cases assume P1: "C = Y" thus ?thesis proof cases assume P2: "C = X" thus ?thesis using P1 by auto next assume "C \ X" thus ?thesis by (smt Par_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) col3 col_permutation_5 l6_21) qed next assume "C \ Y" thus ?thesis by (smt Par_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) col_permutation_2 col_permutation_4 l6_21) qed lemma inter_uniqueness_not_par: assumes "\ Col A B C" and "Col A B P" and "Col C D P" shows "\ A B Par C D" using assms(1) assms(2) assms(3) not_strict_par1 by blast lemma col_not_col_not_par: assumes "\ P. Col A B P \ Col C D P" and "\ Q. Col C D Q \ \Col A B Q" shows "\ A B Par C D" using assms(1) assms(2) colx not_strict_par par_neq2 by blast lemma par_distincts: assumes "A B Par C D" shows "A B Par C D \ A \ B \ C \ D" using assms par_neq1 par_neq2 by blast lemma par_not_col_strict: assumes "A B Par C D" and "Col C D P" and "\ Col A B P" shows "A B ParStrict C D" using Col_cases Par_def assms(1) assms(2) assms(3) col3 by blast lemma col_cop_perp2_pars: assumes "\ Col A B P" and "Col C D P" and "Coplanar A B C D" and "A B Perp P Q" and "C D Perp P Q" shows "A B ParStrict C D" proof - have P1: "C \ D" using assms(5) perp_not_eq_1 by auto then have P2: "Coplanar A B C P" using col_cop__cop assms(2) assms(3) by blast moreover have P3: "Coplanar A B D P" using col_cop__cop using P1 assms(2) assms(3) col2_cop__cop col_trivial_2 by blast have "A B Par C D" proof - have "Coplanar P A Q C" proof - have "\ Col B P A" by (simp add: assms(1) not_col_permutation_1) moreover have "Coplanar B P A Q" by (meson assms(4) ncoplanar_perm_12 perp__coplanar) moreover have "Coplanar B P A C" using P2 ncoplanar_perm_13 by blast ultimately show ?thesis using coplanar_trans_1 by auto qed then have P4: "Coplanar P Q A C" using ncoplanar_perm_2 by blast have "Coplanar P A Q D" proof - have "\ Col B P A" by (simp add: assms(1) not_col_permutation_1) moreover have "Coplanar B P A Q" by (meson assms(4) ncoplanar_perm_12 perp__coplanar) moreover have "Coplanar B P A D" using P3 ncoplanar_perm_13 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed then moreover have "Coplanar P Q A D" using ncoplanar_perm_2 by blast moreover have "Coplanar P Q B C" using P2 assms(1) assms(4) coplanar_perm_1 coplanar_perm_10 coplanar_trans_1 perp__coplanar by blast moreover have "Coplanar P Q B D" by (meson P3 assms(1) assms(4) coplanar_trans_1 ncoplanar_perm_1 ncoplanar_perm_13 perp__coplanar) ultimately show ?thesis using assms(4) assms(5) l12_9 P4 by auto qed thus ?thesis using assms(1) assms(2) par_not_col_strict by auto qed lemma all_one_side_par_strict: assumes "C \ D" and "\ P. Col C D P \ A B OS C P" shows "A B ParStrict C D" proof - have P1: "Coplanar A B C D" by (simp add: assms(2) col_trivial_2 os__coplanar) { assume "\ X. Col X A B \ Col X C D" then obtain X where P2: "Col X A B \ Col X C D" by blast have "A B OS C X" by (simp add: P2 Col_perm assms(2)) then obtain M where "A B TS C M \ A B TS X M" by (meson Col_cases P2 col124__nos) then have "False" using P2 TS_def by blast } thus ?thesis using P1 ParStrict_def by auto qed lemma par_col_par_2: assumes "A \ P" and "Col A B P" and "A B Par C D" shows "A P Par C D" using assms(1) assms(2) assms(3) par_col_par par_symmetry by blast lemma par_col2_par: assumes "E \ F" and "A B Par C D" and "Col C D E" and "Col C D F" shows "A B Par E F" by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_2 not_col_permutation_4 par_col_par par_distincts par_right_comm) lemma par_col2_par_bis: assumes "E \ F" and "A B Par C D" and "Col E F C" and "Col E F D" shows "A B Par E F" by (metis assms(1) assms(2) assms(3) assms(4) col_transitivity_1 not_col_permutation_2 par_col2_par) lemma par_strict_col_par_strict: assumes "C \ E" and "A B ParStrict C D" and "Col C D E" shows "A B ParStrict C E" proof - have P1: "C E Par A B" using Par_def Par_perm assms(1) assms(2) assms(3) par_col_par_2 by blast { assume "C E ParStrict A B" then have "A B ParStrict C E" by (metis par_strict_symmetry) } thus ?thesis using Col_cases Par_def P1 assms(2) par_strict_not_col_1 by blast qed lemma par_strict_col2_par_strict: assumes "E \ F" and "A B ParStrict C D" and "Col C D E" and "Col C D F" shows "A B ParStrict E F" by (smt ParStrict_def assms(1) assms(2) assms(3) assms(4) col2_cop__cop colx not_col_permutation_1 par_strict_neq1 par_strict_symmetry) lemma line_dec: "(Col C1 B1 B2 \ Col C2 B1 B2) \ \ (Col C1 B1 B2 \ Col C2 B1 B2)" by simp lemma par_distinct: assumes "A B Par C D" shows "A \ B \ C \ D" using assms par_neq1 par_neq2 by auto lemma par_col4__par: assumes "E \ F" and "G \ H" and "A B Par C D" and "Col A B E" and "Col A B F" and "Col C D G" and "Col C D H" shows "E F Par G H" proof - have "C D Par E F" using Par_cases assms(1) assms(3) assms(4) assms(5) par_col2_par by blast then have "E F Par C D" by (simp add: \C D Par E F\ par_symmetry) thus ?thesis using assms(2) assms(6) assms(7) par_col2_par by blast qed lemma par_strict_col4__par_strict: assumes "E \ F" and "G \ H" and "A B ParStrict C D" and "Col A B E" and "Col A B F" and "Col C D G" and "Col C D H" shows "E F ParStrict G H" proof - have "C D ParStrict E F" using Par_strict_cases assms(1) assms(3) assms(4) assms(5) par_strict_col2_par_strict by blast then have "E F ParStrict C D" by (simp add: \C D ParStrict E F\ par_strict_symmetry) thus ?thesis using assms(2) assms(6) assms(7) par_strict_col2_par_strict by blast qed lemma par_strict_one_side: assumes "A B ParStrict C D" and "Col C D P" shows "A B OS C P" proof cases assume "C = P" thus ?thesis using assms(1) assms(2) not_col_permutation_5 one_side_reflexivity par_not_col by blast next assume "C \ P" thus ?thesis using assms(1) assms(2) l12_6 par_strict_col_par_strict by blast qed lemma par_strict_all_one_side: assumes "A B ParStrict C D" shows "\ P. Col C D P \ A B OS C P" using assms par_strict_one_side by blast lemma inter_trivial: assumes "\ Col A B X" shows "X Inter A X B X" by (metis Col_perm Inter_def assms col_trivial_1) lemma inter_sym: assumes "X Inter A B C D" shows "X Inter C D A B" proof - obtain P where P1: "Col P C D \ \ Col P A B" using Inter_def assms by auto have P2: "A \ B" using P1 col_trivial_2 by blast then show ?thesis proof cases assume "A = X" have "Col B A B" by (simp add: col_trivial_3) { assume P3: "Col B C D" have "Col P A B" proof - have "C \ D" using Inter_def assms by blast moreover have "Col C D P" using P1 not_col_permutation_2 by blast moreover have "Col C D A" using Inter_def \A = X\ assms by auto moreover have "Col C D B" using P3 not_col_permutation_2 by blast ultimately show ?thesis using col3 by blast qed then have "False" by (simp add: P1) } then have "\ Col B C D" by auto then show ?thesis using Inter_def P2 assms by (meson col_trivial_3) next assume P5: "A \ X" have P6: "Col A A B" using not_col_distincts by blast { assume P7: "Col A C D" have "Col A P X" proof - have "C \ D" using Inter_def assms by auto moreover have "Col C D A" using Col_cases P7 by blast moreover have "Col C D P" using Col_cases P1 by auto moreover have "Col C D X" using Inter_def assms by auto ultimately show ?thesis using col3 by blast qed then have "Col P A B" by (metis (full_types) Col_perm Inter_def P5 assms col_transitivity_2) then have "False" by (simp add: P1) } then have "\ Col A C D" by auto then show ?thesis by (meson Inter_def P2 assms col_trivial_1) qed qed lemma inter_left_comm: assumes "X Inter A B C D" shows "X Inter B A C D" using Col_cases Inter_def assms by auto lemma inter_right_comm: assumes "X Inter A B C D" shows "X Inter A B D C" by (metis assms inter_left_comm inter_sym) lemma inter_comm: assumes "X Inter A B C D" shows "X Inter B A D C" using assms inter_left_comm inter_right_comm by blast lemma l12_17: assumes "A \ B" and "P Midpoint A C" and "P Midpoint B D" shows "A B Par C D" proof cases assume P1: "Col A B P" thus ?thesis proof cases assume "A = P" thus ?thesis using assms(1) assms(2) assms(3) cong_diff_2 is_midpoint_id midpoint_col midpoint_cong not_par_not_col by blast next assume P2: "A \ P" thus ?thesis proof cases assume "B = P" thus ?thesis by (metis assms(1) assms(2) assms(3) midpoint_col midpoint_distinct_2 midpoint_distinct_3 not_par_not_col par_comm) next assume P3: "B \ P" have P4: "Col B P D" using assms(3) midpoint_col not_col_permutation_4 by blast have P5: "Col A P C" using assms(2) midpoint_col not_col_permutation_4 by blast then have P6: "Col B C P" using P1 P2 col_transitivity_2 not_col_permutation_3 not_col_permutation_5 by blast have "C \ D" using assms(1) assms(2) assms(3) l7_9 by blast moreover have "Col A C D" using P1 P3 P4 P6 col3 not_col_permutation_3 not_col_permutation_5 by blast moreover have "Col B C D" using P3 P4 P6 col_trivial_3 colx by blast ultimately show ?thesis by (simp add: Par_def assms(1)) qed qed next assume T1: "\ Col A B P" then obtain E where T2: "Col A B E \ A B Perp P E" using l8_18_existence by blast have T3: "A \ P" using T1 col_trivial_3 by blast then show ?thesis proof cases assume T4: "A = E" then have T5: "Per P A B" using T2 l8_2 perp_per_1 by blast obtain B' where T6: "Bet B A B' \ Cong A B' B A" using segment_construction by blast obtain D' where T7: "Bet B' P D' \ Cong P D' B' P" using segment_construction by blast have T8: "C Midpoint D D'" using T6 T7 assms(2) assms(3) midpoint_def not_cong_3412 symmetry_preserves_midpoint by blast have "Col A B B'" using Col_cases Col_def T6 by blast then have T9: "Per P A B'" using per_col T5 assms(1) by blast obtain B'' where T10: "A Midpoint B B'' \ Cong P B P B''" using Per_def T5 by auto then have "B' = B''" using T6 cong_symmetry midpoint_def symmetric_point_uniqueness by blast then have "Cong P D P D'" by (metis Cong_perm Midpoint_def T10 T7 assms(3) cong_inner_transitivity) then have T12: "Per P C D" using Per_def T8 by auto then have T13: "C PerpAt P C C D" by (metis T3 assms(1) assms(2) assms(3) l7_3_2 per_perp_in sym_preserve_diff) have T14: "P \ C" using T3 assms(2) is_midpoint_id_2 by auto have T15: "C \ D" using assms(1) assms(2) assms(3) l7_9 by auto have T15A: "C C Perp C D \ P C Perp C D" using T12 T14 T15 per_perp by auto { assume "C C Perp C D" then have "A B Par C D" using perp_distinct by auto } { assume "P C Perp C D" have "A B Par C D" proof - have "Coplanar P A A C" using ncop_distincts by blast moreover have "Coplanar P A A D" using ncop_distincts by blast moreover have "Coplanar P A B C" by (simp add: assms(2) coplanar_perm_1 midpoint__coplanar) moreover have "Coplanar P A B D" using assms(3) midpoint_col ncop__ncols by blast moreover have "A B Perp P A" using T2 T4 by auto moreover have "C D Perp P A" proof - have "P A Perp C D" proof - have "P \ A" using T3 by auto moreover have "P C Perp C D" using T14 T15 T12 per_perp by blast moreover have "Col P C A" by (simp add: assms(2) l7_2 midpoint_col) ultimately show ?thesis using perp_col by blast qed then show ?thesis using Perp_perm by blast qed ultimately show ?thesis using l12_9 by blast qed } then show ?thesis using T15A using \C C Perp C D \ A B Par C D\ by blast next assume S1B: "A \ E" obtain F where S2: "Bet E P F \ Cong P F E P" using segment_construction by blast then have S2A: "P Midpoint E F" using midpoint_def not_cong_3412 by blast then have S3: "Col C D F" using T2 assms(2) assms(3) mid_preserves_col by blast obtain A' where S4: "Bet A E A' \ Cong E A' A E" using segment_construction by blast obtain C' where S5: "Bet A' P C' \ Cong P C' A' P" using segment_construction by blast have S6: "F Midpoint C C'" using S4 S5 S2A assms(2) midpoint_def not_cong_3412 symmetry_preserves_midpoint by blast have S7: "Per P E A" using T2 col_trivial_3 l8_16_1 by blast have S8: "Cong P C P C'" proof - have "Cong P C P A" using Cong_perm Midpoint_def assms(2) by blast moreover have "Cong P A P C'" proof - obtain A'' where S9: "E Midpoint A A'' \ Cong P A P A''" using Per_def S7 by blast have S10: "A' = A''" using Cong_perm Midpoint_def S4 S9 symmetric_point_uniqueness by blast then have "Cong P A P A'" using S9 by auto moreover have "Cong P A' P C'" using Cong_perm S5 by blast ultimately show ?thesis using cong_transitivity by blast qed ultimately show ?thesis using cong_transitivity by blast qed then have S9: "Per P F C" using S6 Per_def by blast then have "F PerpAt P F F C" by (metis S2 S2A T1 T2 S1B assms(2) cong_diff_3 l7_9 per_perp_in) then have "F PerpAt F P C F" using Perp_in_perm by blast then have S10: "F P Perp C F \ F F Perp C F" using l8_15_2 perp_in_col by blast { assume S11: "F P Perp C F" have "Coplanar P E A C" proof - have "Col P E P \ Col A C P" using assms(2) col_trivial_3 midpoint_col not_col_permutation_2 by blast then show ?thesis using Coplanar_def by blast qed moreover have "Coplanar P E A D" proof - have "Col P D B \ Col E A B" using Mid_cases T2 assms(3) midpoint_col not_col_permutation_1 by blast then show ?thesis using Coplanar_def by blast qed moreover have "Coplanar P E B C" by (metis S1B T2 calculation(1) col2_cop__cop col_transitivity_1 ncoplanar_perm_5 not_col_permutation_5) moreover have "Coplanar P E B D" by (metis S1B T2 calculation(2) col2_cop__cop col_transitivity_1 ncoplanar_perm_5 not_col_permutation_5) moreover have "C D Perp P E" proof - have "C \ D" using assms(1) assms(2) assms(3) sym_preserve_diff by blast moreover have "P F Perp C F" using Perp_perm S11 by blast moreover have "Col P F E" by (simp add: Col_def S2) moreover have "Col C F D" using Col_perm S3 by blast ultimately show ?thesis using per_col by (smt Perp_cases S2 col_trivial_3 cong_diff perp_col4 perp_not_eq_1) qed ultimately have "A B Par C D" using T2 l12_9 by blast } { assume "F F Perp C F" then have "A B Par C D" using perp_distinct by blast } thus ?thesis using S10 \F P Perp C F \ A B Par C D\ by blast qed qed lemma l12_18_a: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A B Par C D" proof - have "P Midpoint A C \ P Midpoint B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_21 by blast then show ?thesis using assms(3) l12_17 not_col_distincts by blast qed lemma l12_18_b: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "B C Par D A" by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) cong_symmetry inter_uniqueness_not_par l12_18_a l6_21 not_col_distincts) lemma l12_18_c: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "B D TS A C" proof - have "P Midpoint A C \ P Midpoint B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l7_21 by blast then show ?thesis proof - have "A C TS B D" by (metis Col_cases Tarski_neutral_dimensionless.mid_two_sides Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\ assms(3)) then have "\ Col B D A" by (meson Col_cases Tarski_neutral_dimensionless.mid_preserves_col Tarski_neutral_dimensionless.ts__ncol Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\ l7_2) then show ?thesis by (meson Tarski_neutral_dimensionless.mid_two_sides Tarski_neutral_dimensionless_axioms \P Midpoint A C \ P Midpoint B D\) qed qed lemma l12_18_d: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A C TS B D" by (metis (no_types, lifting) Col_cases TS_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l12_18_c not_col_distincts not_cong_2143 not_cong_4321) lemma l12_18: assumes "Cong A B C D" and "Cong B C D A" and "\ Col A B C" and "B \ D" and "Col A P C" and "Col B P D" shows "A B Par C D \ B C Par D A \ B D TS A C \ A C TS B D" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l12_18_a l12_18_b l12_18_c l12_18_d by auto lemma par_two_sides_two_sides: assumes "A B Par C D" and "B D TS A C" shows "A C TS B D" by (metis Par_def TS_def assms(1) assms(2) invert_one_side invert_two_sides l12_6 l9_31 not_col_permutation_4 one_side_symmetry os_ts1324__os pars__os3412) lemma par_one_or_two_sides: assumes "A B ParStrict C D" shows "(A C TS B D \ B D TS A C) \ (A C OS B D \ B D OS A C)" by (smt Par_def assms invert_one_side l12_6 l9_31 not_col_permutation_3 os_ts1324__os par_strict_not_col_1 par_strict_not_col_2 par_two_sides_two_sides pars__os3412 two_sides_cases) lemma l12_21_b: assumes "A C TS B D" and "B A C CongA D C A" shows "A B Par C D" proof - have P1: "\ Col A B C" using TS_def assms(1) not_col_permutation_4 by blast then have P2: "A \ B" using col_trivial_1 by auto have P3: "C \ D" using assms(1) ts_distincts by blast then obtain D' where P4: "C Out D D' \ Cong C D' A B" using P2 segment_construction_3 by blast have P5: "B A C CongA D' C A" proof - have "A Out B B" using P2 out_trivial by auto moreover have "A Out C C" using P1 col_trivial_3 out_trivial by force moreover have "C Out D' D" by (simp add: P4 l6_6) moreover have "C Out A A" using P1 not_col_distincts out_trivial by auto ultimately show ?thesis using assms(2) l11_10 by blast qed then have P6: "Cong D' A B C" using Cong_perm P4 cong_pseudo_reflexivity l11_49 by blast have P7: "A C TS D' B" proof - have "A C TS D B" by (simp add: assms(1) l9_2) moreover have "Col C A C" using col_trivial_3 by auto ultimately show ?thesis using P4 l9_5 by blast qed then obtain M where P8: "Col M A C \ Bet D' M B" using TS_def by blast have "B \ D'" using P7 not_two_sides_id by blast then have "M Midpoint A C \ M Midpoint B D'" by (metis Col_cases P1 P4 P6 P8 bet_col l7_21 not_cong_3412) then have "A B Par C D'" using P2 l12_17 by blast thus ?thesis by (meson P3 P4 Tarski_neutral_dimensionless.par_col_par Tarski_neutral_dimensionless_axioms l6_6 out_col) qed lemma l12_22_aux: assumes "P \ A" and "A \ C" and "Bet P A C" and "P A OS B D" and "B A P CongA D C P" shows "A B Par C D" proof - have P1: "P \ C" using CongA_def assms(5) by blast obtain B' where P2: "Bet B A B' \ Cong A B' B A" using segment_construction by blast have P3: "P A B CongA C A B'" by (metis CongA_def P2 assms(2) assms(3) assms(5) cong_reverse_identity l11_14) have P4: "D C A CongA D C P" by (metis Col_def assms(2) assms(3) assms(4) bet_out_1 col124__nos l6_6 out2__conga out_trivial) have P5: "A B' Par C D" proof - have "\ Col B P A" using assms(4) col123__nos not_col_permutation_2 by blast then have "P A TS B B'" by (metis P2 assms(4) bet__ts cong_reverse_identity invert_two_sides not_col_permutation_3 os_distincts) then have "A C TS B' D" by (meson assms(2) assms(3) assms(4) bet_col bet_col1 col_preserves_two_sides l9_2 l9_8_2) moreover have "B' A C CongA D C A" proof - have "B' A C CongA B A P" by (simp add: P3 conga_comm conga_sym) moreover have "B A P CongA D C A" using P4 assms(5) not_conga not_conga_sym by blast ultimately show ?thesis using not_conga by blast qed ultimately show ?thesis using l12_21_b by blast qed have "C D Par A B" proof - have "A \ B" using assms(4) os_distincts by blast moreover have "C D Par A B'" using P5 par_symmetry by blast moreover have "Col A B' B" by (simp add: Col_def P2) ultimately show ?thesis using par_col_par by blast qed thus ?thesis using Par_cases by blast qed lemma l12_22_b: assumes "P Out A C" and "P A OS B D" and "B A P CongA D C P" shows "A B Par C D" proof cases assume "A = C" then show ?thesis using assms(2) assms(3) conga_comm conga_os__out not_par_not_col os_distincts out_col by blast next assume P1: "A \ C" { assume "Bet P A C" then have "A B Par C D" using P1 assms(2) assms(3) conga_diff2 l12_22_aux by blast } { assume P2: "Bet P C A" have "C D Par A B" proof - have "P C OS D B" using assms(1) assms(2) col_one_side one_side_symmetry out_col out_diff2 by blast moreover have "D C P CongA B A P" using assms(3) not_conga_sym by blast then show ?thesis by (metis P1 P2 assms(1) calculation l12_22_aux out_distinct) qed then have "A B Par C D" using Par_cases by auto } then show ?thesis using Out_def \Bet P A C \ A B Par C D\ assms(1) by blast qed lemma par_strict_par: assumes "A B ParStrict C D" shows "A B Par C D" using Par_def assms by auto lemma par_strict_distinct: assumes "A B ParStrict C D" shows " A \ B \ C \ D" using assms par_strict_neq1 par_strict_neq2 by auto lemma col_par: assumes "A \ B" and "B \ C" and "Col A B C" shows "A B Par B C" by (simp add: Par_def assms(1) assms(2) assms(3) col_trivial_1) lemma acute_col_perp__out: assumes "Acute A B C" and "Col B C A'" and "B C Perp A A'" shows "B Out A' C" proof - { assume P1: "\ Col B C A" then obtain B' where P2: "B C Perp B' B \ B C OS A B'" using assms(2) l10_15 os_distincts by blast have P3: "\ Col B' B C" using P2 col124__nos col_permutation_1 by blast { assume "Col B B' A" then have "A B C LtA A B C" using P2 acute_one_side_aux acute_sym assms(1) one_side_not_col124 by blast then have "False" by (simp add: nlta) } then have P4: "\ Col B B' A" by auto have P5: "B B' ParStrict A A'" proof - have "B B' Par A A'" proof - have "Coplanar B C B A" using ncop_distincts by blast moreover have "Coplanar B C B A'" using ncop_distincts by blast moreover have "Coplanar B C B' A" using P2 coplanar_perm_1 os__coplanar by blast moreover have "Coplanar B C B' A'" using assms(2) ncop__ncols by auto moreover have "B B' Perp B C" using P2 Perp_perm by blast moreover have "A A' Perp B C" using Perp_perm assms(3) by blast ultimately show ?thesis using l12_9 by auto qed moreover have "Col A A' A" by (simp add: col_trivial_3) moreover have "\ Col B B' A" by (simp add: P4) ultimately show ?thesis using par_not_col_strict by auto qed then have P6: "\ Col B B' A'" using P5 par_strict_not_col_4 by auto then have "B B' OS A' C" proof - have "B B' OS A' A" using P5 l12_6 one_side_symmetry by blast moreover have "B B' OS A C" using P2 acute_one_side_aux acute_sym assms(1) one_side_symmetry by blast ultimately show ?thesis using one_side_transitivity by blast qed then have "B Out A' C" using Col_cases assms(2) col_one_side_out by blast } then show ?thesis using assms(2) assms(3) perp_not_col2 by blast qed lemma acute_col_perp__out_1: assumes "Acute A B C" and "Col B C A'" and "B A Perp A A'" shows "B Out A' C" proof - obtain A0 where P1: "Bet A B A0 \ Cong B A0 A B" using segment_construction by blast obtain C0 where P2: "Bet C B C0 \ Cong B C0 C B" using segment_construction by blast have P3: "\ Col B A A'" using assms(3) col_trivial_2 perp_not_col2 by blast have "Bet A' B C0" proof - have P4: "Col A' B C0" using P2 acute_distincts assms(1) assms(2) bet_col col_transitivity_2 not_col_permutation_4 by blast { assume P5: "B Out A' C0" have "B Out A A0" proof - have "Bet C B A'" by (smt Bet_perm Col_def P2 P5 assms(2) between_exchange3 not_bet_and_out outer_transitivity_between2) then have "A B C CongA A0 B A'" using P1 P3 acute_distincts assms(1) cong_diff_4 l11_14 not_col_distincts by blast then have "Acute A' B A0" using acute_conga__acute acute_sym assms(1) by blast moreover have "B A0 Perp A' A" proof - have "B \ A0" using P1 P3 col_trivial_1 cong_reverse_identity by blast moreover have "B A Perp A' A" using Perp_perm assms(3) by blast moreover have "Col B A A0" using P1 bet_col not_col_permutation_4 by blast ultimately show ?thesis using perp_col by blast qed ultimately show ?thesis using Col_cases P1 acute_col_perp__out bet_col by blast qed then have "False" using P1 not_bet_and_out by blast } moreover then have "\ B Out A' C0" by auto ultimately show ?thesis using l6_4_2 P4 by blast qed then show ?thesis by (metis P2 P3 acute_distincts assms(1) cong_diff_3 l6_2 not_col_distincts) qed lemma conga_inangle_per2__inangle: assumes "Per A B C" and "T InAngle A B C" and "P B A CongA P B C" and "Per B P T" and "Coplanar A B C P" shows "P InAngle A B C" proof cases assume "P = T" then show ?thesis by (simp add: assms(2)) next assume P1: "P \ T" obtain P' where P2: "P' InAngle A B C \ P' B A CongA P' B C" using CongA_def angle_bisector assms(3) by presburger have P3: "Acute P' B A" using P2 acute_sym assms(1) conga_inangle_per__acute by blast have P4: "\ Col A B C" using assms(1) assms(3) conga_diff2 conga_diff56 l8_9 by blast have P5: "Col B P P'" proof - have "\ B Out A C" using Col_cases P4 out_col by blast moreover have "Coplanar A B P P'" proof - have T1: "\ Col C A B" using Col_perm P4 by blast moreover have "Coplanar C A B P" using assms(5) ncoplanar_perm_8 by blast moreover have "Coplanar C A B P'" using P2 inangle__coplanar ncoplanar_perm_21 by blast ultimately show ?thesis using coplanar_trans_1 by blast qed moreover have "Coplanar B C P P'" proof - have "Coplanar A B C P" by (meson P2 bet__coplanar calculation(1) calculation(2) col_in_angle_out coplanar_perm_18 coplanar_trans_1 inangle__coplanar l11_21_a l6_6 l6_7 not_col_permutation_4 not_col_permutation_5) have "Coplanar A B C P'" using P2 inangle__coplanar ncoplanar_perm_18 by blast then show ?thesis using P4 \Coplanar A B C P\ coplanar_trans_1 by blast qed ultimately show ?thesis using conga2_cop2__col P2 assms(3) by blast qed have "B Out P P'" proof - have "Acute T B P'" using P2 acute_sym assms(1) assms(2) conga_inangle2_per__acute by blast moreover have "B P' Perp T P" by (metis P1 P5 acute_distincts assms(3) assms(4) calculation col_per_perp conga_distinct l8_2 not_col_permutation_4) ultimately show ?thesis using Col_cases P5 acute_col_perp__out by blast qed then show ?thesis using Out_cases P2 in_angle_trans inangle_distincts out341__inangle by blast qed lemma perp_not_par: assumes "A B Perp X Y" shows "\ A B Par X Y" proof - obtain P where P1: "P PerpAt A B X Y" using Perp_def assms by blast { assume P2: "A B Par X Y" { assume P3: "A B ParStrict X Y" then have "False" proof - have "Col P A B" using Col_perm P1 perp_in_col by blast moreover have "Col P X Y" using P1 col_permutation_2 perp_in_col by blast ultimately show ?thesis using P3 par_not_col by blast qed } { assume P4: "A \ B \ X \ Y \ Col A X Y \ Col B X Y" then have "False" proof cases assume "A = Y" thus ?thesis using P4 assms not_col_permutation_1 perp_not_col by blast next assume "A \ Y" thus ?thesis using Col_perm P4 Perp_perm assms perp_not_col2 by blast qed } then have "False" using Par_def P2 \A B ParStrict X Y \ False\ by auto } thus ?thesis by auto qed lemma cong_conga_perp: assumes "B P TS A C" and "Cong A B C B" and "A B P CongA C B P" shows "A C Perp B P" proof - have P1: " \ Col A B P" using TS_def assms(1) by blast then have P2: "B \ P" using col_trivial_2 by blast have P3: "A \ B" using assms(1) ts_distincts by blast have P4: "C \ B" using assms(1) ts_distincts by auto have P5: "A \ C" using assms(1) not_two_sides_id by auto show ?thesis proof cases assume P6: "Bet A B C" then have "Per P B A" by (meson Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms assms(3) l11_18_2) then show ?thesis using P2 P3 P5 Per_perm P6 bet_col per_perp perp_col by blast next assume P7: "\ Bet A B C" obtain T where P7A: "Col T B P \ Bet A T C" using TS_def assms(1) by auto then have P8: "B \ T" using P7 by blast then have P9: "T B A CongA T B C" by (meson Col_cases P7A Tarski_neutral_dimensionless.col_conga__conga Tarski_neutral_dimensionless.conga_comm Tarski_neutral_dimensionless_axioms assms(3)) then have P10: "Cong T A T C" using assms(2) cong2_conga_cong cong_reflexivity not_cong_2143 by blast then have P11: "T Midpoint A C" using P7A midpoint_def not_cong_2134 by blast have P12: "Per B T A" using P11 Per_def assms(2) not_cong_2143 by blast then show ?thesis proof - have "A C Perp B T" by (metis P11 P12 P5 P8 col_per_perp midpoint_col midpoint_distinct_1) moreover have "B \ T" by (simp add: P8) moreover have "T \ A" using P1 P7A by blast moreover have "C \ T" using P10 P5 cong_identity by blast moreover have "C \ A" using P5 by auto moreover have "Col T A C" by (meson P7A bet_col not_col_permutation_4) ultimately show ?thesis using P2 P7A not_col_permutation_4 perp_col1 by blast qed qed qed lemma perp_inter_exists: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P" proof - obtain P where "P PerpAt A B C D" using Perp_def assms by auto then show ?thesis using perp_in_col by blast qed lemma perp_inter_perp_in: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P \ P PerpAt A B C D" by (meson Perp_def Tarski_neutral_dimensionless.perp_in_col Tarski_neutral_dimensionless_axioms assms) end context Tarski_2D begin lemma l12_9_2D: assumes "A1 A2 Perp C1 C2" and "B1 B2 Perp C1 C2" shows "A1 A2 Par B1 B2" using l12_9 all_coplanar assms(1) assms(2) by auto end context Tarski_neutral_dimensionless begin subsection "Tarski: Chapter 13" subsubsection "Introduction" lemma per2_col_eq: assumes "A \ P" and "A \ P'" and "Per A P B" and "Per A P' B" and "Col P A P'" shows "P = P'" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) col_per2_cases l6_16_1 l8_2 not_col_permutation_3) lemma per2_preserves_diff: assumes "PO \ A'" and "PO \ B'" and "Col PO A' B'" and "Per PO A' A" and "Per PO B' B" and "A' \ B'" shows "A \ B" using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) not_col_permutation_4 per2_col_eq by blast lemma per23_preserves_bet: assumes "Bet A B C" and "A \ B'" and "A \ C'" and "Col A B' C'" and "Per A B' B" and "Per A C' C" shows "Bet A B' C'" proof - have P1: "Col A B C" by (simp add: assms(1) bet_col) show ?thesis proof cases assume P2: "B = B'" then have "Col A C' C" using P1 assms(2) assms(4) col_transitivity_1 by blast then have P4: "A = C' \ C = C'" by (simp add: assms(6) l8_9) { assume "A = C'" then have "Bet A B' C'" using assms(3) by auto } { assume "C = C'" then have "Bet A B' C'" using P2 assms(1) by auto } then show ?thesis using P4 assms(3) by auto next assume T1: "B \ B'" have T2: "A \ C" using assms(3) assms(6) l8_8 by auto have T3: "C \ C'" using P1 T1 assms(2) assms(3) assms(4) assms(5) col_trivial_3 colx l8_9 not_col_permutation_5 by blast have T3A: "A B' Perp B' B" using T1 assms(2) assms(5) per_perp by auto have T3B: "A C' Perp C' C" using T3 assms(3) assms(6) per_perp by auto have T4: "B B' Par C C'" proof - have "Coplanar A B' B C" using P1 ncop__ncols by blast moreover have "Coplanar A B' B C'" using assms(4) ncop__ncols by blast moreover have "Coplanar A B' B' C" using ncop_distincts by blast moreover have "B B' Perp A B'" using Perp_perm \A B' Perp B' B\ by blast moreover have "C C' Perp A B'" using Col_cases Perp_cases T3B assms(2) assms(4) perp_col1 by blast ultimately show ?thesis using l12_9 bet__coplanar between_trivial by auto qed moreover have "Bet A B' C'" proof cases assume "B = C" then show ?thesis by (metis T1 Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(4) assms(5) calculation l6_16_1 l6_6 or_bet_out out_diff1 par_id) next assume T6: "B \ C" have T7: "\ Col A B' B" using T1 assms(2) assms(5) l8_9 by blast have T8: "\ Col A C' C" using T3 assms(3) assms(6) l8_9 by blast have T9: "B' \ C'" using P1 T6 assms(2) assms(5) assms(6) col_per2__per col_permutation_1 l8_2 l8_8 by blast have T10: "B B' ParStrict C C' \ (B \ B' \ C \ C' \ Col B C C' \ Col B' C C')" using Par_def calculation by blast { assume T11: "B B' ParStrict C C'" then have T12: "B B' OS C' C" using l12_6 one_side_symmetry by blast have "B B' TS A C" using Col_cases T6 T7 assms(1) bet__ts by blast then have "Bet A B' C'" using T12 assms(4) l9_5 l9_9 not_col_distincts or_bet_out by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A B' C'" using Col_def T6 T8 assms(1) col_transitivity_2 by blast } then show ?thesis using T10 \B B' ParStrict C C' \ Bet A B' C'\ by blast qed ultimately show ?thesis by (smt P1 Par_def T1 T2 assms(4) col_transitivity_2 not_col_permutation_1 par_strict_not_col_2) qed qed lemma per23_preserves_bet_inv: assumes "Bet A B' C'" and "A \ B'" and "Col A B C" and "Per A B' B" and "Per A C' C" shows "Bet A B C" proof cases assume T1: "B = B'" then have "Col A C' C" using Col_def assms(1) assms(2) assms(3) col_transitivity_1 by blast then have T2: "A = C' \ C = C'" by (simp add: assms(5) l8_9) { assume "A = C'" then have "Bet A B C" using assms(1) assms(2) between_identity by blast } { assume "C = C'" then have "Bet A B C" by (simp add: T1 assms(1)) } then show ?thesis using T2 \A = C' \ Bet A B C\ by auto next assume P1: "B \ B'" then have P2: "A B' Perp B' B" using assms(2) assms(4) per_perp by auto have "Per A C' C" by (simp add: assms(5)) then have P2: "C' PerpAt A C' C' C" by (metis (mono_tags, lifting) Col_cases P1 assms(1) assms(2) assms(3) assms(4) bet_col bet_neq12__neq col_transitivity_1 l8_9 per_perp_in) then have P3: "A C' Perp C' C" using perp_in_perp by auto then have "C' \ C" using \A C' Perp C' C\ perp_not_eq_2 by auto have "C' PerpAt C' A C C'" by (simp add: Perp_in_perm P2) then have "(C' A Perp C C') \ (C' C' Perp C C')" using Perp_def by blast have "A \ C'" using assms(1) assms(2) between_identity by blast { assume "C' A Perp C C'" have "Col A B' C'" using assms(1) by (simp add: Col_def) have "A B' Perp C' C" using Col_cases \A C' Perp C' C\ \Col A B' C'\ assms(2) perp_col by blast have P7: "B' B Par C' C" proof - have "Coplanar A B' B' C'" using ncop_distincts by blast moreover have "Coplanar A B' B' C" using ncop_distincts by auto moreover have "Coplanar A B' B C'" using Bet_perm assms(1) bet__coplanar ncoplanar_perm_20 by blast moreover have "Coplanar A B' B C" using assms(3) ncop__ncols by auto moreover have "B' B Perp A B'" by (metis P1 Perp_perm assms(2) assms(4) per_perp) moreover have "C' C Perp A B'" using Perp_cases \A B' Perp C' C\ by auto ultimately show ?thesis using l12_9 by blast qed have "Bet A B C" proof cases assume "B = C" then show ?thesis by (simp add: between_trivial) next assume T1: "B \ C" have T2: "B' B ParStrict C' C \ (B' \ B \ C' \ C \ Col B' C' C \ Col B C' C)" using P7 Par_def by auto { assume T3: "B' B ParStrict C' C" then have "B' \ C'" using not_par_strict_id by auto have "\ X. Col X B' B \ Col X B' C" using col_trivial_1 by blast have "B' B OS C' C" by (simp add: T3 l12_6) have "B' B TS A C'" by (metis Bet_cases T3 assms(1) assms(2) bet__ts l9_2 par_strict_not_col_1) then have T8: "B' B TS C A" using \B' B OS C' C\ l9_2 l9_8_2 by blast then obtain T where T9: "Col T B' B \ Bet C T A" using TS_def by auto have "\ Col A C B'" using T8 assms(3) not_col_permutation_2 not_col_permutation_3 ts__ncol by blast then have "T = B" by (metis Col_def Col_perm T9 assms(3) colx) then have "Bet A B C" using Bet_cases T9 by auto } { assume "B' \ B \ C' \ C \ Col B' C' C \ Col B C' C" then have "Col A B' B" by (metis Col_perm T1 assms(3) l6_16_1) then have "A = B' \ B = B'" using assms(4) l8_9 by auto then have "Bet A B C" by (simp add: P1 assms(2)) } then show ?thesis using T2 \B' B ParStrict C' C \ Bet A B C\ by auto qed } then show ?thesis by (simp add: P3 perp_comm) qed lemma per13_preserves_bet: assumes "Bet A B C" and "B \ A'" and "B \ C'" and "Col A' B C'" and "Per B A' A" and "Per B C' C" shows "Bet A' B C'" by (smt Col_cases Tarski_neutral_dimensionless.per23_preserves_bet_inv Tarski_neutral_dimensionless_axioms assms(1) assms(4) assms(5) assms(6) bet_col between_equality between_symmetry per_distinct third_point) lemma per13_preserves_bet_inv: assumes "Bet A' B C'" and "B \ A'" and "B \ C'" and "Col A B C" and "Per B A' A" and "Per B C' C" shows "Bet A B C" proof - have P1: "Col A' B C'" by (simp add: Col_def assms(1)) show ?thesis proof cases assume "A = A'" then show ?thesis using P1 assms(1) assms(3) assms(4) assms(6) col_transitivity_2 l8_9 not_bet_distincts by blast next assume "A \ A'" show ?thesis by (metis Col_cases P1 Tarski_neutral_dimensionless.per23_preserves_bet Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) between_equality between_symmetry third_point) qed qed lemma per3_preserves_bet1: assumes "Col PO A B" and "Bet A B C" and "PO \ A'" and "PO \ B'" and "PO \ C'" and "Per PO A' A" and "Per PO B' B" and "Per PO C' C" and "Col A' B' C'" and "Col PO A' B'" shows "Bet A' B' C'" proof cases assume "A = B" then show ?thesis using assms(10) assms(3) assms(4) assms(6) assms(7) between_trivial2 per2_preserves_diff by blast next assume P1: "A \ B" show ?thesis proof cases assume P2: "A = A'" show ?thesis proof cases assume P3: "B = B'" then have "Col PO C C'" by (metis (no_types, opaque_lifting) Col_def P1 P2 assms(1) assms(2) assms(9) col_transitivity_1) then have "C = C'" using assms(5) assms(8) l8_9 not_col_permutation_5 by blast then show ?thesis using P2 P3 assms(2) by blast next assume P4: "B \ B'" show ?thesis proof cases assume "A = B'" then show ?thesis using P2 between_trivial2 by auto next assume "A \ B'" have "A \ C" using P1 assms(2) between_identity by blast have P7: "\ Col PO B' B" using P4 assms(4) assms(7) l8_9 by blast show ?thesis using P2 P7 assms(1) assms(10) assms(3) col_transitivity_1 by blast qed qed next assume R1: "A \ A'" show ?thesis proof cases assume R2: "A' = B'" then show ?thesis by (simp add: between_trivial2) next assume R3: "A' \ B'" show ?thesis proof cases assume "B = C" have "B' = C'" by (metis Tarski_neutral_dimensionless.per2_col_eq Tarski_neutral_dimensionless_axioms \A' \ B'\ \B = C\ assms(10) assms(4) assms(5) assms(7) assms(8) assms(9) col_transitivity_2 not_col_permutation_2) then show ?thesis by (simp add: between_trivial) next assume R4: "B \ C" show ?thesis proof cases assume "B = B'" then show ?thesis by (metis R1 assms(1) assms(10) assms(3) assms(4) assms(6) l6_16_1 l8_9 not_col_permutation_2) next assume R5: "B \ B'" show ?thesis proof cases assume "A' = B" then show ?thesis using R5 assms(10) assms(4) assms(7) col_permutation_5 l8_9 by blast next assume R5A: "A' \ B" have R6: "C \ C'" by (metis P1 R1 R3 assms(1) assms(10) assms(2) assms(3) assms(5) assms(6) assms(9) bet_col col_permutation_1 col_trivial_2 l6_21 l8_9) have R7: "A A' Perp PO A'" by (metis Perp_cases R1 assms(3) assms(6) per_perp) have R8: "C C' Perp PO A'" by (smt Perp_cases R3 R6 assms(10) assms(3) assms(5) assms(8) assms(9) col2__eq col3 col_per_perp col_trivial_2 l8_2 per_perp) have "A A' Par C C'" proof - have "Coplanar PO A' A C" using P1 assms(1) assms(2) bet_col col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' A C'" using R3 assms(10) assms(9) col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' A' C" using ncop_distincts by blast moreover have "Coplanar PO A' A' C'" using ncop_distincts by blast ultimately show ?thesis using l12_9 R7 R8 by blast qed have S1: "B B' Perp PO A'" by (metis Col_cases Per_cases Perp_perm R5 assms(10) assms(3) assms(4) assms(7) col_per_perp) have "A A' Par B B'" proof - have "Coplanar PO A' A B" using assms(1) ncop__ncols by auto moreover have "Coplanar PO A' A B'" using assms(10) ncop__ncols by auto moreover have "Coplanar PO A' A' B" using ncop_distincts by auto moreover have "Coplanar PO A' A' B'" using ncop_distincts by auto moreover have "A A' Perp PO A'" by (simp add: R7) moreover have "B B' Perp PO A'" by (simp add: S1) ultimately show ?thesis using l12_9 by blast qed { assume "A A' ParStrict B B'" then have "A A' OS B B'" by (simp add: l12_6) have "B B' TS A C" using R4 \A A' ParStrict B B'\ assms(2) bet__ts par_strict_not_col_3 by auto have "B B' OS A A'" using \A A' ParStrict B B'\ pars__os3412 by auto have "B B' TS A' C" using \B B' OS A A'\ \B B' TS A C\ l9_8_2 by blast have "Bet A' B' C'" proof cases assume "C = C'" then show ?thesis using R6 by auto next assume "C \ C'" have "C C' Perp PO A'" by (simp add: R8) have Q2: "B B' Par C C'" proof - have "Coplanar PO A' B C" by (metis P1 assms(1) assms(2) bet_col col_transitivity_1 colx ncop__ncols not_col_permutation_5) moreover have "Coplanar PO A' B C'" using R3 assms(10) assms(9) col_trivial_2 colx ncop__ncols by blast moreover have "Coplanar PO A' B' C" by (simp add: assms(10) col__coplanar) moreover have "Coplanar PO A' B' C'" using assms(10) col__coplanar by auto moreover have "B B' Perp PO A'" by (simp add: S1) moreover have "C C' Perp PO A'" by (simp add: R8) ultimately show ?thesis using l12_9 by auto qed then have Q3: "(B B' ParStrict C C') \ (B \ B' \ C \ C' \ Col B C C' \ Col B' C C')" by (simp add: Par_def) { assume "B B' ParStrict C C'" then have "B B' OS C C'" using l12_6 by auto then have "B B' TS C' A'" using \B B' TS A' C\ l9_2 l9_8_2 by blast then obtain T where Q4: "Col T B B' \ Bet C' T A'" using TS_def by blast have "T = B'" proof - have "\ Col B B' A'" using \B B' OS A A'\ col124__nos by auto moreover have "A' \ C'" using \B B' TS C' A'\ not_two_sides_id by auto moreover have "Col B B' T" using Col_cases Q4 by auto moreover have "Col B B' B'" using not_col_distincts by blast moreover have "Col A' C' T" by (simp add: Col_def Q4) ultimately show ?thesis by (meson assms(9) col_permutation_5 l6_21) qed then have "Bet A' B' C'" using Q4 between_symmetry by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A' B' C'" using TS_def \B B' TS A C\ l6_16_1 not_col_permutation_2 by blast } then show ?thesis using Q3 \B B' ParStrict C C' \ Bet A' B' C'\ by blast qed } { assume R8: "A \ A' \ B \ B' \ Col A B B' \ Col A' B B'" have "A' A Perp PO A'" by (simp add: R7 perp_left_comm) have "\ Col A' A PO" using Col_cases R8 assms(3) assms(6) l8_9 by blast then have "Bet A' B' C'" using Col_perm P1 R8 assms(1) l6_16_1 by blast } then show ?thesis using Par_def \A A' Par B B'\ \A A' ParStrict B B' \ Bet A' B' C'\ by auto qed qed qed qed qed qed lemma per3_preserves_bet2_aux: assumes "Col PO A C" and "A \ C'" and "Bet A B' C'" and "PO \ A" and "PO \ B'" and "PO \ C'" and "Per PO B' B" and "Per PO C' C" and "Col A B C" and "Col PO A C'" shows "Bet A B C" proof cases assume "A = B" then show ?thesis by (simp add: between_trivial2) next assume P1: "A \ B" show ?thesis proof cases assume "B = C" then show ?thesis by (simp add: between_trivial) next assume P2: "B \ C" have P3: "Col PO A B'" by (metis Col_def assms(10) assms(2) assms(3) l6_16_1) then have P4: "Col PO B' C'" using assms(10) assms(4) col_transitivity_1 by blast show ?thesis proof cases assume "B = B'" thus ?thesis by (metis Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(1) assms(10) assms(3) assms(4) assms(6) assms(8) col_transitivity_1) next assume P5: "B \ B'" have P6: "C = C'" using assms(1) assms(10) assms(4) assms(6) assms(8) col_transitivity_1 l8_9 by blast then have "False" by (metis P3 P5 P6 Tarski_neutral_dimensionless.per_col_eq Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(4) assms(5) assms(7) assms(9) col_transitivity_1 l6_16_1 not_col_permutation_4) then show ?thesis by blast qed qed qed lemma per3_preserves_bet2: assumes "Col PO A C" and "A' \ C'" and "Bet A' B' C'" and "PO \ A'" and "PO \ B'" and "PO \ C'" and "Per PO A' A" and "Per PO B' B" and "Per PO C' C" and "Col A B C" and "Col PO A' C'" shows "Bet A B C" proof cases assume "A = A'" then show ?thesis using assms(1) assms(10) assms(11) assms(2) assms(3) assms(4) assms(5) assms(6) assms(8) assms(9) per3_preserves_bet2_aux by blast next assume P1: "A \ A'" show ?thesis proof cases assume "C = C'" thus ?thesis by (metis P1 assms(1) assms(11) assms(4) assms(6) assms(7) col_trivial_3 l6_21 l8_9 not_col_permutation_2) next assume P2: "C \ C'" then have P3: "PO A' Perp C C'" by (metis assms(11) assms(4) assms(6) assms(9) col_per_perp l8_2 not_col_permutation_1) have P4: "PO A' Perp A A'" using P1 assms(4) assms(7) per_perp perp_right_comm by auto have "A A' Par C C'" proof - have "Coplanar PO A' A C" using assms(1) ncop__ncols by blast moreover have "Coplanar PO A' A C'" by (meson assms(11) ncop__ncols) moreover have "Coplanar PO A' A' C" using ncop_distincts by blast moreover have "Coplanar PO A' A' C'" using ncop_distincts by blast moreover have "A A' Perp PO A'" using P4 Perp_cases by blast moreover have "C C' Perp PO A'" using P3 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume P5: "A A' ParStrict C C'" then have P6: "A A' OS C C'" by (simp add: l12_6) have P7: "C C' OS A A'" by (simp add: P5 pars__os3412) have "Bet A B C" proof cases assume P8: "B = B'" then have "A' A OS B C'" by (metis P6 assms(10) assms(3) bet_out col123__nos col124__nos invert_one_side out_one_side) then have "A A' OS B C'" by (simp add: invert_one_side) then have "A A' OS B C" using P6 one_side_symmetry one_side_transitivity by blast then have P12: "A Out B C" using assms(10) col_one_side_out by blast have "C' C OS B A'" by (metis Col_perm P5 P7 P8 assms(10) assms(3) bet_out_1 col123__nos out_one_side par_strict_not_col_2) then have "C C' OS B A" by (meson P7 invert_one_side one_side_symmetry one_side_transitivity) then have "C C' OS A B" using one_side_symmetry by blast then have "C Out A B" using assms(10) col_one_side_out col_permutation_2 by blast then show ?thesis by (simp add: P12 out2__bet) next assume T1: "B \ B'" have T2: "PO A' Perp B B'" proof - have "Per PO B' B" by (simp add: assms(8)) then have "B' PerpAt PO B' B' B" using T1 assms(5) per_perp_in by auto then have "B' PerpAt B' PO B B'" by (simp add: perp_in_comm) then have T4: "B' PO Perp B B' \ B' B' Perp B B'" using Perp_def by auto { assume T5: "B' PO Perp B B'" have "Col A' B' C'" by (simp add: assms(3) bet_col) then have "Col PO B' A'" using assms(11) assms(2) col2__eq col_permutation_4 col_permutation_5 by blast then have "PO A' Perp B B'" by (metis T5 assms(4) col_trivial_3 perp_col2 perp_comm) } { assume "B' B' Perp B B'" then have "PO A' Perp B B'" using perp_distinct by auto } then show ?thesis using T4 \B' PO Perp B B' \ PO A' Perp B B'\ by linarith qed have T6: "B B' Par A A'" proof - have "Coplanar PO A' B A" by (metis Col_cases P7 assms(1) assms(10) col_transitivity_2 ncop__ncols os_distincts) moreover have "Coplanar PO A' B A'" using ncop_distincts by blast moreover have "Coplanar PO A' B' A" proof - have "(Bet PO A' C' \ Bet PO C' A') \ Bet C' PO A'" by (meson assms(11) third_point) then show ?thesis by (meson Bet_perm assms(3) bet__coplanar between_exchange2 l5_3 ncoplanar_perm_8) qed moreover have "Coplanar PO A' B' A'" using ncop_distincts by auto moreover have "B B' Perp PO A'" using Perp_cases T2 by blast moreover have "A A' Perp PO A'" using P4 Perp_cases by blast ultimately show ?thesis using l12_9 by blast qed { assume "B B' ParStrict A A'" then have "B B' OS A A'" by (simp add: l12_6) have "B B' Par C C'" proof - have "Coplanar PO A' B C" by (metis Col_cases P7 assms(1) assms(10) col2__eq ncop__ncols os_distincts) moreover have "Coplanar PO A' B C'" using assms(11) ncop__ncols by auto moreover have "Coplanar PO A' B' C" by (metis Out_def assms(11) assms(2) assms(3) col_trivial_2 l6_16_1 ncop__ncols not_col_permutation_1 out_col) moreover have "Coplanar PO A' B' C'" using assms(11) ncop__ncols by blast moreover have "B B' Perp PO A'" using Perp_cases T2 by blast moreover have "C C' Perp PO A'" using P3 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume T9: "B B' ParStrict C C'" then have T10: "B B' OS C C'" by (simp add: l12_6) have T11: "B B' TS A' C'" by (metis Col_cases T10 \B B' ParStrict A A'\ assms(3) bet__ts invert_two_sides os_distincts par_strict_not_col_4) have T12: "B B' TS A C'" using \B B' OS A A'\ \B B' TS A' C'\ l9_8_2 one_side_symmetry by blast then have T12A: "B B' TS C A" using T10 l9_2 l9_8_2 one_side_symmetry by blast then obtain T where T13: "Col T B B' \ Bet C T A" using TS_def by auto then have "B = T" by (metis Col_perm TS_def T12A assms(10) bet_col1 col_transitivity_2 col_two_sides_bet) then have "Bet A B C" using Bet_perm T13 by blast } { assume "B \ B' \ C \ C' \ Col B C C' \ Col B' C C'" then have "Bet A B C" by (metis Col_cases P5 assms(10) col3 col_trivial_2 not_bet_distincts par_strict_not_col_3) } then have "Bet A B C" using Par_def \B B' Par C C'\ \B B' ParStrict C C' \ Bet A B C\ by auto } { assume "B \ B' \ A \ A' \ Col B A A' \ Col B' A A'" then have "Bet A B C" by (smt P6 assms(10) col123__nos l6_16_1 not_bet_distincts not_col_permutation_1) } then show ?thesis using Par_def T6 \B B' ParStrict A A' \ Bet A B C\ by auto qed } { assume "A \ A' \ C \ C' \ Col A C C' \ Col A' C C'" then have "Bet A B C" by (metis Col_perm P3 Par_def assms(11) assms(2) assms(4) col_transitivity_1 perp_not_par) } thus ?thesis using Par_def \A A' Par C C'\ \A A' ParStrict C C' \ Bet A B C\ by auto qed qed lemma symmetry_preserves_per: assumes "Per B P A" and "B Midpoint A A'" and "B Midpoint P P'" shows "Per B P' A'" proof - obtain C where P1: "P Midpoint A C" using symmetric_point_construction by blast obtain C' where P2: "B Midpoint C C'" using symmetric_point_construction by blast have P3: "P' Midpoint A' C'" using P1 P2 assms(2) assms(3) symmetry_preserves_midpoint by blast have "Cong B A' B C'" by (meson P1 P2 assms(1) assms(2) l7_16 l7_3_2 per_double_cong) then show ?thesis using P3 Per_def by blast qed lemma l13_1_aux: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" and "R Midpoint A B" shows "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" proof - have P1: "Q \ C" using assms(1) assms(3) midpoint_not_midpoint not_col_distincts by blast have P2: "P \ C" using assms(1) assms(2) is_midpoint_id_2 not_col_distincts by blast then have "Q \ R" using assms(2) assms(3) assms(4) l7_3 symmetric_point_uniqueness by blast have "R \ B" using assms(1) assms(4) midpoint_not_midpoint not_col_distincts by blast { assume V1: "Col P Q C" have V2: "Col B P C" by (simp add: assms(2) bet_col midpoint_bet) have V3: "Col A Q C" by (simp add: assms(3) bet_col midpoint_bet) have "Col A R B" using assms(4) midpoint_col not_col_permutation_4 by blast then have "Col A B C" using V1 V2 V3 by (metis P1 P2 col2__eq col_permutation_5) then have "False" using assms(1) by auto } then have P2A: "\ Col P Q C" by auto then obtain C' where P3: "Col P Q C' \ P Q Perp C C'" using l8_18_existence by blast obtain A' where P4: "Q Midpoint C' A'" using symmetric_point_construction by auto obtain B' where P5: "P Midpoint C' B'" using symmetric_point_construction by auto have P6: "Cong C C' B B'" using Mid_cases P5 assms(2) l7_13 by blast have P7: "Cong C C' A A'" using P4 assms(3) l7_13 l7_2 by blast have P8: "Per P B' B" proof cases assume "P = C'" then show ?thesis using P5 Per_cases is_midpoint_id l8_5 by blast next assume "P \ C'" then have "P C' Perp C C'" using P3 perp_col by blast then have "Per P C' C" using Perp_perm perp_per_2 by blast then show ?thesis using symmetry_preserves_per Mid_perm P5 assms(2) by blast qed have P8A: "Per Q A' A" proof cases assume "Q = C'" then show ?thesis using P4 Per_cases is_midpoint_id l8_5 by blast next assume "Q \ C'" then have "C' Q Perp C C'" using P3 col_trivial_2 perp_col2 by auto then have "Per Q C' C" by (simp add: perp_per_1) then show ?thesis by (meson Mid_cases P4 assms(3) l7_3_2 midpoint_preserves_per) qed have P9: "Col A' C' Q" using P4 midpoint_col not_col_permutation_3 by blast have P10: "Col B' C' P" using P5 midpoint_col not_col_permutation_3 by blast have P11: "P \ Q" using P2A col_trivial_1 by auto then have P12: "A' \ B'" using P4 P5 l7_17 by blast have P13: "Col A' B' P" by (metis P10 P3 P4 P5 P9 col2__eq col_permutation_5 midpoint_distinct_1 not_col_distincts) have P14: "Col A' B' Q" by (smt P10 P3 P4 P5 P9 col3 col_permutation_1 midpoint_distinct_1 not_col_distincts) have P15: "Col A' B' C'" using P11 P13 P14 P3 colx by blast have P16: "C \ C'" using P2A P3 by blast then have P17: "A \ A'" using P7 cong_diff by blast have P18: "B \ B'" using P16 P6 cong_diff by blast have P19: "Per P A' A" proof cases assume P20: "A' = Q" then have "A' P Perp C A'" by (metis P3 P4 Perp_cases midpoint_not_midpoint) then have "Per P A' C" by (simp add: perp_per_1) then show ?thesis using P20 assms(3) l7_2 l8_4 by blast next assume "A' \ Q" then show ?thesis by (meson P12 P13 P14 P8A col_transitivity_1 l8_2 per_col) qed have "Per Q B' B" proof cases assume P21: "P = B'" then have P22: "C' = B'" using P5 is_midpoint_id_2 by auto then have "Per Q B' C" using P3 P21 perp_per_1 by auto thus ?thesis by (metis Col_perm P16 P21 P22 assms(2) midpoint_col per_col) next assume P23: "P \ B'" have "Col B' P Q" using P12 P13 P14 col_transitivity_2 by blast then have "Per B B' Q" using P8 P23 l8_2 l8_3 by blast thus ?thesis using Per_perm by blast qed then have P24: "Per A' B' B" using P11 P13 P14 P8 l8_3 not_col_permutation_2 by blast have P25: "Per A A' B'" using P11 P13 P14 P19 P8A l8_2 l8_3 not_col_permutation_5 by blast then have "Per B' A' A" using Per_perm by blast then have "\ Col B' A' A" using P12 P17 P25 per_not_col by auto then have P26: "\ Col A' B' A" using Col_cases by auto have "\ Col A' B' B" using P12 P18 P24 l8_9 by auto obtain X where P28: "X Midpoint A' B'" using midpoint_existence by blast then have P28A: "Col A' B' X" using midpoint_col not_col_permutation_2 by blast then have "\ Q. A' B' Perp Q X \ A' B' OS A Q" by (simp add: P26 l10_15) then obtain y where P29: "A' B' Perp y X \ A' B' OS A y" by blast then obtain B'' where P30: "(X y Perp A B'' \ A = B'') \ (\ M. (Col X y M \ M Midpoint A B''))" using ex_sym by blast then have P31: "B'' A ReflectL X y" using P30 ReflectL_def by blast have P32: "X \ y" using P29 P28A col124__nos by blast then have "X \ y \ B'' A ReflectL X y \ X = y \ X Midpoint A B''" using P31 by auto then have P33: "B'' A Reflect X y" by (simp add: Reflect_def) have P33A: "X \ y \ A' B' ReflectL X y" using P28 P29 Perp_cases ReflectL_def P32 col_trivial_3 l10_4_spec by blast then have P34: "A' B' Reflect X y" using Reflect_def by blast have P34A: "A B'' Reflect X y" using P33 l10_4 by blast then have P35: "Cong B'' B' A A'" using P34 l10_10 by auto have "Per A' B' B''" proof - have R1: "X \ y \ A B'' ReflectL X y \ X = y \ X Midpoint B'' A" by (simp add: P31 P32 l10_4_spec) have R2: "X \ y \ A' B' ReflectL X y \ X = y \ X Midpoint B' A'" using P33A by linarith { assume "X \ y \ A B'' ReflectL X y \ X \ y \ A' B' ReflectL X y" then have "Per A' B' B''" using \Per B' A' A\ image_spec_preserves_per l10_4_spec by blast } { assume "X \ y \ A B'' ReflectL X y \ X = y \ X Midpoint B' A'" then have "Per A' B' B''" by blast } { assume "X = y \ X Midpoint B'' A \ X \ y \ A' B' ReflectL X y" then have "Per A' B' B''" by blast } { assume "X = y \ X Midpoint B'' A \ X = y \ X Midpoint B' A'" then have "Per A' B' B''" using P32 by blast } then show ?thesis using R1 R2 using \X \ y \ A B'' ReflectL X y \ X \ y \ A' B' ReflectL X y \ Per A' B' B''\ by auto qed have "A' B' OS A B''" proof - { assume S1: "X y Perp A B''" have "Coplanar A y A' X" by (metis P28A P29 col_one_side coplanar_perm_16 ncop_distincts os__coplanar) have "Coplanar A y B' X" by (smt P12 P28A P29 col2_cop__cop col_transitivity_1 ncoplanar_perm_22 not_col_permutation_5 os__coplanar) have S2: "\ Col A X y" using Col_perm P34A S1 local.image_id perp_distinct by blast have "A' B' Par A B''" proof - have "Coplanar X y A' A" using \Coplanar A y A' X\ ncoplanar_perm_21 by blast moreover have "Coplanar X y A' B''" proof - have "Coplanar A X y A'" using \Coplanar X y A' A\ ncoplanar_perm_9 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using S2 coplanar_trans_1 by auto qed moreover have "Coplanar X y B' A" proof - have "\ Col A X y" by (simp add: S2) moreover have "Coplanar A X y B'" using \Coplanar A y B' X\ ncoplanar_perm_3 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using ncoplanar_perm_18 by blast qed moreover have "Coplanar X y B' B''" proof - have "\ Col A X y" by (simp add: S2) moreover have "Coplanar A X y B'" using \Coplanar X y B' A\ ncoplanar_perm_9 by blast moreover have "Coplanar A X y B''" using Coplanar_def S1 perp_inter_exists by blast ultimately show ?thesis using coplanar_trans_1 by blast qed ultimately show ?thesis using l12_9 using P29 Perp_cases S1 by blast qed have "A' B' OS A B''" proof - { assume "A' B' ParStrict A B''" have "A' B' OS A B''" using l12_6 using \A' B' ParStrict A B''\ by blast } { assume "A' \ B' \ A \ B'' \ Col A' A B'' \ Col B' A B''" have "A' B' OS A B''" using P26 \A' B' Par A B''\ \A' B' ParStrict A B'' \ A' B' OS A B''\ col_trivial_3 par_not_col_strict by blast } then show ?thesis using Par_def \A' B' Par A B''\ \A' B' ParStrict A B'' \ A' B' OS A B''\ by auto qed } { assume "A = B''" then have "A' B' OS A B''" using P12 P25 \Per A' B' B''\ l8_2 l8_7 by blast } then show ?thesis using P30 \X y Perp A B'' \ A' B' OS A B''\ by blast qed have "A' B' OS A B" proof - have "A' B' TS A C" proof - have "\ Col A A' B'" using Col_perm \\ Col B' A' A\ by blast moreover have "\ Col C A' B'" by (metis P13 P14 P2A \\ Col B' A' A\ col3 not_col_distincts not_col_permutation_3 not_col_permutation_4) moreover have "\ T. Col T A' B' \ Bet A T C" using P14 assms(3) midpoint_bet not_col_permutation_1 by blast ultimately show ?thesis by (simp add: TS_def) qed moreover have "A' B' TS B C" by (metis Col_cases P13 TS_def \\ Col A' B' B\ assms(2) calculation midpoint_bet) ultimately show ?thesis using OS_def by blast qed have "Col B B'' B'" proof - have "Coplanar A' B B'' B'" proof - have "Coplanar A' B' B B''" proof - have "\ Col A A' B'" using Col_perm \\ Col B' A' A\ by blast moreover have "Coplanar A A' B' B" using \A' B' OS A B\ ncoplanar_perm_8 os__coplanar by blast moreover have "Coplanar A A' B' B''" using \A' B' OS A B''\ ncoplanar_perm_8 os__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by blast qed then show ?thesis using ncoplanar_perm_4 by blast qed moreover have "A' \ B'" by (simp add: P12) moreover have "Per B B' A'" by (simp add: P24 l8_2) moreover have "Per B'' B' A'" using Per_cases \Per A' B' B''\ by auto ultimately show ?thesis using cop_per2__col by blast qed have "Cong B B' A A'" using P6 P7 cong_inner_transitivity by blast have "B = B'' \ B' Midpoint B B''" proof - have "Col B B' B''" using \Col B B'' B'\ not_col_permutation_5 by blast moreover have "Cong B' B B' B''" by (metis Cong_perm P35 P6 P7 cong_inner_transitivity) ultimately show ?thesis using l7_20 by simp qed { assume "B = B''" then obtain M where S1: "Col X y M \ M Midpoint A B" using P30 by blast then have "R = M" using assms(4) l7_17 by auto have "A \ B" using assms(1) col_trivial_1 by auto have "Col R A B" by (simp add: assms(4) midpoint_col) have "X \ R" using Midpoint_def P28 \A' B' OS A B''\ \B = B''\ assms(4) midpoint_col one_side_chara by auto then have "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" proof - have "R PerpAt R X A B" proof - have "R X Perp A B" using P30 S1 \A \ B\ \B = B''\ \R = M\ \X \ R\ perp_col perp_left_comm by blast then show ?thesis using \Col R A B\ l8_14_2_1b_bis not_col_distincts by blast qed moreover have "R X Perp P Q" proof - have "X R Perp P Q" proof - have "X y Perp P Q" proof - have "P Q Perp X y" using P11 P13 P14 P29 P33A col_trivial_2 col_trivial_3 perp_col4 by blast then show ?thesis using Perp_perm by blast qed moreover have "Col X y R" by (simp add: S1 \R = M\) ultimately show ?thesis using \X \ R\ perp_col by blast qed then show ?thesis using Perp_perm by blast qed moreover have "Coplanar A B C R" using \Col R A B\ ncop__ncols not_col_permutation_2 by blast moreover have "Coplanar A B C X" proof - have "Col P Q X" using P12 P13 P14 P28A col3 by blast moreover have "\ Col P Q C" by (simp add: P2A) moreover have "Coplanar P Q C A" using assms(3) coplanar_perm_19 midpoint__coplanar by blast moreover have "Coplanar P Q C B" using assms(2) midpoint_col ncop__ncols not_col_permutation_5 by blast moreover have "Coplanar P Q C C" using ncop_distincts by auto moreover have "Coplanar P Q C X" using calculation(1) ncop__ncols by blast ultimately show ?thesis using coplanar_pseudo_trans by blast qed ultimately show ?thesis by blast qed } { assume "B' Midpoint B B''" have "A' B' TS B B''" proof - have "\ Col B A' B'" using Col_perm \\ Col A' B' B\ by blast moreover have "\ Col B'' A' B'" using \A' B' OS A B''\ col124__nos not_col_permutation_2 by blast moreover have "\ T. Col T A' B' \ Bet B T B''" using \B' Midpoint B B''\ col_trivial_3 midpoint_bet by blast ultimately show ?thesis by (simp add: TS_def) qed have "A' B' OS B B''" using \A' B' OS A B''\ \A' B' OS A B\ one_side_symmetry one_side_transitivity by blast have "\ A' B' OS B B''" using \A' B' TS B B''\ l9_9_bis by blast then have "False" by (simp add: \A' B' OS B B''\) then have "\ X Y. (R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y)" by auto } then show ?thesis using \B = B'' \ \X Y. R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y\ \B = B'' \ B' Midpoint B B''\ by blast qed lemma l13_1: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" and "R Midpoint A B" shows "\ X Y.(R PerpAt X Y A B \ X Y Perp P Q)" proof - obtain X Y where "R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y" using l13_1_aux assms(1) assms(2) assms(3) assms(4) by blast then show ?thesis by blast qed lemma per_lt: assumes "A \ B" and "C \ B" and "Per A B C" shows "A B Lt A C \ C B Lt A C" proof - have "B A Lt A C \ B C Lt A C" using assms(1) assms(2) assms(3) l11_46 by auto then show ?thesis using lt_left_comm by blast qed lemma cong_perp_conga: assumes "Cong A B C B" and "A C Perp B P" shows "A B P CongA C B P \ B P TS A C" proof - have P1: "A \ C" using assms(2) perp_distinct by auto have P2: "B \ P" using assms(2) perp_distinct by auto have P3: "A \ B" by (metis P1 assms(1) cong_diff_3) have P4: "C \ B" using P3 assms(1) cong_diff by blast show ?thesis proof cases assume P5: "Col A B C" have P6: "\ Col B A P" using P3 P5 assms(2) col_transitivity_1 not_col_permutation_4 not_col_permutation_5 perp_not_col2 by blast have "Per P B A" using P3 P5 Perp_perm assms(2) not_col_permutation_5 perp_col1 perp_per_1 by blast then have P8: "Per A B P" using Per_cases by blast have "Per P B C" using P3 P5 P8 col_per2__per l8_2 l8_5 by blast then have P10: "Per C B P" using Per_perm by blast show ?thesis proof - have "A B P CongA C B P" using P2 P3 P4 P8 P10 l11_16 by auto moreover have "B P TS A C" by (metis Col_cases P1 P5 P6 assms(1) bet__ts between_cong not_cong_2143 not_cong_4321 third_point) ultimately show ?thesis by simp qed next assume T1: "\ Col A B C" obtain T where T2: "T PerpAt A C B P" using assms(2) perp_inter_perp_in by blast then have T3: "Col A C T \ Col B P T" using perp_in_col by auto have T4: "B \ T" using Col_perm T1 T3 by blast have T5: "B T Perp A C" using Perp_cases T3 T4 assms(2) perp_col1 by blast { assume T5_1: "A = T" have "B A Lt B C \ C A Lt B C" proof - have "B \ A" using P3 by auto moreover have "C \ A" using P1 by auto moreover have "Per B A C" using T5 T5_1 perp_comm perp_per_1 by blast ultimately show ?thesis by (simp add: per_lt) qed then have "False" using Cong_perm assms(1) cong__nlt by blast } then have T6: "A \ T" by auto { assume T6_1: "C = T" have "B C Lt B A \ A C Lt B A" proof - have "B \ C" using P4 by auto moreover have "A \ C" by (simp add: P1) moreover have "Per B C A" using T5 T6_1 perp_left_comm perp_per_1 by blast ultimately show ?thesis by (simp add: per_lt) qed then have "False" using Cong_perm assms(1) cong__nlt by blast } then have T7: "C \ T" by auto have T8: "T PerpAt B T T A" by (metis Perp_in_cases T2 T3 T4 T6 perp_in_col_perp_in) have T9: "T PerpAt B T T C" by (metis Col_cases T3 T7 T8 perp_in_col_perp_in) have T10: "Cong T A T C \ T A B CongA T C B \ T B A CongA T B C" proof - have "A T B CongA C T B" proof - have "Per A T B" using T2 perp_in_per_1 by auto moreover have "Per C T B" using T2 perp_in_per_3 by auto ultimately show ?thesis by (simp add: T4 T6 T7 l11_16) qed moreover have "Cong A B C B" by (simp add: assms(1)) moreover have "Cong T B T B" by (simp add: cong_reflexivity) moreover have "T B Le A B" proof - have "Per B T A" using T8 perp_in_per by auto then have "B T Lt B A \ A T Lt B A" using T4 T6 per_lt by blast then show ?thesis using Le_cases Lt_def by blast qed ultimately show ?thesis using l11_52 by blast qed show ?thesis proof - have T11: "A B P CongA C B P" proof - have "P B A CongA P B C" using Col_cases P2 T10 T3 col_conga__conga by blast thus ?thesis using conga_comm by blast qed moreover have "B P TS A C" proof - have T12: "A = C \ T Midpoint A C" using T10 T3 l7_20_bis not_col_permutation_5 by blast { assume "T Midpoint A C" then have "B P TS A C" by (smt Col_perm P2 T1 T3 \A = T \ False\ \C = T \ False\ col2__eq l9_18 midpoint_bet) } then show ?thesis using P1 T12 by auto qed ultimately show ?thesis by simp qed qed qed lemma perp_per_bet: assumes "\ Col A B C" and (* "Col A P C" and *) "Per A B C" and "P PerpAt P B A C" shows "Bet A P C" proof - have "A \ C" using assms(1) col_trivial_3 by auto then show ?thesis using assms(2) assms(3) l11_47 perp_in_left_comm by blast qed lemma ts_per_per_ts: assumes "A B TS C D" and "Per B C A" and "Per B D A" shows "C D TS A B" proof - have P1: "\ Col C A B" using TS_def assms(1) by blast have P2: "A \ B" using P1 col_trivial_2 by auto obtain P where P3: "Col P A B \ Bet C P D" using TS_def assms(1) by blast have P4: "C \ D" using assms(1) not_two_sides_id by auto show ?thesis proof - { assume "Col A C D" then have "C = D" by (metis assms(1) assms(2) assms(3) col_per2_cases col_permutation_2 not_col_distincts ts_distincts) then have "False" using P4 by auto } then have "\ Col A C D" by auto moreover have "\ Col B C D" using assms(1) assms(2) assms(3) per2_preserves_diff ts_distincts by blast moreover have "\ T. Col T C D \ Bet A T B" proof - have "Col P C D" using Col_def Col_perm P3 by blast moreover have "Bet A P B" proof - have "\ X. Col A B X \ A B Perp C X" using Col_perm P1 l8_18_existence by blast then obtain C' where P5: "Col A B C' \ A B Perp C C'" by blast have "\ X. Col A B X \ A B Perp D X" by (metis (no_types) Col_perm TS_def assms(1) l8_18_existence) then obtain D' where P6: "Col A B D' \ A B Perp D D'" by blast have P7: "A \ C'" using P5 assms(2) l8_7 perp_not_eq_2 perp_per_1 by blast have P8: "A \ D'" using P6 assms(3) l8_7 perp_not_eq_2 perp_per_1 by blast have P9: "Bet A C' B" proof - have "\ Col A C B" using Col_cases P1 by blast moreover have "Per A C B" by (simp add: assms(2) l8_2) moreover have "C' PerpAt C' C A B" using P5 Perp_in_perm l8_15_1 by blast ultimately show ?thesis using perp_per_bet by blast qed have P10: "Bet A D' B" proof - have "\ Col A D B" using P6 col_permutation_5 perp_not_col2 by blast moreover have "Per A D B" by (simp add: assms(3) l8_2) moreover have "D' PerpAt D' D A B" using P6 Perp_in_perm l8_15_1 by blast ultimately show ?thesis using perp_per_bet by blast qed show ?thesis proof cases assume "P = C'" then show ?thesis by (simp add: P9) next assume "P \ C'" show ?thesis proof cases assume "P = D'" then show ?thesis by (simp add: P10) next assume "P \ D'" show ?thesis proof cases assume "A = P" then show ?thesis by (simp add: between_trivial2) next assume "A \ P" show ?thesis proof cases assume "B = P" then show ?thesis using between_trivial by auto next assume "B \ P" have "Bet C' P D'" proof - have "Bet C P D" by (simp add: P3) moreover have "P \ C'" by (simp add: \P \ C'\) moreover have "P \ D'" by (simp add: \P \ D'\) moreover have "Col C' P D'" by (meson P2 P3 P5 P6 col3 col_permutation_2) moreover have "Per P C' C" using P3 P5 l8_16_1 l8_2 not_col_permutation_3 not_col_permutation_4 by blast moreover have "Per P D' D" by (metis P3 P6 calculation(3) not_col_permutation_2 perp_col2 perp_per_1) ultimately show ?thesis using per13_preserves_bet by blast qed then show ?thesis using P10 P9 bet3__bet by blast qed qed qed qed qed ultimately show ?thesis by auto qed ultimately show ?thesis by (simp add: TS_def) qed qed lemma l13_2_1: assumes "A B TS C D" and "Per B C A" and "Per B D A" and "Col C D E" and "A E Perp C D" and "C A B CongA D A B" shows "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof - have P1: "\ Col C A B" using TS_def assms(1) by auto have P2: "A \ C" using P1 col_trivial_1 by blast have P3: "A \ B" using P1 col_trivial_2 by auto have P4: "A \ D" using assms(1) ts_distincts by auto have P5: "Cong B C B D \ Cong A C A D \ C B A CongA D B A" proof - have "\ Col B A C" by (simp add: P1 not_col_permutation_3) moreover have "A C B CongA A D B" using assms(1) assms(2) assms(3) l11_16 l8_2 ts_distincts by blast moreover have "B A C CongA B A D" by (simp add: assms(6) conga_comm) moreover have "Cong B A B A" by (simp add: cong_reflexivity) ultimately show ?thesis using l11_50_2 by blast qed then have P6: "C D Perp A B" using assms(1) assms(6) cong_conga_perp not_cong_2143 by blast then have P7: "C D TS A B" by (simp add: assms(1) assms(2) assms(3) ts_per_per_ts) obtain T1 where P8: "Col T1 C D \ Bet A T1 B" using P7 TS_def by auto obtain T where P9: "Col T A B \ Bet C T D" using TS_def assms(1) by blast have P10: "T1 = T" by (metis (no_types) Col_def P1 P3 P8 P9 between_equality_2 between_trivial2 l6_16_1) have P11: "T = E" proof - have "\ Col A B C" using Col_perm P1 by blast moreover have "C \ D" using assms(1) ts_distincts by blast moreover have "Col A B T" using Col_cases P9 by auto moreover have "Col A B E" by (metis P7 Perp_cases P6 assms(1) assms(5) col_perp2_ncol_col col_trivial_3 not_col_permutation_3 one_side_not_col123 os_ts1324__os ts_ts_os) moreover have "Col C D T" using NCol_cases P9 bet_col by blast moreover have "Col C D E" by (simp add: assms(4)) ultimately show ?thesis using l6_21 by blast qed show ?thesis proof - have "B A C CongA D A E" proof - have "A Out C C" using P2 out_trivial by auto moreover have "A Out B B" using P3 out_trivial by auto moreover have "A Out D D" using P4 out_trivial by auto moreover have "A Out E B" by (metis P10 P11 P7 P8 TS_def bet_out) ultimately show ?thesis by (meson assms(6) conga_comm conga_right_comm l11_10) qed moreover have "B A D CongA C A E" proof - have "C A E CongA D A B" by (meson Perp_cases P5 assms(5) assms(6) calculation cong_perp_conga conga_right_comm conga_trans not_cong_2143 not_conga_sym) then have "C A E CongA B A D" by (simp add: conga_right_comm) then show ?thesis by (simp add: conga_sym) qed moreover have "Bet C E D" using P11 P9 by auto ultimately show ?thesis by simp qed qed lemma triangle_mid_par: assumes "\ Col A B C" and "P Midpoint B C" and "Q Midpoint A C" shows "A B ParStrict Q P" proof - obtain R where P1: "R Midpoint A B" using midpoint_existence by auto then obtain X Y where P2: "R PerpAt X Y A B \ X Y Perp P Q \ Coplanar A B C X \ Coplanar A B C Y" using l13_1_aux assms(1) assms(2) assms(3) by blast have P3: "Coplanar X Y A P \ Coplanar X Y A Q \ Coplanar X Y B P \ Coplanar X Y B Q" proof - have "Coplanar A B C A" using ncop_distincts by auto moreover have "Coplanar A B C B" using ncop_distincts by auto moreover have "Coplanar A B C P" using assms(2) coplanar_perm_21 midpoint__coplanar by blast moreover have "Coplanar A B C Q" using assms(3) coplanar_perm_11 midpoint__coplanar by blast ultimately show ?thesis using P2 assms(1) coplanar_pseudo_trans by blast qed have P4: "Col X Y R \ Col A B R" using P2 perp_in_col by blast have P5: "R Y Perp A B \ X R Perp A B" using P2 perp_in_perp_bis by auto have P6: "Col A R B" using Col_perm P4 by blast have P7: "X \ Y" using P2 perp_not_eq_1 by auto { assume P8: "R Y Perp A B" have "Col Y R X" using P4 not_col_permutation_2 by blast then have "Y X Perp A B" using P2 Perp_cases perp_in_perp by blast then have P10: "X Y Perp A B" using Perp_cases by blast have "A B Par P Q" proof - have "Coplanar X Y A P" by (simp add: P3) moreover have "Coplanar X Y A Q" by (simp add: P3) moreover have "Coplanar X Y B P" by (simp add: P3) moreover have "Coplanar X Y B Q" by (simp add: P3) moreover have "A B Perp X Y" using P10 Perp_cases by auto moreover have "P Q Perp X Y" using P2 Perp_cases by auto ultimately show ?thesis using l12_9 by blast qed { assume "A B ParStrict P Q" then have "A B ParStrict Q P" using Par_strict_perm by blast } { assume "A \ B \ P \ Q \ Col A P Q \ Col B P Q" then have "Col A B P" using l6_16_1 not_col_permutation_1 by blast then have "P = B" by (metis Col_perm assms(1) assms(2) l6_16_1 midpoint_col) then have "A B ParStrict Q P" using assms(1) assms(2) col_trivial_2 is_midpoint_id by blast } then have "A B ParStrict Q P" using Par_def \A B Par P Q\ \A B ParStrict P Q \ A B ParStrict Q P\ by auto } { assume P10: "X R Perp A B" have "Col X R Y" by (simp add: Col_perm P4) then have P11: "X Y Perp A B" using P7 P10 perp_col by blast have "A B Par P Q" proof - have "A B Perp X Y" using P11 Perp_perm by blast moreover have "P Q Perp X Y" using P2 Perp_perm by blast ultimately show ?thesis using P3 l12_9 by blast qed { assume "A B ParStrict P Q" then have "A B ParStrict Q P" by (simp add: par_strict_right_comm) } { assume "A \ B \ P \ Q \ Col A P Q \ Col B P Q" then have "Col A B P" using Col_perm l6_16_1 by blast then have "P = B" by (metis Col_perm assms(1) assms(2) l6_16_1 midpoint_col) then have "A B ParStrict Q P" using assms(1) assms(2) col_trivial_2 is_midpoint_id by blast } then have "A B ParStrict Q P" using Par_def \A B Par P Q\ \A B ParStrict P Q \ A B ParStrict Q P\ by auto } then show ?thesis using P5 \R Y Perp A B \ A B ParStrict Q P\ by blast qed lemma cop4_perp_in2__col: assumes "Coplanar X Y A A'" and "Coplanar X Y A B'" and "Coplanar X Y B A'" and "Coplanar X Y B B'" and "P PerpAt A B X Y" and "P PerpAt A' B' X Y" shows "Col A B A'" proof - have P1: "Col A B P \ Col X Y P" using assms(5) perp_in_col by auto show ?thesis proof cases assume P2: "A = P" show ?thesis proof cases assume P3: "P = X" have "Col B A' P" proof - have "Coplanar Y B A' P" using P3 assms(3) ncoplanar_perm_18 by blast moreover have "Y \ P" using P3 assms(6) perp_in_distinct by blast moreover have "Per B P Y" using assms(5) perp_in_per_4 by auto moreover have "Per A' P Y" using assms(6) perp_in_per_2 by auto ultimately show ?thesis using cop_per2__col by auto qed then show ?thesis using Col_perm P2 by blast next assume P4: "P \ X" have "Col B A' P" proof - have "Coplanar X B A' P" by (metis P1 assms(3) assms(6) col2_cop__cop col_trivial_3 ncoplanar_perm_9 perp_in_distinct) moreover have "Per B P X" using assms(5) perp_in_per_3 by auto moreover have "Per A' P X" using assms(6) perp_in_per_1 by auto ultimately show ?thesis using cop_per2__col P4 by auto qed then show ?thesis using Col_perm P2 by blast qed next assume P5: "A \ P" have P6: "Per A P Y" using assms(5) perp_in_per_2 by auto show ?thesis proof cases assume P7: "P = A'" have P8: "Per B' P Y" using assms(6) perp_in_per_4 by auto have "Col A B' P" proof - have "Coplanar Y A B' P" using assms(2) by (metis P1 assms(6) col_transitivity_2 coplanar_trans_1 ncop__ncols perp_in_distinct) then show ?thesis using P6 P8 cop_per2__col by (metis assms(2) assms(5) assms(6) col_permutation_4 coplanar_perm_5 perp_in_distinct perp_in_per_1 perp_in_per_3) qed then show ?thesis using P1 P7 by auto next assume T1: "P \ A'" show ?thesis proof cases assume T2: "Y = P" { assume R1: "Coplanar X P A A' \ P PerpAt A B X P \ P PerpAt A' B' X P \ A \ P" then have R2: "Per A P X" using perp_in_per_1 by auto have "Per A' P X" using R1 perp_in_per_1 by auto then have "Col A B A'" by (metis R1 R2 PerpAt_def col_permutation_3 col_transitivity_2 cop_per2__col ncoplanar_perm_5) } then show ?thesis using P5 T1 T2 assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) by blast next assume P10: "Y \ P" have "Col A A' P" proof - have "Coplanar Y A A' P" by (metis P1 assms(1) assms(6) col2_cop__cop col_trivial_2 ncoplanar_perm_9 perp_in_distinct) moreover have "Per A P Y" by (simp add: P6) moreover have "Per A' P Y" using assms(6) perp_in_per_2 by auto ultimately show ?thesis using cop_per2__col P10 by auto qed then show ?thesis using P1 P5 col2__eq col_permutation_4 by blast qed qed qed qed lemma l13_2: assumes "A B TS C D" and "Per B C A" and "Per B D A" and "Col C D E" and "A E Perp C D" shows "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof - have P2: "\ Col C A B" using TS_def assms(1) by auto have P3: "C \ D" using assms(1) not_two_sides_id by blast have P4: "\ C'. B A C CongA D A C' \ D A OS C' B" proof - have "\ Col B A C" using Col_cases P2 by auto moreover have "\ Col D A B" using TS_def assms(1) by blast ultimately show ?thesis by (simp add: angle_construction_1) qed then obtain E' where P5: "B A C CongA D A E' \ D A OS E' B" by blast have P6: "A \ B" using P2 not_col_distincts by blast have P7: "A \ C" using P2 not_col_distincts by blast have P8: "A \ D" using P5 os_distincts by blast have P9: "((A B TS C E' \ A E' TS D B) \ (A B OS C E' \ A E' OS D B \ C A B CongA D A E' \ B A E' CongA E' A B)) \ C A E' CongA D A B" by (metis P5 P6 conga_diff56 conga_left_comm conga_pseudo_refl l11_22) have P10: "C D TS A B" by (simp add: assms(1) assms(2) assms(3) ts_per_per_ts) have P11: "\ Col A C D" using P10 TS_def by auto obtain T where P12: "Col T A B \ Bet C T D" using TS_def assms(1) by blast obtain T2 where P13: "Col T2 C D \ Bet A T2 B" using P10 TS_def by auto then have P14: "T = T2" by (metis Col_def Col_perm P12 P2 P3 P6 l6_16_1) have P15: "B InAngle D A C" using P10 assms(1) l11_24 ts2__inangle by blast have P16: "C A B LeA C A D" by (simp add: P10 assms(1) inangle__lea ts2__inangle) have P17: "E' InAngle D A C" proof - have "D A E' LeA D A C" using P16 P5 P7 P8 conga_left_comm conga_pseudo_refl l11_30 by presburger moreover have "D A OS C E'" by (meson P11 P15 P5 col124__nos in_angle_one_side invert_one_side not_col_permutation_2 one_side_symmetry one_side_transitivity) ultimately show ?thesis by (simp add: lea_in_angle) qed obtain E'' where P18: "Bet D E'' C \ (E'' = A \ A Out E'' E')" using InAngle_def P17 by auto { assume "E'' = A" then have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" using Col_def P11 P18 by auto } { assume P19: "A Out E'' E'" then have P20: "B A C CongA D A E''" by (meson OS_def P5 Tarski_neutral_dimensionless.out2__conga Tarski_neutral_dimensionless_axioms col_one_side_out col_trivial_2 l9_18_R1 not_conga one_side_reflexivity) have P21: "A \ T" using P11 P13 P14 by auto have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" proof cases assume P22: "E'' = T" have P23: "C A B CongA D A B" proof - have "C A B CongA D A T" using P22 P20 conga_left_comm by blast moreover have "A Out C C" using P7 out_trivial by presburger moreover have "A Out B B" using P6 out_trivial by auto moreover have "A Out D D" using P8 out_trivial by auto moreover have "A Out B T" using Out_def P13 P14 P6 P21 by blast ultimately show ?thesis using l11_10 by blast qed then show ?thesis using assms(1) assms(2) assms(3) assms(4) assms(5) l13_2_1 by blast next assume P23A: "E'' \ T" have P24: "D \ E''" using P2 P20 col_trivial_3 ncol_conga_ncol not_col_permutation_3 by blast { assume P24A: "C = E''" have P24B: "C A OS B D" by (meson P10 assms(1) invert_one_side ts_ts_os) have P24C: "A Out B D" proof - have "C A B CongA C A D" using P20 P24A conga_comm by blast moreover have "C A OS B D" by (simp add: P24B) ultimately show ?thesis using conga_os__out by blast qed then have "False" using Col_def P5 one_side_not_col124 out_col by blast } then have P25: "C \ E''" by auto have P26: "A \ E''" using P19 out_diff1 by auto { assume "Col E'' A B" then have "E'' = T" by (smt P13 P14 P18 P2 P3 bet_col l6_21 not_col_permutation_2 not_col_permutation_3) then have "False" using P23A by auto } then have P27: "\ Col E'' A B" by auto have "(A B TS C E'' \ A E'' TS D B) \ (A B OS C E'' \ A E'' OS D B \ C A B CongA D A E'' \ B A E'' CongA E'' A B)" proof cases assume P27_0: "A B OS C E''" have "A E'' OS D B" proof - have P27_1: "A E'' TS D C" by (metis Col_def P10 P18 P24 TS_def P25 bet__ts invert_two_sides l6_16_1) moreover have "A E'' TS B C" proof - have "A E'' TS T C" proof - have "\ Col T A E''" by (metis NCol_cases P13 P14 P21 P27 bet_col col3 col_trivial_2) moreover have "\ Col C A E''" using P27_1 TS_def by auto moreover have "\ T0. (Col T0 A E'' \ Bet T T0 C)" by (meson P12 P18 P27_0 between_symmetry col_trivial_3 l5_3 one_side_chara) ultimately show ?thesis by (simp add: TS_def) qed moreover have "A Out T B" using Out_def P13 P14 P21 P6 by auto ultimately show ?thesis using col_trivial_1 l9_5 by blast qed ultimately show ?thesis using OS_def by auto qed thus ?thesis using P20 P27_0 conga_distinct conga_left_comm conga_pseudo_refl by blast next assume P27_2: "\ A B OS C E''" show ?thesis proof - have P27_3: "A B TS C E''" using P18 P2 P27_2 P27 assms(1) bet_cop__cop between_symmetry cop_nos__ts ts__coplanar by blast moreover have "A E'' TS D B" proof - have P27_3: "A B OS D E''" using P18 bet_ts__os between_symmetry calculation one_side_symmetry by blast have P27_4: "A E'' TS T D" proof - have "\ Col T A E''" by (metis NCol_cases P13 P14 P21 P27 bet_col col3 col_trivial_2) moreover have "\ Col D A E''" by (smt Col_def P11 P18 P24 P27_3 bet3__bet bet_col1 col3 col_permutation_5 col_two_sides_bet l5_1) moreover have "\ T0. (Col T0 A E'' \ Bet T T0 D)" by (meson Bet_perm P12 P18 P27_3 bet_col1 bet_out__bet between_exchange3 col_trivial_3 not_bet_out one_side_chara) ultimately show ?thesis by (simp add: TS_def) qed have "A E'' TS B D" proof - have "A E'' TS T D" using P27_4 by simp moreover have "Col A A E''" using col_trivial_1 by auto moreover have "A Out T B" using P13 P14 P21 bet_out by auto ultimately show ?thesis using l9_5 by blast qed thus ?thesis by (simp add: l9_2) qed ultimately show ?thesis by simp qed qed then have P28: "C A E'' CongA D A B" using l11_22 by (metis P20 P26 P6 conga_left_comm conga_pseudo_refl) obtain C' where P29: "Bet B C C' \ Cong C C' B C" using segment_construction by blast obtain D' where P30: "Bet B D D' \ Cong D D' B D" using segment_construction by blast have P31: "B A D Cong3 D' A D" proof - have "Per A D B" by (simp add: assms(3) l8_2) then obtain D'' where P31_2: "D Midpoint B D'' \ Cong A B A D''" using Per_def by auto have "D Midpoint B D'" using Cong_perm Midpoint_def P30 by blast then have "D' = D''" using P31_2 symmetric_point_uniqueness by auto thus ?thesis using Cong3_def Cong_perm P30 P31_2 cong_reflexivity by blast qed then have P32: "B A D CongA D' A D" using P6 P8 cong3_conga by auto have "B A C Cong3 C' A C" proof - obtain C'' where P33_1: "C Midpoint B C'' \ Cong A B A C''" using Per_def assms(2) l8_2 by blast have "C Midpoint B C'" using Cong_perm Midpoint_def P29 by blast then have "C' = C''" using P33_1 symmetric_point_uniqueness by auto thus ?thesis using Cong3_def Cong_perm P29 P33_1 cong_reflexivity by blast qed then have P34: "B A C CongA C' A C" using P6 P7 cong3_conga by auto have P35: "E'' A C' CongA D' A E''" proof - have "(A C TS E'' C' \ A D TS D' E'') \ (A C OS E'' C' \ A D OS D' E'')" proof - have P35_1: "C A OS D E''" by (metis Col_perm P11 P18 P25 bet_out between_symmetry one_side_symmetry out_one_side) have P35_2: "C A OS B D" using P10 assms(1) one_side_symmetry ts_ts_os by blast have P35_3: "C A TS B C'" by (metis P2 P29 bet__ts cong_diff_4 not_col_distincts) have P35_4: "C A OS B E''" using P35_1 P35_2 one_side_transitivity by blast have P35_5: "D A OS C E''" by (metis Col_perm P18 P24 P35_1 bet2__out l5_1 one_side_not_col123 out_one_side) have P35_6: "D A OS B C" by (simp add: P10 assms(1) invert_two_sides l9_2 one_side_symmetry ts_ts_os) have P35_7: "D A TS B D'" by (metis P30 TS_def assms(1) bet__ts cong_diff_3 ts_distincts) have P35_8: "D A OS B E''" using P35_5 P35_6 one_side_transitivity by blast have P35_9: "A C TS E'' C'" using P35_3 P35_4 invert_two_sides l9_8_2 by blast have "A D TS D' E''" using P35_7 P35_8 invert_two_sides l9_2 l9_8_2 by blast thus ?thesis using P35_9 by simp qed moreover have "E'' A C CongA D' A D" proof - have "E'' A C CongA B A D" by (simp add: P28 conga_comm) moreover have "B A D CongA D' A D" by (simp add: P32) ultimately show ?thesis using conga_trans by blast qed moreover have "C A C' CongA D A E''" proof - have "D A E'' CongA C A C'" proof - have "D A E'' CongA B A C" by (simp add: P20 conga_sym) moreover have "B A C CongA C A C'" by (simp add: P34 conga_right_comm) ultimately show ?thesis using conga_trans by blast qed thus ?thesis using not_conga_sym by blast qed ultimately show ?thesis using l11_22 by auto qed have P36: "D' \ B" using P30 assms(1) bet_neq32__neq ts_distincts by blast have P37: "C' \ B" using P29 assms(1) bet_neq32__neq ts_distincts by blast then have P38: "\ Col C' D' B" by (metis Col_def P10 P29 P30 P36 TS_def col_transitivity_2) have P39: "C' D' ParStrict C D" proof - have "\ Col C' D' B" by (simp add: P38) moreover have "D Midpoint D' B" using P30 l7_2 midpoint_def not_cong_3412 by blast moreover have "C Midpoint C' B" using P29 l7_2 midpoint_def not_cong_3412 by blast ultimately show ?thesis using triangle_mid_par by auto qed have P40: "A E'' TS C D" by (metis Bet_perm Col_def P10 P18 P24 TS_def \C = E'' \ False\ bet__ts col_transitivity_2 invert_two_sides) have P41: "B A TS C D" by (simp add: assms(1) invert_two_sides) have P42: "A B OS C C'" proof - have "\ Col A B C" by (simp add: P2 not_col_permutation_1) moreover have "Col A B B" by (simp add: col_trivial_2) moreover have "B Out C C'" by (metis P29 P37 bet_out cong_identity) ultimately show ?thesis using out_one_side_1 by blast qed have P43: "A B OS D D'" using out_one_side_1 by (metis Col_perm P30 TS_def assms(1) bet_out col_trivial_1) then have P44: "A B OS D D'" using invert_two_sides by blast have P45: "A B TS C' D" using P42 assms(1) l9_8_2 by blast then have P46: "A B TS C' D'" using P44 l9_2 l9_8_2 by blast have P47: "C' D' Perp A E''" proof - have "A E'' TS C' D'" proof - have "A Out C' D' \ E'' A TS C' D'" proof - have "E'' A C' CongA E'' A D'" by (simp add: P35 conga_right_comm) moreover have "Coplanar E'' A C' D'" proof - have f1: "B A OS C C'" by (metis P42 invert_one_side) have f2: "Coplanar B A C' C" by (meson P42 ncoplanar_perm_7 os__coplanar) have f3: "Coplanar D' A C' D" by (meson P44 P46 col124__nos coplanar_trans_1 invert_one_side ncoplanar_perm_7 os__coplanar ts__coplanar) have "Coplanar D' A C' C" using f2 f1 by (meson P46 col124__nos coplanar_trans_1 ncoplanar_perm_6 ncoplanar_perm_8 ts__coplanar) then show ?thesis using f3 by (meson P18 bet_cop2__cop ncoplanar_perm_6 ncoplanar_perm_7 ncoplanar_perm_8) qed ultimately show ?thesis using conga_cop__or_out_ts by simp qed then show ?thesis using P46 col_two_sides_bet invert_two_sides not_bet_and_out out_col by blast qed moreover have "Cong C' A D' A" using Cong3_def P31 \B A C Cong3 C' A C\ cong_inner_transitivity by blast moreover have "C' A E'' CongA D' A E''" by (simp add: P35 conga_left_comm) ultimately show ?thesis by (simp add: cong_conga_perp) qed have T1: "Cong A C' A D'" proof - have "Cong A C' A B" using Cong3_def Cong_perm \B A C Cong3 C' A C\ by blast moreover have "Cong A D' A B" using Cong3_def P31 not_cong_4321 by blast ultimately show ?thesis using Cong_perm \Cong A C' A B\ \Cong A D' A B\ cong_inner_transitivity by blast qed obtain R where T2: "R Midpoint C' D'" using midpoint_existence by auto have "\ X Y. (R PerpAt X Y C' D' \ X Y Perp D C \ Coplanar C' D' B X \ Coplanar C' D' B Y)" proof - have "\ Col C' D' B" by (simp add: P38) moreover have "D Midpoint D' B" using P30 l7_2 midpoint_def not_cong_3412 by blast moreover have "C Midpoint C' B" using Cong_perm Mid_perm Midpoint_def P29 by blast moreover have "R Midpoint C' D'" by (simp add: T2) ultimately show ?thesis using l13_1_aux by blast qed then obtain X Y where T3: "R PerpAt X Y C' D' \ X Y Perp D C \ Coplanar C' D' B X \ Coplanar C' D' B Y" by blast then have "X \ Y" using perp_not_eq_1 by blast have "C D Perp A E''" proof cases assume "A = R" then have W1: "A PerpAt C' D' A E''" using Col_def P47 T2 between_trivial2 l8_14_2_1b_bis midpoint_col by blast have "Coplanar B C' D' E''" proof - have "\ Col B C D" using P10 TS_def by auto moreover have "Coplanar B C D B" using ncop_distincts by auto moreover have "Coplanar B C D C'" using P29 bet_col ncop__ncols by blast moreover have "Coplanar B C D D'" using P30 bet_col ncop__ncols by blast moreover have "Coplanar B C D E''" by (simp add: P18 bet__coplanar coplanar_perm_22) ultimately show ?thesis using coplanar_pseudo_trans by blast qed have "Coplanar C' D' X E''" proof - have "\ Col B C' D'" by (simp add: P38 not_col_permutation_2) moreover have "Coplanar B C' D' X" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' E''" by (simp add: \Coplanar B C' D' E''\) ultimately show ?thesis using coplanar_trans_1 by blast qed have "Coplanar C' D' Y E''" proof - have "\ Col B C' D'" by (simp add: P38 not_col_permutation_2) moreover have "Coplanar B C' D' Y" by (simp add: T3 coplanar_perm_12) moreover have "Coplanar B C' D' E''" by (simp add: \Coplanar B C' D' E''\) ultimately show ?thesis using coplanar_trans_1 by blast qed have "Coplanar C' D' X A" proof - have "Col C' D' A" using T2 \A = R\ midpoint_col not_col_permutation_2 by blast moreover have "Col X A A" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Coplanar C' D' Y A" proof - have "Col C' D' A" using T2 \A = R\ midpoint_col not_col_permutation_2 by blast moreover have "Col Y A A" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Col X Y A" proof - have "Coplanar C' D' X A" by (simp add: \Coplanar C' D' X A\) moreover have "Coplanar C' D' X E''" by (simp add: \Coplanar C' D' X E''\) moreover have "Coplanar C' D' Y A" by (simp add: \Coplanar C' D' Y A\) moreover have "Coplanar C' D' Y E''" by (simp add: \Coplanar C' D' Y E''\) moreover have "A PerpAt X Y C' D'" using T3 \A = R\ Perp_in_cases by auto moreover have "A PerpAt A E'' C' D'" using Perp_in_cases \A PerpAt C' D' A E''\ by blast ultimately show ?thesis using cop4_perp_in2__col by blast qed have "Col X Y E''" proof - have "Coplanar C' D' X E''" using \Coplanar C' D' X E''\ by auto moreover have "Coplanar C' D' X A" by (simp add: \Coplanar C' D' X A\) moreover have "Coplanar C' D' Y E''" by (simp add: \Coplanar C' D' Y E''\) moreover have "Coplanar C' D' Y A" using \Coplanar C' D' Y A\ by auto moreover have "A PerpAt X Y C' D'" using T3 \A = R\ Perp_in_cases by auto moreover have "A PerpAt E'' A C' D'" using Perp_in_perm W1 by blast ultimately show ?thesis using cop4_perp_in2__col by blast qed have "A E'' Perp C D" proof cases assume "Y = A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A X Perp C D" using T3 Perp_cases \Y = A\ by blast moreover have "Col A X E''" using Col_perm \Col X Y E''\ \Y = A\ by blast ultimately show ?thesis using perp_col by blast qed next assume "Y \ A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A Y Perp C D" proof - have "Y X Perp C D" using T3 by (simp add: perp_comm) then have "Y A Perp C D" using \Col X Y A\ \Y \ A\ col_trivial_2 perp_col2 perp_left_comm by blast then show ?thesis using Perp_cases by blast qed moreover have "Col A Y E''" using Col_perm \Col X Y A\ \Col X Y E''\ \X \ Y\ col_transitivity_2 by blast ultimately show ?thesis using perp_col by blast qed qed thus ?thesis using Perp_perm by blast next assume "A \ R" have "R \ C'" using P46 T2 is_midpoint_id ts_distincts by blast have "Per A R C'" using T1 T2 Per_def by blast then have "R PerpAt A R R C'" by (simp add: \A \ R\ \R \ C'\ per_perp_in) then have "R PerpAt R C' A R" using Perp_in_perm by blast then have "R C' Perp A R \ R R Perp A R" using perp_in_perp by auto { assume "R C' Perp A R" then have "C' R Perp A R" by (simp add: \R C' Perp A R\ Perp_perm) have "C' D' Perp R A" by (metis P47 T2 \A \ R\ \Per A R C'\ \R \ C'\ col_per_perp midpoint_col perp_distinct perp_right_comm) then have "R PerpAt C' D' R A" using T2 l8_14_2_1b_bis midpoint_col not_col_distincts by blast have "Col B D D'" by (simp add: Col_def P30) have "Col B C C'" using Col_def P29 by auto have "Col D E'' C" using P18 bet_col by auto have "Col R C' D'" using \R PerpAt C' D' R A\ by (simp add: T2 midpoint_col) have "Col A E'' E'" by (simp add: P19 out_col) have "Coplanar C' D' X A" proof - have "\ Col B C' D'" using Col_perm P38 by blast moreover have "Coplanar B C' D' X" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' A" using P46 ncoplanar_perm_18 ts__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by auto qed have "Coplanar C' D' Y A" proof - have "\ Col B C' D'" using Col_perm P38 by blast moreover have "Coplanar B C' D' Y" using T3 ncoplanar_perm_8 by blast moreover have "Coplanar B C' D' A" using P46 ncoplanar_perm_18 ts__coplanar by blast ultimately show ?thesis using coplanar_trans_1 by auto qed have "Coplanar C' D' X R" proof - have "Col C' D' R" using Col_perm \Col R C' D'\ by blast moreover have "Col X R R" by (simp add: col_trivial_2) ultimately show ?thesis using ncop__ncols by blast qed have "Coplanar C' D' Y R" using Col_perm T2 midpoint_col ncop__ncols by blast have "Col X Y A" proof - have "R PerpAt X Y C' D'" using T3 by simp moreover have "R PerpAt A R C' D'" using Perp_in_perm \R PerpAt C' D' R A\ by blast ultimately show ?thesis using \Coplanar C' D' Y R\ \Coplanar C' D' X R\ cop4_perp_in2__col \Coplanar C' D' X A\ \Coplanar C' D' Y A\ by blast qed have Z1: "Col X Y R" using T3 perp_in_col by blast have "Col A E'' R" proof - have "Coplanar C' D' E'' R" using Col_cases \Col R C' D'\ ncop__ncols by blast moreover have "A E'' Perp C' D'" using P47 Perp_perm by blast moreover have "A R Perp C' D'" using Perp_perm \C' D' Perp R A\ by blast ultimately show ?thesis using cop_perp2__col by blast qed then have "Col X Y E''" using Z1 by (metis (full_types) \A \ R\ \Col X Y A\ col_permutation_4 col_trivial_2 l6_21) have "Col A E'' R" proof - have "Coplanar C' D' E'' R" using Col_cases \Col R C' D'\ ncop__ncols by blast moreover have "A E'' Perp C' D'" using P47 Perp_perm by blast moreover have "A R Perp C' D'" using Perp_perm \C' D' Perp R A\ by blast ultimately show ?thesis using cop_perp2__col by blast qed have "Col A R X" using \Col X Y A\ \Col X Y R\ \X \ Y\ col_transitivity_1 not_col_permutation_3 by blast have "Col A R Y" using \Col X Y A\ \Col X Y R\ \X \ Y\ col_transitivity_2 not_col_permutation_3 by blast have "A E'' Perp C D" proof cases assume "X = A" show ?thesis proof - have "A \ E''" by (simp add: P26) moreover have "A Y Perp C D" using T3 \X = A\ perp_right_comm by blast moreover have "Col A Y E''" using Col_perm \A \ R\ \Col A E'' R\ \Col A R Y\ col_transitivity_1 by blast ultimately show ?thesis using perp_col by auto qed next assume "X \ A" show ?thesis proof - have "A X Perp C D" by (smt P3 T3 \Col X Y A\ \X \ A\ col_trivial_2 col_trivial_3 perp_col4) moreover have "Col A X E''" using Col_perm \A \ R\ \Col A E'' R\ \Col A R X\ col_transitivity_1 by blast ultimately show ?thesis using P26 perp_col by blast qed qed } { assume "R R Perp A R" then have "A E'' Perp C D" using perp_distinct by blast } then have "A E'' Perp C D" using Perp_cases \R C' Perp A R \ A E'' Perp C D\ \R C' Perp A R \ R R Perp A R\ by auto then show ?thesis using Perp_perm by blast qed show ?thesis proof - have "Col A E E''" proof - have "Coplanar C D E E'" using assms(4) col__coplanar by auto moreover have "A E Perp C D" using assms(5) by auto moreover have "A E'' Perp C D" using Perp_perm \C D Perp A E''\ by blast ultimately show ?thesis by (meson P11 col_perp2_ncol_col col_trivial_3 not_col_permutation_2) qed moreover have "E'' = E" proof - have f1: "C = E'' \ Col C E'' D" by (metis P18 bet_out_1 out_col) then have f2: "C = E'' \ Col C E'' E" using Col_perm P3 assms(4) col_transitivity_1 by blast have "\p. (C = E'' \ Col C p D) \ \ Col C E'' p" using f1 by (meson col_transitivity_1) then have "\p. \ Col E'' p A \ Col E'' E p" using f2 by (metis (no_types) Col_perm P11 assms(4)) then show ?thesis using Col_perm calculation col_transitivity_1 by blast qed ultimately show ?thesis by (metis Bet_perm P18 P20 P28 Tarski_neutral_dimensionless.conga_left_comm Tarski_neutral_dimensionless_axioms not_conga_sym) qed qed then have "B A C CongA D A E \ B A D CongA C A E \ Bet C E D" by blast } thus ?thesis using P18 \E'' = A \ B A C CongA D A E \ B A D CongA C A E \ Bet C E D\ by blast qed lemma perp2_refl: assumes "A \ B" shows "P Perp2 A B A B" proof cases assume "Col A B P" obtain X where "\ Col A B X" using assms not_col_exists by blast then obtain Q where "A B Perp Q P \ A B OS X Q" using \Col A B P\ l10_15 by blast thus ?thesis using Perp2_def Perp_cases col_trivial_3 by blast next assume "\ Col A B P" then obtain Q where "Col A B Q \ A B Perp P Q" using l8_18_existence by blast thus ?thesis using Perp2_def Perp_cases col_trivial_3 by blast qed lemma perp2_sym: assumes "P Perp2 A B C D" shows "P Perp2 C D A B" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def by blast qed lemma perp2_left_comm: assumes "P Perp2 A B C D" shows "P Perp2 B A C D" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def perp_right_comm by blast qed lemma perp2_right_comm: assumes "P Perp2 A B C D" shows "P Perp2 A B D C" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using Perp2_def perp_right_comm by blast qed lemma perp2_comm: assumes "P Perp2 A B C D" shows "P Perp2 B A D C" proof - obtain X Y where "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms by auto thus ?thesis using assms perp2_left_comm perp2_right_comm by blast qed lemma perp2_pseudo_trans: assumes "P Perp2 A B C D" and "P Perp2 C D E F" and "\ Col C D P" shows "P Perp2 A B E F" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by auto obtain X' Y' where P2: "Col P X' Y' \ X' Y' Perp C D \ X' Y' Perp E F" using Perp2_def assms(2) by auto have "X Y Par X' Y'" proof - have "Coplanar P C D X" proof cases assume "X = P" thus ?thesis using ncop_distincts by blast next assume "X \ P" then have "X P Perp C D" using Col_cases P1 perp_col by blast then have "Coplanar X P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D Y" proof cases assume "Y = P" thus ?thesis using ncop_distincts by blast next assume "Y \ P" then have "Y P Perp C D" by (metis (full_types) Col_cases P1 Perp_cases col_transitivity_2 perp_col2) then have "Coplanar Y P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D X'" proof cases assume "X' = P" thus ?thesis using ncop_distincts by blast next assume "X' \ P" then have "X' P Perp C D" using Col_cases P2 perp_col by blast then have "Coplanar X' P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed have "Coplanar P C D Y'" proof cases assume "Y' = P" thus ?thesis using ncop_distincts by blast next assume "Y' \ P" then have "Y' P Perp C D" by (metis (full_types) Col_cases P2 Perp_cases col_transitivity_2 perp_col2) then have "Coplanar Y' P C D" by (simp add: perp__coplanar) thus ?thesis using ncoplanar_perm_18 by blast qed show ?thesis proof - have "Coplanar C D X X'" using Col_cases \Coplanar P C D X'\ \Coplanar P C D X\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D X Y'" using Col_cases \Coplanar P C D X\ \Coplanar P C D Y'\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D Y X'" using Col_cases \Coplanar P C D X'\ \Coplanar P C D Y\ assms(3) coplanar_trans_1 by blast moreover have "Coplanar C D Y Y'" using Col_cases \Coplanar P C D Y'\ \Coplanar P C D Y\ assms(3) coplanar_trans_1 by blast ultimately show ?thesis using l12_9 P1 P2 by blast qed qed thus ?thesis proof - { assume "X Y ParStrict X' Y'" then have "Col X X' Y'" using P1 P2 \X Y ParStrict X' Y'\ par_not_col by blast } then have "Col X X' Y'" using Par_def \X Y Par X' Y'\ by blast moreover have "Col Y X' Y'" proof - { assume "X Y ParStrict X' Y'" then have "Col Y X' Y'" using P1 P2 \X Y ParStrict X' Y'\ par_not_col by blast } thus ?thesis using Par_def \X Y Par X' Y'\ by blast qed moreover have "X \ Y" using P1 perp_not_eq_1 by auto ultimately show ?thesis by (meson Perp2_def P1 P2 col_permutation_1 perp_col2) qed qed lemma col_cop_perp2__pars_bis: assumes "\ Col A B P" and "Col C D P" and "Coplanar A B C D" and "P Perp2 A B C D" shows "A B ParStrict C D" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(4) by auto then have "Col X Y P" using Col_perm by blast obtain Q where "X \ Q \ Y \ Q \ P \ Q \ Col X Y Q" using \Col X Y P\ diff_col_ex3 by blast thus ?thesis by (smt P1 Perp_perm assms(1) assms(2) assms(3) col_cop_perp2_pars col_permutation_1 col_transitivity_2 not_col_distincts perp_col4 perp_distinct) qed lemma perp2_preserves_bet23: assumes "Bet PO A B" and "Col PO A' B'" and "\ Col PO A A'" and "PO Perp2 A A' B B'" shows "Bet PO A' B'" proof - have "A \ A'" using assms(3) not_col_distincts by auto show ?thesis proof cases assume "A' = B'" thus ?thesis using between_trivial by auto next assume "A' \ B'" { assume "A = B" then obtain X Y where P1: "Col PO X Y \ X Y Perp A A' \ X Y Perp A B'" using Perp2_def assms(4) by blast have "Col A A' B'" proof - have "Coplanar X Y A' B'" using Col_cases Coplanar_def P1 assms(2) by auto moreover have "A A' Perp X Y" using P1 Perp_perm by blast moreover have "A B' Perp X Y" using P1 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by blast qed then have "False" using Col_perm \A' \ B'\ assms(2) assms(3) l6_16_1 by blast } then have "A \ B" by auto have "A A' Par B B'" proof - obtain X Y where P2: "Col PO X Y \ X Y Perp A A' \ X Y Perp B B'" using Perp2_def assms(4) by auto then have "Coplanar X Y A B" using Coplanar_def assms(1) bet_col not_col_permutation_2 by blast show ?thesis proof - have "Coplanar X Y A B'" by (metis (full_types) Col_cases P2 assms(2) assms(3) col_cop2__cop col_trivial_3 ncop__ncols perp__coplanar) moreover have "Coplanar X Y A' B" proof cases assume "Col A X Y" then have "Col Y X A" by (metis (no_types) Col_cases) then show ?thesis by (metis Col_cases P2 assms(1) assms(3) bet_col colx ncop__ncols not_col_distincts) next assume "\ Col A X Y" moreover have "Coplanar A X Y A'" using Coplanar_def P2 perp_inter_exists by blast moreover have "Coplanar A X Y B" using \Coplanar X Y A B\ ncoplanar_perm_8 by blast ultimately show ?thesis using coplanar_trans_1 by auto qed moreover have "Coplanar X Y A' B'" using Col_cases Coplanar_def P2 assms(2) by auto moreover have "A A' Perp X Y" using P2 Perp_perm by blast moreover have "B B' Perp X Y" using P2 Perp_perm by blast ultimately show ?thesis using \Coplanar X Y A B\ l12_9 by auto qed qed { assume "A A' ParStrict B B'" then have "A A' OS B B'" by (simp add: l12_6) have "A A' TS PO B" using Col_cases \A \ B\ assms(1) assms(3) bet__ts by blast then have "A A' TS B' PO" using \A A' OS B B'\ l9_2 l9_8_2 by blast then have "Bet PO A' B'" using Col_cases assms(2) between_symmetry col_two_sides_bet invert_two_sides by blast } thus ?thesis by (metis Col_cases Par_def \A A' Par B B'\ \A \ B\ assms(1) assms(3) bet_col col_trivial_3 l6_21) qed qed lemma perp2_preserves_bet13: assumes "Bet B PO C" and "Col PO B' C'" and "\ Col PO B B'" and "PO Perp2 B C' C B'" shows "Bet B' PO C'" proof cases assume "C' = PO" thus ?thesis using not_bet_distincts by blast next assume "C' \ PO" show ?thesis proof cases assume "B' = PO" thus ?thesis using between_trivial2 by auto next assume "B' \ PO" have "B \ PO" using assms(3) col_trivial_1 by auto have "Col B PO C" by (simp add: Col_def assms(1)) show ?thesis proof cases assume "B = C" thus ?thesis using \B = C\ \B \ PO\ assms(1) between_identity by blast next assume "B \ C" have "B C' Par C B'" proof - obtain X Y where P1: "Col PO X Y \ X Y Perp B C' \ X Y Perp C B'" using Perp2_def assms(4) by auto have "Coplanar X Y B C" by (meson P1 \Col B PO C\ assms(1) l9_18_R2 ncop__ncols not_col_permutation_2 not_col_permutation_5 ts__coplanar) have "Coplanar X Y C' B'" using Col_cases Coplanar_def P1 assms(2) by auto show ?thesis proof - have "Coplanar X Y B C" by (simp add: \Coplanar X Y B C\) moreover have "Coplanar X Y B B'" by (metis P1 \C' \ PO\ assms(1) assms(2) bet_cop__cop calculation col_cop2__cop not_col_permutation_5 perp__coplanar) moreover have "Coplanar X Y C' C" by (smt P1 \B \ PO\ \Col B PO C\ \Coplanar X Y C' B'\ assms(2) col2_cop__cop col_cop2__cop col_permutation_1 col_transitivity_2 coplanar_perm_1 perp__coplanar) moreover have "Coplanar X Y C' B'" by (simp add: \Coplanar X Y C' B'\) moreover have "B C' Perp X Y" using P1 Perp_perm by blast moreover have "C B' Perp X Y" by (simp add: P1 Perp_perm) ultimately show ?thesis using l12_9 by blast qed qed have "B C' ParStrict C B'" by (metis Out_def Par_def \B C' Par C B'\ \B \ C\ \B \ PO\ assms(1) assms(3) col_transitivity_1 not_col_permutation_4 out_col) have "B' \ PO" by (simp add: \B' \ PO\) obtain X Y where P5: "Col PO X Y \ X Y Perp B C' \ X Y Perp C B'" using Perp2_def assms(4) by auto have "X \ Y" using P5 perp_not_eq_1 by auto show ?thesis proof cases assume "Col X Y B" have "Col X Y C" using P5 \B \ PO\ \Col B PO C\ \Col X Y B\ col_permutation_1 colx by blast show ?thesis proof - have "Col B' PO C'" using Col_cases assms(2) by auto moreover have "Per PO C B'" by (metis P5 \B C' ParStrict C B'\ \Col X Y C\ assms(2) col_permutation_2 par_strict_not_col_2 perp_col2 perp_per_2) moreover have "Per PO B C'" using P5 \B \ PO\ \Col X Y B\ col_permutation_1 perp_col2 perp_per_2 by blast ultimately show ?thesis by (metis Tarski_neutral_dimensionless.per13_preserves_bet_inv Tarski_neutral_dimensionless_axioms \B C' ParStrict C B'\ assms(1) assms(3) between_symmetry not_col_distincts not_col_permutation_3 par_strict_not_col_2) qed next assume "\ Col X Y B" then obtain B0 where U1: "Col X Y B0 \ X Y Perp B B0" using l8_18_existence by blast have "\ Col X Y C" by (smt P5 \B C' ParStrict C B'\ \Col B PO C\ \\ Col X Y B\ assms(2) col_permutation_2 colx par_strict_not_col_2) then obtain C0 where U2: "Col X Y C0 \ X Y Perp C C0" using l8_18_existence by blast have "B0 \ PO" by (metis P5 Perp_perm \Col B PO C\ \Col X Y B0 \ X Y Perp B B0\ \\ Col X Y C\ assms(3) col_permutation_2 col_permutation_3 col_perp2_ncol_col) { assume "C0 = PO" then have "C PO Par C B'" by (metis P5 Par_def Perp_cases \Col X Y C0 \ X Y Perp C C0\ \\ Col X Y C\ col_perp2_ncol_col not_col_distincts not_col_permutation_3 perp_distinct) then have "False" by (metis \B C' ParStrict C B'\ assms(2) assms(3) col3 not_col_distincts par_id_2 par_strict_not_col_2) } then have "C0 \ PO" by auto have "Bet B0 PO C0" proof - have "Bet B PO C" by (simp add: assms(1)) moreover have "PO \ B0" using \B0 \ PO\ by auto moreover have "PO \ C0" using \C0 \ PO\ by auto moreover have "Col B0 PO C0" using U1 U2 P5 \X \ Y\ col3 not_col_permutation_2 by blast moreover have "Per PO B0 B" proof - have "B0 PerpAt PO B0 B0 B" proof cases assume "X = B0" have "B0 PO Perp B B0" by (metis P5 U1 calculation(2) col3 col_trivial_2 col_trivial_3 perp_col2) show ?thesis proof - have "B0 \ PO" using calculation(2) by auto moreover have "B0 Y Perp B B0" using U1 \X = B0\ by auto moreover have "Col B0 Y PO" using Col_perm P5 \X = B0\ by blast ultimately show ?thesis using \B0 PO Perp B B0\ perp_in_comm perp_perp_in by blast qed next assume "X \ B0" have "X B0 Perp B B0" using U1 \X \ B0\ perp_col by blast have "B0 PO Perp B B0" by (metis P5 U1 calculation(2) not_col_permutation_2 perp_col2) then have "B0 PerpAt B0 PO B B0" by (simp add: perp_perp_in) thus ?thesis using Perp_in_perm by blast qed then show ?thesis by (simp add: perp_in_per) qed moreover have "Per PO C0 C" proof - have "C0 PO Perp C C0" by (metis P5 U2 calculation(3) col3 col_trivial_2 col_trivial_3 perp_col2) then have "C0 PerpAt PO C0 C0 C" by (simp add: perp_in_comm perp_perp_in) thus ?thesis using perp_in_per_2 by auto qed ultimately show ?thesis using per13_preserves_bet by blast qed show ?thesis proof cases assume "C' = B0" have "B' = C0" proof - have "\ Col C' PO C" using P5 U1 \B0 \ PO\ \C' = B0\ \\ Col X Y C\ colx not_col_permutation_3 not_col_permutation_4 by blast moreover have "C \ C0" using U2 \\ Col X Y C\ by auto moreover have "Col C C0 B'" proof - have "Coplanar X Y C0 B'" proof - have "Col X Y C0" by (simp add: U2) moreover have "Col C0 B' C0" by (simp add: col_trivial_3) ultimately show ?thesis using ncop__ncols by blast qed moreover have "C C0 Perp X Y" using Perp_perm U2 by blast moreover have "C B' Perp X Y" using P5 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by auto qed ultimately show ?thesis by (metis Col_def \C' = B0\ \Bet B0 PO C0\ assms(2) colx) qed show ?thesis using Bet_cases \B' = C0\ \C' = B0\ \Bet B0 PO C0\ by blast next assume "C' \ B0" then have "B' \ C0" by (metis P5 U1 U2 \C0 \ PO\ assms(2) col_permutation_1 colx l8_18_uniqueness) have "B C' Par B B0" proof - have "Coplanar X Y B B" using ncop_distincts by auto moreover have "Coplanar X Y B B0" using U1 ncop__ncols by blast moreover have "Coplanar X Y C' B" using P5 ncoplanar_perm_1 perp__coplanar by blast moreover have "Coplanar X Y C' B0" using \\ Col X Y B\ calculation(2) calculation(3) col_permutation_1 coplanar_perm_12 coplanar_perm_18 coplanar_trans_1 by blast moreover have "B C' Perp X Y" using P5 Perp_perm by blast moreover have "B B0 Perp X Y" using Perp_perm U1 by blast ultimately show ?thesis using l12_9 by blast qed { assume "B C' ParStrict B B0" have "Col B B0 C'" by (simp add: \B C' Par B B0\ par_id_3) } then have "Col B B0 C'" using \B C' Par B B0\ par_id_3 by blast have "Col C C0 B'" proof - have "Coplanar X Y C0 B'" by (simp add: U2 col__coplanar) moreover have "C C0 Perp X Y" by (simp add: Perp_perm U2) moreover have "C B' Perp X Y" using P5 Perp_perm by blast ultimately show ?thesis using cop_perp2__col by auto qed show ?thesis proof - have "Col B' PO C'" using assms(2) not_col_permutation_4 by blast moreover have "Per PO C0 B'" proof - have "C0 PerpAt PO C0 C0 B'" proof cases assume "X = C0" have "C0 PO Perp C B'" proof - have "C0 \ PO" by (simp add: \C0 \ PO\) moreover have "C0 Y Perp C B'" using P5 \X = C0\ by auto moreover have "Col C0 Y PO" using Col_perm P5 \X = C0\ by blast ultimately show ?thesis using perp_col by blast qed then have "B' C0 Perp C0 PO" using Perp_perm \B' \ C0\ \Col C C0 B'\ not_col_permutation_1 perp_col1 by blast then have "C0 PerpAt C0 B' PO C0" using Perp_perm perp_perp_in by blast thus ?thesis using Perp_in_perm by blast next assume "X \ C0" then have "X C0 Perp C B'" using P5 U2 perp_col by blast have "C0 PO Perp C B'" using Col_cases P5 U2 \C0 \ PO\ perp_col2 by blast then have "B' C0 Perp C0 PO" using Perp_cases \B' \ C0\ \Col C C0 B'\ col_permutation_2 perp_col by blast thus ?thesis using Perp_in_perm Perp_perm perp_perp_in by blast qed then show ?thesis using perp_in_per_2 by auto qed moreover have "Per PO B0 C'" proof - have "B0 PerpAt PO B0 B0 C'" proof - have "Col C' B B0" using Col_cases \Col B B0 C'\ by blast then have "C' B0 Perp X Y" using perp_col P5 Perp_cases \C' \ B0\ by blast show ?thesis proof - have "PO B0 Perp B0 C'" by (smt P5 U1 \B0 \ PO\ \C' \ B0\ \Col B B0 C'\ col_trivial_2 not_col_permutation_2 perp_col4) then show ?thesis using Perp_in_cases Perp_perm perp_perp_in by blast qed qed thus ?thesis by (simp add: perp_in_per) qed ultimately show ?thesis using \B0 \ PO\ \C0 \ PO\ \Bet B0 PO C0\ between_symmetry per13_preserves_bet_inv by blast qed qed qed qed qed qed lemma is_image_perp_in: assumes "A \ A'" and "X \ Y" and "A A' Reflect X Y" shows "\ P. P PerpAt A A' X Y" by (metis Perp_def Tarski_neutral_dimensionless.Perp_perm Tarski_neutral_dimensionless_axioms assms(1) assms(2) assms(3) ex_sym1 l10_6_uniqueness) lemma perp_inter_perp_in_n: assumes "A B Perp C D" shows "\ P. Col A B P \ Col C D P \ P PerpAt A B C D" by (simp add: assms perp_inter_perp_in) lemma perp2_perp_in: assumes "PO Perp2 A B C D" and "\ Col PO A B" and "\ Col PO C D" shows "\ P Q. Col A B P \ Col C D Q \ Col PO P Q \ P PerpAt PO P A B \ Q PerpAt PO Q C D" proof - obtain X Y where P1: "Col PO X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by blast have "X \ Y" using P1 perp_not_eq_1 by auto obtain P where P2: "Col X Y P \ Col A B P \ P PerpAt X Y A B" using P1 perp_inter_perp_in_n by blast obtain Q where P3: "Col X Y Q \ Col C D Q \ Q PerpAt X Y C D" using P1 perp_inter_perp_in_n by blast have "Col A B P" using P2 by simp moreover have "Col C D Q" using P3 by simp moreover have "Col PO P Q" using P2 P3 P1 \X \ Y\ col3 not_col_permutation_2 by blast moreover have "P PerpAt PO P A B" proof cases assume "X = PO" thus ?thesis by (metis P2 assms(2) not_col_permutation_3 not_col_permutation_4 perp_in_col_perp_in perp_in_sym) next assume "X \ PO" then have "P PerpAt A B X PO" by (meson Col_cases P1 P2 perp_in_col_perp_in perp_in_sym) then have "P PerpAt A B PO X" using Perp_in_perm by blast then have "P PerpAt A B PO P" by (metis Col_cases assms(2) perp_in_col perp_in_col_perp_in) thus ?thesis by (simp add: perp_in_sym) qed moreover have "Q PerpAt PO Q C D" by (metis P1 P3 \X \ Y\ assms(3) col_trivial_2 colx not_col_permutation_3 not_col_permutation_4 perp_in_col_perp_in perp_in_right_comm perp_in_sym) ultimately show ?thesis by blast qed lemma l13_8: assumes "U \ PO" and "V \ PO" and "Col PO P Q" and "Col PO U V" and "Per P U PO" and "Per Q V PO" shows "PO Out P Q \ PO Out U V" by (smt Out_def assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) l8_2 not_col_permutation_5 per23_preserves_bet per23_preserves_bet_inv per_distinct_1) lemma perp_in_rewrite: assumes "P PerpAt A B C D" shows "P PerpAt A P P C \ P PerpAt A P P D \ P PerpAt B P P C \ P PerpAt B P P D" by (metis assms per_perp_in perp_in_distinct perp_in_per_1 perp_in_per_3 perp_in_per_4) lemma perp_out_acute: assumes "B Out A C'" and "A B Perp C C'" shows "Acute A B C" proof - have "A \ B" using assms(1) out_diff1 by auto have "C' \ B" using Out_def assms(1) by auto then have "B C' Perp C C'" by (metis assms(1) assms(2) out_col perp_col perp_comm perp_right_comm) then have "Per C C' B" using Perp_cases perp_per_2 by blast then have "Acute C' C B \ Acute C' B C" by (metis \C' \ B\ assms(2) l11_43 perp_not_eq_2) have "C \ B" using \B C' Perp C C'\ l8_14_1 by auto show ?thesis proof - have "B Out A C'" by (simp add: assms(1)) moreover have "B Out C C" by (simp add: \C \ B\ out_trivial) moreover have "Acute C' B C" by (simp add: \Acute C' C B \ Acute C' B C\) ultimately show ?thesis using acute_out2__acute by auto qed qed lemma perp_bet_obtuse: assumes "B \ C'" and "A B Perp C C'" and "Bet A B C'" shows "Obtuse A B C" proof - have "Acute C' B C" proof - have "B Out C' C'" using assms(1) out_trivial by auto moreover have "Col A B C'" by (simp add: Col_def assms(3)) then have "C' B Perp C C'" using Out_def assms(2) assms(3) bet_col1 calculation perp_col2 by auto ultimately show ?thesis using perp_out_acute by blast qed thus ?thesis using acute_bet__obtuse assms(2) assms(3) between_symmetry perp_not_eq_1 by blast qed end subsubsection "Part 1: 2D" context Tarski_2D begin lemma perp_in2__col: assumes "P PerpAt A B X Y" and "P PerpAt A' B' X Y" shows "Col A B A'" using cop4_perp_in2__col all_coplanar assms by blast lemma perp2_trans: assumes "P Perp2 A B C D" and "P Perp2 C D E F" shows "P Perp2 A B E F" proof - obtain X Y where P1: "Col P X Y \ X Y Perp A B \ X Y Perp C D" using Perp2_def assms(1) by blast obtain X' Y' where P2: "Col P X' Y' \ X' Y' Perp C D \ X' Y' Perp E F" using Perp2_def assms(2) by blast { assume "X Y Par X' Y'" then have P3: "X Y ParStrict X' Y' \ (X \ Y \ X' \ Y' \ Col X X' Y' \ Col Y X' Y')" using Par_def by blast { assume "X Y ParStrict X' Y'" then have "P Perp2 A B E F" using P1 P2 par_not_col by auto } { assume "X \ Y \ X' \ Y' \ Col X X' Y' \ Col Y X' Y'" then have "P Perp2 A B E F" by (meson P1 P2 Perp2_def col_permutation_1 perp_col2) } then have "P Perp2 A B E F" using P3 \X Y ParStrict X' Y' \ P Perp2 A B E F\ by blast } { assume "\ X Y Par X' Y'" then have "P Perp2 A B E F" using P1 P2 l12_9_2D by blast } thus ?thesis using \X Y Par X' Y' \ P Perp2 A B E F\ by blast qed lemma perp2_par: assumes "PO Perp2 A B C D" shows "A B Par C D" using Perp2_def l12_9_2D Perp_perm assms by blast end subsubsection "Part 2: length" context Tarski_neutral_dimensionless begin lemma lg_exists: "\ l. (QCong l \ l A B)" using QCong_def cong_pseudo_reflexivity by blast lemma lg_cong: assumes "QCong l" and "l A B" and "l C D" shows "Cong A B C D" by (metis QCong_def assms(1) assms(2) assms(3) cong_inner_transitivity) lemma lg_cong_lg: assumes "QCong l" and "l A B" and "Cong A B C D" shows "l C D" by (metis QCong_def assms(1) assms(2) assms(3) cong_transitivity) lemma lg_sym: assumes "QCong l" and "l A B" shows "l B A" using assms(1) assms(2) cong_pseudo_reflexivity lg_cong_lg by blast lemma ex_points_lg: assumes "QCong l" shows "\ A B. l A B" using QCong_def assms cong_pseudo_reflexivity by fastforce lemma is_len_cong: assumes "TarskiLen A B l" and "TarskiLen C D l" shows "Cong A B C D" using TarskiLen_def assms(1) assms(2) lg_cong by auto lemma is_len_cong_is_len: assumes "TarskiLen A B l" and "Cong A B C D" shows "TarskiLen C D l" using TarskiLen_def assms(1) assms(2) lg_cong_lg by fastforce lemma not_cong_is_len: assumes "\ Cong A B C D" and "TarskiLen A B l" shows "\ l C D" using TarskiLen_def assms(1) assms(2) lg_cong by auto lemma not_cong_is_len1: assumes "\ Cong A B C D" and "TarskiLen A B l" shows "\ TarskiLen C D l" using assms(1) assms(2) is_len_cong by blast lemma lg_null_instance: assumes "QCongNull l" shows "l A A" by (metis QCongNull_def QCong_def assms cong_diff cong_trivial_identity) lemma lg_null_trivial: assumes "QCong l" and "l A A" shows "QCongNull l" using QCongNull_def assms(1) assms(2) by auto lemma lg_null_dec: (*assumes "QCong l" *) shows "QCongNull l \ \ QCongNull l" by simp lemma ex_point_lg: assumes "QCong l" shows "\ B. l A B" by (metis QCong_def assms not_cong_3412 segment_construction) lemma ex_point_lg_out: assumes "A \ P" and "QCong l" and "\ QCongNull l" shows "\ B. (l A B \ A Out B P)" proof - obtain X Y where P1: "\ X0 Y0. (Cong X Y X0 Y0 \ l X0 Y0)" using QCong_def assms(2) by auto then have "l X Y" using cong_reflexivity by auto then have "X \ Y" using assms(2) assms(3) lg_null_trivial by auto then obtain B where "A Out P B \ Cong A B X Y" using assms(1) segment_construction_3 by blast thus ?thesis using Cong_perm Out_cases P1 by blast qed lemma ex_point_lg_bet: assumes "QCong l" shows "\ B. (l M B \ Bet A M B)" proof - obtain X Y where P1: "\ X0 Y0. (Cong X Y X0 Y0 \ l X0 Y0)" using QCong_def assms by auto then have "l X Y" using cong_reflexivity by blast obtain B where "Bet A M B \ Cong M B X Y" using segment_construction by blast thus ?thesis using Cong_perm P1 by blast qed lemma ex_points_lg_not_col: assumes "QCong l" and "\ QCongNull l" shows "\ A B. (l A B \ \ Col A B P)" proof - have "\ B::'p. A \ B" using another_point by blast then obtain A::'p where "P \ A" by metis then obtain Q where "\ Col P A Q" using not_col_exists by auto then have "A \ Q" using col_trivial_2 by auto then obtain B where "l A B \ A Out B Q" using assms(1) assms(2) ex_point_lg_out by blast thus ?thesis by (metis \\ Col P A Q\ col_transitivity_1 not_col_permutation_1 out_col out_diff1) qed lemma ex_eql: assumes "\ A B. (TarskiLen A B l1 \ TarskiLen A B l2)" shows "l1 = l2" proof - obtain A B where P1: "TarskiLen A B l1 \ TarskiLen A B l2" using assms by auto have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" proof - have "\ A0 B0. (l1 A0 B0 \ l2 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) moreover have "\ A0 B0. (l2 A0 B0 \ l1 A0 B0)" by (metis TarskiLen_def \TarskiLen A B l1 \ TarskiLen A B l2\ lg_cong lg_cong_lg) ultimately show ?thesis by blast qed thus ?thesis by blast qed lemma all_eql: assumes "TarskiLen A B l1" and "TarskiLen A B l2" shows "l1 = l2" using assms(1) assms(2) ex_eql by auto lemma null_len: assumes "TarskiLen A A la" and "TarskiLen B B lb" shows "la = lb" by (metis TarskiLen_def all_eql assms(1) assms(2) lg_null_instance lg_null_trivial) lemma eqL_equivalence: assumes "QCong la" and "QCong lb" and "QCong lc" shows "la = la \ (la = lb \ lb = la) \ (la = lb \ lb = lc \ la = lc)" by simp lemma ex_lg: "\ l. (QCong l \ l A B)" by (simp add: lg_exists) lemma lg_eql_lg: assumes "QCong l1" and "l1 = l2" shows "QCong l2" using assms(1) assms(2) by auto lemma ex_eqL: assumes "QCong l1" and "QCong l2" and "\ A B. (l1 A B \ l2 A B)" shows "l1 = l2" using TarskiLen_def all_eql assms(1) assms(2) assms(3) by auto subsubsection "Part 3 : angles" lemma ang_exists: assumes "A \ B" and "C \ B" shows "\ a. (QCongA a \ a A B C)" proof - have "A B C CongA A B C" by (simp add: assms(1) assms(2) conga_refl) thus ?thesis using QCongA_def assms(1) assms(2) by auto qed lemma ex_points_eng: assumes "QCongA a" shows "\ A B C. (a A B C)" proof - obtain A B C where "A \ B \ C \ B \ (\ X Y Z. (A B C CongA X Y Z \ a X Y Z))" using QCongA_def assms by auto thus ?thesis using conga_pseudo_refl by blast qed lemma ang_conga: assumes "QCongA a" and "a A B C" and "a A' B' C'" shows "A B C CongA A' B' C'" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongA_def assms(1) by auto thus ?thesis by (meson assms(2) assms(3) not_conga not_conga_sym) qed lemma is_ang_conga: assumes "A B C Ang a" and "A' B' C' Ang a" shows "A B C CongA A' B' C'" using Ang_def ang_conga assms(1) assms(2) by auto lemma is_ang_conga_is_ang: assumes "A B C Ang a" and "A B C CongA A' B' C'" shows "A' B' C' Ang a" proof - have "QCongA a" using Ang_def assms(1) by auto then obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongA_def by auto thus ?thesis by (metis Ang_def assms(1) assms(2) not_conga) qed lemma not_conga_not_ang: assumes "QCongA a" and "\ A B C CongA A' B' C'" and "a A B C" shows "\ a A' B' C'" using ang_conga assms(1) assms(2) assms(3) by auto lemma not_conga_is_ang: assumes "\ A B C CongA A' B' C'" and "A B C Ang a" shows "\ a A' B' C'" using Ang_def ang_conga assms(1) assms(2) by auto lemma not_cong_is_ang1: assumes "\ A B C CongA A' B' C'" and "A B C Ang a" shows "\ A' B' C' Ang a" using assms(1) assms(2) is_ang_conga by blast lemma ex_eqa: assumes "\ A B C.(A B C Ang a1 \ A B C Ang a2)" shows "a1 = a2" proof - obtain A B C where P1: "A B C Ang a1 \ A B C Ang a2" using assms by auto { fix x y z assume "a1 x y z" then have "x y z Ang a1" using Ang_def assms by auto then have "x y z CongA A B C" using P1 not_cong_is_ang1 by blast then have "x y z Ang a2" using P1 is_ang_conga_is_ang not_conga_sym by blast then have "a2 x y z" using Ang_def assms by auto } { fix x y z assume "a2 x y z" then have "x y z Ang a2" using Ang_def assms by auto then have "x y z CongA A B C" using P1 not_cong_is_ang1 by blast then have "x y z Ang a1" using P1 is_ang_conga_is_ang not_conga_sym by blast then have "a1 x y z" using Ang_def assms by auto } then have "\ x y z. (a1 x y z) \ (a2 x y z)" using \\z y x. a1 x y z \ a2 x y z\ by blast then have "\x y. (\ z. (a1 x y) z = (a2 x y) z)" by simp then have "\ x y. (a1 x y) = (a2 x y)" using fun_eq_iff by auto thus ?thesis using fun_eq_iff by auto qed lemma all_eqa: assumes "A B C Ang a1" and "A B C Ang a2" shows "a1 = a2" using assms(1) assms(2) ex_eqa by blast lemma is_ang_distinct: assumes "A B C Ang a" shows "A \ B \ C \ B" using assms conga_diff1 conga_diff2 is_ang_conga by blast lemma null_ang: assumes "A B A Ang a1" and "C D C Ang a2" shows "a1 = a2" using all_eqa assms(1) assms(2) conga_trivial_1 is_ang_conga_is_ang is_ang_distinct by auto lemma flat_ang: assumes "Bet A B C" and "Bet A' B' C'" and "A B C Ang a1" and "A' B' C' Ang a2" shows "a1 = a2" proof - have "A B C Ang a2" proof - have "A' B' C' Ang a2" by (simp add: assms(4)) moreover have "A' B' C' CongA A B C" by (metis assms(1) assms(2) assms(3) calculation conga_line is_ang_distinct) ultimately show ?thesis using is_ang_conga_is_ang by blast qed then show ?thesis using assms(3) all_eqa by auto qed lemma ang_distinct: assumes "QCongA a" and "a A B C" shows "A \ B \ C \ B" proof - have "A B C Ang a" by (simp add: Ang_def assms(1) assms(2)) thus ?thesis using is_ang_distinct by auto qed lemma ex_ang: assumes "B \ A" and "B \ C" shows "\ a. (QCongA a \ a A B C)" using ang_exists assms(1) assms(2) by auto lemma anga_exists: assumes "A \ B" and "C \ B" and "Acute A B C" shows "\ a. (QCongAAcute a \ a A B C)" proof - have "A B C CongA A B C" by (simp add: assms(1) assms(2) conga_refl) thus ?thesis using assms(1) QCongAAcute_def assms(3) by blast qed lemma anga_is_ang: assumes "QCongAAcute a" shows "QCongA a" proof - obtain A0 B0 C0 where P1: "Acute A0 B0 C0 \ (\ X Y Z.(A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongAAcute_def assms by auto thus ?thesis using QCongA_def by (metis acute_distincts) qed lemma ex_points_anga: assumes "QCongAAcute a" shows "\ A B C. a A B C" by (simp add: anga_is_ang assms ex_points_eng) lemma anga_conga: assumes "QCongAAcute a" and "a A B C" and "a A' B' C'" shows "A B C CongA A' B' C'" by (meson Tarski_neutral_dimensionless.ang_conga Tarski_neutral_dimensionless_axioms anga_is_ang assms(1) assms(2) assms(3)) lemma is_anga_to_is_ang: assumes "A B C AngAcute a" shows "A B C Ang a" using AngAcute_def Ang_def anga_is_ang assms by auto lemma is_anga_conga: assumes "A B C AngAcute a" and "A' B' C' AngAcute a" shows "A B C CongA A' B' C'" using AngAcute_def anga_conga assms(1) assms(2) by auto lemma is_anga_conga_is_anga: assumes "A B C AngAcute a" and "A B C CongA A' B' C'" shows "A' B' C' AngAcute a" using Tarski_neutral_dimensionless.AngAcute_def Tarski_neutral_dimensionless.Ang_def Tarski_neutral_dimensionless.is_ang_conga_is_ang Tarski_neutral_dimensionless_axioms assms(1) assms(2) is_anga_to_is_ang by fastforce lemma not_conga_is_anga: assumes "\ A B C CongA A' B' C'" and "A B C AngAcute a" shows "\ a A' B' C'" using AngAcute_def anga_conga assms(1) assms(2) by auto lemma not_cong_is_anga1: assumes "\ A B C CongA A' B' C'" and "A B C AngAcute a" shows "\ A' B' C' AngAcute a" using assms(1) assms(2) is_anga_conga by auto lemma ex_eqaa: assumes "\ A B C. (A B C AngAcute a1 \ A B C AngAcute a2)" shows "a1 = a2" using all_eqa assms is_anga_to_is_ang by blast lemma all_eqaa: assumes "A B C AngAcute a1" and "A B C AngAcute a2" shows "a1 = a2" using assms(1) assms(2) ex_eqaa by blast lemma is_anga_distinct: assumes "A B C AngAcute a" shows "A \ B \ C \ B" using assms is_ang_distinct is_anga_to_is_ang by blast lemma null_anga: assumes "A B A AngAcute a1" and "C D C AngAcute a2" shows "a1 = a2" using assms(1) assms(2) is_anga_to_is_ang null_ang by blast lemma anga_distinct: assumes "QCongAAcute a" and "a A B C" shows "A \ B \ C \ B" using ang_distinct anga_is_ang assms(1) assms(2) by blast lemma out_is_len_eq: assumes "A Out B C" and "TarskiLen A B l" and "TarskiLen A C l" shows "B = C" using Out_def assms(1) assms(2) assms(3) between_cong not_cong_is_len1 by fastforce lemma out_len_eq: assumes "QCong l" and "A Out B C" and "l A B" and "l A C" shows "B = C" using out_is_len_eq using TarskiLen_def assms(1) assms(2) assms(3) assms(4) by auto lemma ex_anga: assumes "Acute A B C" shows "\ a. (QCongAAcute a \ a A B C)" using acute_distincts anga_exists assms by blast lemma not_null_ang_ang: assumes "QCongAnNull a" shows "QCongA a" using QCongAnNull_def assms by blast lemma not_null_ang_def_equiv: "QCongAnNull a \ (QCongA a \ (\ A B C. (a A B C \ \ B Out A C)))" proof - { assume "QCongAnNull a" have "QCongA a \ (\ A B C. (a A B C \ \ B Out A C))" using QCongAnNull_def \QCongAnNull a\ ex_points_eng by fastforce } { assume "QCongA a \ (\ A B C. (a A B C \ \ B Out A C))" have "QCongAnNull a" by (metis Ang_def QCongAnNull_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms \QCongA a \ (\A B C. a A B C \ \ B Out A C)\ not_conga_is_ang) } thus ?thesis using \QCongAnNull a \ QCongA a \ (\A B C. a A B C \ \ B Out A C)\ by blast qed lemma not_flat_ang_def_equiv: "QCongAnFlat a \ (QCongA a \ (\ A B C. (a A B C \ \ Bet A B C)))" proof - { assume "QCongAnFlat a" then have "QCongA a \ (\ A B C. (a A B C \ \ Bet A B C))" using QCongAnFlat_def ex_points_eng by fastforce } { assume "QCongA a \ (\ A B C. (a A B C \ \ Bet A B C))" have "QCongAnFlat a" proof - obtain pp :: 'p and ppa :: 'p and ppb :: 'p where f1: "QCongA a \ a pp ppa ppb \ \ Bet pp ppa ppb" - using \QCongA a \ (\A B C. a A B C \ \ Bet A B C)\ by moura + using \QCongA a \ (\A B C. a A B C \ \ Bet A B C)\ by blast then have f2: "\p pa pb. pp ppa ppb CongA pb pa p \ \ a pb pa p" by (metis (no_types) Ang_def Tarski_neutral_dimensionless.not_cong_is_ang1 Tarski_neutral_dimensionless_axioms) then have f3: "\p pa pb. (Col pp ppa ppb \ \ a pb pa p) \ \ Bet pb pa p" by (metis (no_types) Col_def Tarski_neutral_dimensionless.ncol_conga_ncol Tarski_neutral_dimensionless_axioms) have f4: "\p pa pb. (\ Bet ppa ppb pp \ \ Bet pb pa p) \ \ a pb pa p" using f2 f1 by (metis Col_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms not_bet_and_out not_out_bet) have f5: "\p pa pb. (\ Bet ppb pp ppa \ \ Bet pb pa p) \ \ a pb pa p" using f2 f1 by (metis Col_def Tarski_neutral_dimensionless.l11_21_a Tarski_neutral_dimensionless_axioms not_bet_and_out not_out_bet) { assume "Bet ppa ppb pp" then have ?thesis using f4 f1 QCongAnFlat_def by blast } moreover { assume "Bet ppb pp ppa" then have ?thesis using f5 f1 QCongAnFlat_def by blast } ultimately show ?thesis using f3 f1 Col_def QCongAnFlat_def by blast qed } thus ?thesis using \QCongAnFlat a \ QCongA a \ (\A B C. a A B C \ \ Bet A B C)\ by blast qed lemma ang_const: assumes "QCongA a" and "A \ B" shows "\ C. a A B C" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(1)) then have "(A0 B0 C0 CongA A0 B0 C0) \ a A0 B0 C0" by (simp add: conga_refl) then have "a A0 B0 C0" using \A0 \ B0 \ C0 \ B0 \ (\X Y Z. A0 B0 C0 CongA X Y Z \ a X Y Z)\ conga_refl by blast then show ?thesis using \A0 \ B0 \ C0 \ B0 \ (\X Y Z. A0 B0 C0 CongA X Y Z \ a X Y Z)\ angle_construction_3 assms(2) by blast qed lemma ang_sym: assumes "QCongA a" and "a A B C" shows "a C B A" proof - obtain A0 B0 C0 where "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(1)) then show ?thesis by (metis Tarski_neutral_dimensionless.ang_conga Tarski_neutral_dimensionless_axioms assms(1) assms(2) conga_left_comm conga_refl not_conga_sym) qed lemma ang_not_null_lg: assumes "QCongA a" and "QCong l" and "a A B C" and "l A B" shows "\ QCongNull l" by (metis QCongNull_def TarskiLen_def ang_distinct assms(1) assms(3) assms(4) cong_reverse_identity not_cong_is_len) lemma ang_distincts: assumes "QCongA a" and "a A B C" shows "A \ B \ C \ B" using ang_distinct assms(1) assms(2) by auto lemma anga_sym: assumes "QCongAAcute a" and "a A B C" shows "a C B A" by (simp add: ang_sym anga_is_ang assms(1) assms(2)) lemma anga_not_null_lg: assumes "QCongAAcute a" and "QCong l" and "a A B C" and "l A B" shows "\ QCongNull l" using ang_not_null_lg anga_is_ang assms(1) assms(2) assms(3) assms(4) by blast lemma anga_distincts: assumes "QCongAAcute a" and "a A B C" shows "A \ B \ C \ B" using anga_distinct assms(1) assms(2) by blast lemma ang_const_o: assumes "\ Col A B P" and "QCongA a" and "QCongAnNull a" and "QCongAnFlat a" shows "\ C. a A B C \ A B OS C P" proof - obtain A0 B0 C0 where P1: "A0 \ B0 \ C0 \ B0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" by (metis QCongA_def assms(2)) then have "a A0 B0 C0" by (simp add: conga_refl) then have T1: "A0 \ C0" using P1 Tarski_neutral_dimensionless.QCongAnNull_def Tarski_neutral_dimensionless_axioms assms(3) out_trivial by fastforce have "A \ B" using assms(1) col_trivial_1 by blast have "A0 \ B0 \ B0 \ C0" using P1 by auto then obtain C where P2: "A0 B0 C0 CongA A B C \ (A B OS C P \ Col A B C)" using angle_construction_2 assms(1) by blast then have "a A B C" by (simp add: P1) have P3: "A B OS C P \ Col A B C" using P2 by simp have P4: "\ A B C. (a A B C \ \ B Out A C)" using assms(3) by (simp add: QCongAnNull_def) have P5: "\ A B C. (a A B C \ \ Bet A B C)" using assms(4) QCongAnFlat_def by auto { assume "Col A B C" have "\ B Out A C" using P4 by (simp add: \a A B C\) have "\ Bet A B C" using P5 by (simp add: \a A B C\) then have "A B OS C P" using \Col A B C\ \\ B Out A C\ l6_4_2 by blast then have "\ C1. (a A B C1 \ A B OS C1 P)" using \a A B C\ by blast } then have "\ C1. (a A B C1 \ A B OS C1 P)" using P3 \a A B C\ by blast then show ?thesis by simp qed lemma anga_const: assumes "QCongAAcute a" and "A \ B" shows "\ C. a A B C" using Tarski_neutral_dimensionless.ang_const Tarski_neutral_dimensionless_axioms anga_is_ang assms(1) assms(2) by fastforce lemma null_anga_null_angaP: "QCongANullAcute a \ IsNullAngaP a" proof - have "QCongANullAcute a \ IsNullAngaP a" using IsNullAngaP_def QCongANullAcute_def ex_points_anga by fastforce moreover have "IsNullAngaP a \ QCongANullAcute a" by (metis IsNullAngaP_def QCongAnNull_def Tarski_neutral_dimensionless.QCongANullAcute_def Tarski_neutral_dimensionless_axioms anga_is_ang not_null_ang_def_equiv) ultimately show ?thesis by blast qed lemma is_null_anga_out: assumes (*"QCongAAcute a" and *) "a A B C" and "QCongANullAcute a" shows "B Out A C" using QCongANullAcute_def assms(1) assms(2) by auto lemma acute_not_bet: assumes "Acute A B C" shows "\ Bet A B C" using acute_col__out assms bet_col not_bet_and_out by blast lemma anga_acute: assumes "QCongAAcute a" and "a A B C" shows "Acute A B C" by (smt Tarski_neutral_dimensionless.QCongAAcute_def Tarski_neutral_dimensionless_axioms acute_conga__acute assms(1) assms(2)) lemma not_null_not_col: assumes "QCongAAcute a" and "\ QCongANullAcute a" and "a A B C" shows "\ Col A B C" proof - have "Acute A B C" using anga_acute assms(1) assms(3) by blast then show ?thesis using Tarski_neutral_dimensionless.IsNullAngaP_def Tarski_neutral_dimensionless_axioms acute_col__out assms(1) assms(2) assms(3) null_anga_null_angaP by blast qed lemma ang_cong_ang: assumes "QCongA a" and "a A B C" and "A B C CongA A' B' C'" shows "a A' B' C'" by (metis QCongA_def assms(1) assms(2) assms(3) not_conga) lemma is_null_ang_out: assumes (*"QCongA a" and *) "a A B C" and "QCongANull a" shows "B Out A C" proof - have "a A B C \ B Out A C" using QCongANull_def assms(2) by auto then show ?thesis by (simp add: assms(1)) qed lemma out_null_ang: assumes "QCongA a" and "a A B C" and "B Out A C" shows "QCongANull a" by (metis QCongANull_def QCongAnNull_def assms(1) assms(2) assms(3) not_null_ang_def_equiv) lemma bet_flat_ang: assumes "QCongA a" and "a A B C" and "Bet A B C" shows "AngFlat a" by (metis AngFlat_def QCongAnFlat_def assms(1) assms(2) assms(3) not_flat_ang_def_equiv) lemma out_null_anga: assumes "QCongAAcute a" and "a A B C" and "B Out A C" shows "QCongANullAcute a" using IsNullAngaP_def assms(1) assms(2) assms(3) null_anga_null_angaP by auto lemma anga_not_flat: assumes "QCongAAcute a" shows "QCongAnFlat a" by (metis (no_types, lifting) Tarski_neutral_dimensionless.QCongAnFlat_def Tarski_neutral_dimensionless.anga_is_ang Tarski_neutral_dimensionless_axioms assms bet_col is_null_anga_out not_bet_and_out not_null_not_col) lemma anga_const_o: assumes "\ Col A B P" and "\ QCongANullAcute a" and "QCongAAcute a" shows "\ C. (a A B C \ A B OS C P)" proof - have "QCongA a" by (simp add: anga_is_ang assms(3)) moreover have "QCongAnNull a" using QCongANullAcute_def assms(2) assms(3) calculation not_null_ang_def_equiv by auto moreover have "QCongAnFlat a" by (simp add: anga_not_flat assms(3)) ultimately show ?thesis by (simp add: ang_const_o assms(1)) qed lemma anga_conga_anga: assumes "QCongAAcute a" and "a A B C" and "A B C CongA A' B' C'" shows "a A' B' C'" using ang_cong_ang anga_is_ang assms(1) assms(2) assms(3) by blast lemma anga_out_anga: assumes "QCongAAcute a" and "a A B C" and "B Out A A'" and "B Out C C'" shows "a A' B C'" proof - have "A B C CongA A' B C'" by (simp add: assms(3) assms(4) l6_6 out2__conga) thus ?thesis using anga_conga_anga assms(1) assms(2) by blast qed lemma out_out_anga: assumes "QCongAAcute a" and "B Out A C" and "B' Out A' C'" and "a A B C" shows "a A' B' C'" proof - have "A B C CongA A' B' C'" by (simp add: assms(2) assms(3) l11_21_b) thus ?thesis using anga_conga_anga assms(1) assms(4) by blast qed lemma is_null_all: assumes "A \ B" and "QCongANullAcute a" shows "a A B A" proof - obtain A0 B0 C0 where "Acute A0 B0 C0 \ (\ X Y Z. (A0 B0 C0 CongA X Y Z \ a X Y Z))" using QCongAAcute_def QCongANullAcute_def assms(2) by auto then have "a A0 B0 C0" using acute_distincts conga_refl by blast thus ?thesis by (smt QCongANullAcute_def assms(1) assms(2) out_out_anga out_trivial) qed lemma anga_col_out: assumes "QCongAAcute a" and "a A B C" and "Col A B C" shows "B Out A C" proof - have "Acute A B C" using anga_acute assms(1) assms(2) by auto then have P1: "Bet A B C \ B Out A C" using acute_not_bet by auto then have "Bet C A B \ B Out A C" using assms(3) l6_4_2 by auto thus ?thesis using P1 assms(3) l6_4_2 by blast qed lemma ang_not_lg_null: assumes "QCong la" and "QCong lc" and "QCongA a" and "la A B" and "lc C B" and "a A B C" shows "\ QCongNull la \ \ QCongNull lc" by (metis ang_not_null_lg ang_sym assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)) lemma anga_not_lg_null: assumes (*"QCong la" and "QCong lc" and*) "QCongAAcute a" and "la A B" and "lc C B" and "a A B C" shows "\ QCongNull la \ \ QCongNull lc" by (metis QCongNull_def anga_not_null_lg anga_sym assms(1) assms(2) assms(3) assms(4)) lemma anga_col_null: assumes "QCongAAcute a" and "a A B C" and "Col A B C" shows "B Out A C \ QCongANullAcute a" using anga_col_out assms(1) assms(2) assms(3) out_null_anga by blast lemma eqA_preserves_ang: assumes "QCongA a" and "a = b" shows "QCongA b" using assms(1) assms(2) by auto lemma eqA_preserves_anga: assumes "QCongAAcute a" and (* "QCongA b" and*) "a = b" shows "QCongAAcute b" using assms(1) assms(2) by auto section "Some postulates of the parallels" lemma euclid_5__original_euclid: assumes "Euclid5" shows "EuclidSParallelPostulate" proof - { fix A B C D P Q R assume P1: "B C OS A D \ SAMS A B C B C D \ A B C B C D SumA P Q R \ \ Bet P Q R" obtain M where P2: "M Midpoint B C" using midpoint_existence by auto obtain D' where P3: "C Midpoint D D'" using symmetric_point_construction by auto obtain E where P4: "M Midpoint D' E" using symmetric_point_construction by auto have P5: "A \ B" using P1 os_distincts by blast have P6: "B \ C" using P1 os_distincts by blast have P7: "C \ D" using P1 os_distincts by blast have P10: "M \ B" using P2 P6 is_midpoint_id by auto have P11: "M \ C" using P2 P6 is_midpoint_id_2 by auto have P13: "C \ D'" using P3 P7 is_midpoint_id_2 by blast have P16: "\ Col B C A" using one_side_not_col123 P1 by blast have "B C OS D A" using P1 one_side_symmetry by blast then have P17: "\ Col B C D" using one_side_not_col123 P1 by blast then have P18: "\ Col M C D" using P2 Col_perm P11 col_transitivity_2 midpoint_col by blast then have P19: "\ Col M C D'" by (metis P13 P3 Col_perm col_transitivity_2 midpoint_col) then have P20: "\ Col D' C B" by (metis Col_perm P13 P17 P3 col_transitivity_2 midpoint_col) then have P21: "\ Col M C E" by (metis P19 P4 bet_col col2__eq col_permutation_4 midpoint_bet midpoint_distinct_2) have P22: "M C D' CongA M B E \ M D' C CongA M E B" using P13 l11_49 by (metis Cong_cases P19 P2 P4 l11_51 l7_13_R1 l7_2 midpoint_cong not_col_distincts) have P23: "Cong C D' B E" using P11 P2 P4 l7_13_R1 l7_2 by blast have P27: "C B TS D D'" by (simp add: P13 P17 P3 bet__ts midpoint_bet not_col_permutation_4) have P28: "A InAngle C B E" proof - have "C B A LeA C B E" proof - have "A B C LeA B C D'" proof - have "Bet D C D'" by (simp add: P3 midpoint_bet) then show ?thesis using P1 P7 P13 sams_chara by (metis sams_left_comm sams_sym) qed moreover have "A B C CongA C B A" using P5 P6 conga_pseudo_refl by auto moreover have "B C D' CongA C B E" by (metis CongA_def Mid_cases P2 P22 P4 P6 symmetry_preserves_conga) ultimately show ?thesis using l11_30 by blast qed moreover have "C B OS E A" proof - have "C B TS E D'" using P2 P20 P4 l7_2 l9_2 mid_two_sides not_col_permutation_1 by blast moreover have "C B TS A D'" using P27 \B C OS D A\ invert_two_sides l9_8_2 by blast ultimately show ?thesis using OS_def by blast qed ultimately show ?thesis using lea_in_angle by simp qed obtain A' where P30: "Bet C A' E \ (A' = B \ B Out A' A)" using P28 InAngle_def by auto { assume "A' = B" then have "Col D' C B" by (metis Col_def P2 P21 P30 P6 col_transitivity_1 midpoint_col) then have "False" by (simp add: P20) then have "\ Y. B Out A Y \ C Out D Y" by auto } { assume P31: "B Out A' A" have "\ I. BetS D' C I \ BetS B A' I" proof - have P32: "BetS B M C" using BetS_def Midpoint_def P10 P11 P2 by auto moreover have "BetS E M D'" using BetS_def Bet_cases P19 P21 P4 midpoint_bet not_col_distincts by fastforce moreover have "BetS C A' E" proof - have P32A: "C \ A'" using P16 P31 out_col by auto { assume "A' = E" then have P33: "B Out A E" using P31 l6_6 by blast then have "A B C B C D SumA D' C D" proof - have "D' C B CongA A B C" proof - have "D' C M CongA E B M" by (simp add: P22 conga_comm) moreover have "C Out D' D'" using P13 out_trivial by auto moreover have "C Out B M" using BetSEq Out_cases P32 bet_out_1 by blast moreover have "B Out A E" using P33 by auto moreover have "B Out C M" using BetSEq Out_def P32 by blast ultimately show ?thesis using l11_10 by blast qed moreover have "D' C B B C D SumA D' C D" by (simp add: P27 l9_2 ts__suma_1) moreover have "B C D CongA B C D" using P6 P7 conga_refl by auto moreover have "D' C D CongA D' C D" using P13 P7 conga_refl by presburger ultimately show ?thesis using conga3_suma__suma by blast qed then have "D' C D CongA P Q R" using P1 suma2__conga by auto then have "Bet P Q R" using Bet_cases P3 bet_conga__bet midpoint_bet by blast then have "False" using P1 by simp } then have "A' \ E" by auto then show ?thesis by (simp add: BetS_def P30 P32A) qed moreover have "\ Col B C D'" by (simp add: P20 not_col_permutation_3) moreover have "Cong B M C M" using Midpoint_def P2 not_cong_1243 by blast moreover have "Cong E M D' M" using Cong_perm Midpoint_def P4 by blast ultimately show ?thesis using euclid_5_def assms by blast qed then obtain Y where P34: "Bet D' C Y \ BetS B A' Y" using BetSEq by blast then have "\ Y. B Out A Y \ C Out D Y" proof - have P35: "B Out A Y" by (metis BetSEq Out_def P31 P34 l6_7) moreover have "C Out D Y" proof - have "D \ C" using P7 by auto moreover have "Y \ C" using P16 P35 l6_6 out_col by blast moreover have "D' \ C" using P13 by auto moreover have "Bet D C D'" by (simp add: P3 midpoint_bet) moreover have "Bet Y C D'" by (simp add: Bet_perm P34) ultimately show ?thesis using l6_2 by blast qed ultimately show ?thesis by auto qed } then have "\ Y. B Out A Y \ C Out D Y" using P30 \A' = B \ \Y. B Out A Y \ C Out D Y\ by blast } then show ?thesis using euclid_s_parallel_postulate_def by blast qed lemma tarski_s_euclid_implies_euclid_5: assumes "TarskiSParallelPostulate" shows "Euclid5" proof - { fix P Q R S T U assume P1: "BetS P T Q \ BetS R T S \ BetS Q U R \ \ Col P Q S \ Cong P T Q T \ Cong R T S T" have P1A: "BetS P T Q" using P1 by simp have P1B: "BetS R T S" using P1 by simp have P1C: "BetS Q U R" using P1 by simp have P1D: "\ Col P Q S" using P1 by simp have P1E: "Cong P T Q T" using P1 by simp have P1F: "Cong R T S T" using P1 by simp obtain V where P2: "P Midpoint R V" using symmetric_point_construction by auto have P3: "Bet V P R" using Mid_cases P2 midpoint_bet by blast then obtain W where P4: "Bet P W Q \ Bet U W V" using inner_pasch using BetSEq P1C by blast { assume "P = W" have "P \ V" by (metis BetSEq Bet_perm Col_def Cong_perm Midpoint_def P1A P1B P1D P1E P1F P2 between_trivial is_midpoint_id_2 l7_9) have "Col P Q S" proof - have f1: "Col V P R" by (meson Col_def P3) have f2: "Col U R Q" by (simp add: BetSEq Col_def P1) have f3: "Bet P T Q" using BetSEq P1 by fastforce have f4: "R = P \ Col V P U" by (metis (no_types) Col_def P4 \P = W\ \P \ V\ l6_16_1) have f5: "Col Q P T" using f3 by (meson Col_def) have f6: "Col T Q P" using f3 by (meson Col_def) have f7: "Col P T Q" using f3 by (meson Col_def) have f8: "Col P Q P" using Col_def P4 \P = W\ by blast have "Col R T S" by (meson BetSEq Col_def P1) then have "T = P \ Q = P" using f8 f7 f6 f5 f4 f2 f1 by (metis (no_types) BetSEq P1 \P \ V\ colx l6_16_1) then show ?thesis by (metis BetSEq P1) qed then have "False" by (simp add: P1D) } then have P5: "P \ W" by auto have "Bet V W U" using Bet_cases P4 by auto then obtain X Y where P7: "Bet P V X \ Bet P U Y \ Bet X Q Y" using assms(1) P1 P4 P5 tarski_s_parallel_postulate_def by blast have "Q S Par P R" proof - have "Q \ S" using P1D col_trivial_2 by auto moreover have "T Midpoint Q P" using BetSEq P1A P1E l7_2 midpoint_def not_cong_1243 by blast moreover have "T Midpoint S R" using BetSEq P1B P1F l7_2 midpoint_def not_cong_1243 by blast ultimately show ?thesis using l12_17 by auto qed then have P9: "Q S ParStrict P R" using P1D Par_def par_strict_symmetry par_symmetry by blast have P10: "Q S TS P Y" proof - have P10A: "P \ R" using P9 par_strict_distinct by auto then have P11: "P \ X" by (metis P2 P7 bet_neq12__neq midpoint_not_midpoint) have P12: "\ Col X Q S" proof - have "Q S ParStrict P R" by (simp add: P9) then have "Col P R X" by (metis P2 P3 P7 bet_col between_symmetry midpoint_not_midpoint not_col_permutation_4 outer_transitivity_between) then have "P X ParStrict Q S" using P9 Par_strict_perm P11 par_strict_col_par_strict by blast then show ?thesis using par_strict_not_col_2 by auto qed { assume W1: "Col Y Q S" have W2: "Q = Y" by (metis P12 P7 W1 bet_col bet_col1 colx) then have "\ Col Q P R" using P9 W1 par_not_col by auto then have W3: "Q = U" by (smt BetS_def Col_def P1C P7 W2 col_transitivity_2) then have "False" using BetS_def P1C by auto } then have "\ Col Y Q S" by auto then have "Q S TS X Y" by (metis P7 P12 bet__ts not_col_distincts not_col_permutation_1) moreover have "Q S OS X P" proof - have "P \ V" using P10A P2 is_midpoint_id_2 by blast then have "Q S ParStrict P X" by (meson Bet_perm P3 P7 P9 P11 bet_col not_col_permutation_4 par_strict_col_par_strict) then have "Q S ParStrict X P" by (simp add: par_strict_right_comm) then show ?thesis by (simp add: l12_6) qed ultimately show ?thesis using l9_8_2 by auto qed then obtain I where W4: "Col I Q S \ Bet P I Y" using TS_def by blast have "\ I. (BetS S Q I \ BetS P U I)" proof - have "BetS P U I" proof - have "P \ Y" using P10 not_two_sides_id by auto have W4A: "Bet P U I" proof - have W5: "Col P U I" using P7 W4 bet_col1 by auto { assume W6: "Bet U I P" have W7: "Q S OS P U" proof - have "Q S OS R U" proof - have "\ Col Q S R" using P9 par_strict_not_col_4 by auto moreover have "Q Out R U" using BetSEq Out_def P1C by blast ultimately show ?thesis by (simp add: out_one_side) qed moreover have "Q S OS P R" by (simp add: P9 l12_6) ultimately show ?thesis using one_side_transitivity by blast qed have W8: "I Out P U \ \ Col Q S P" by (simp add: P1D not_col_permutation_1) have "False" proof - have "I Out U P" using W4 W6 W7 between_symmetry one_side_chara by blast then show ?thesis using W6 not_bet_and_out by blast qed } { assume V1: "Bet I P U" have "P R OS I U" proof - have "P R OS I Q" proof - { assume "Q = I" then have "Col P Q S" by (metis BetSEq Col_def P1C P7 P9 V1 W4 between_equality outer_transitivity_between par_not_col) then have "False" using P1D by blast } then have "Q \ I" by blast moreover have "P R ParStrict Q S" using P9 par_strict_symmetry by blast moreover have "Col Q S I" using Col_cases W4 by blast ultimately show ?thesis using one_side_symmetry par_strict_all_one_side by blast qed moreover have "P R OS Q U" proof - have "Q S ParStrict P R" using P9 by blast have "R Out Q U \ \ Col P R Q" by (metis BetSEq Bet_cases Out_def P1C calculation col124__nos) then show ?thesis by (metis P7 V1 W4 \Bet U I P \ False\ between_equality col_permutation_2 not_bet_distincts out_col outer_transitivity_between) qed ultimately show ?thesis using one_side_transitivity by blast qed then have V2: "P Out I U" using P7 W4 bet2__out os_distincts by blast then have "Col P I U" using V1 not_bet_and_out by blast then have "False" using V1 V2 not_bet_and_out by blast } then moreover have "\ (Bet U I P \ Bet I P U)" using \Bet U I P \ False\ by auto ultimately show ?thesis using Col_def W5 by blast qed { assume "P = U" then have "Col P R Q" using BetSEq Col_def P1C by blast then have "False" using P9 par_strict_not_col_3 by blast } then have V6: "P \ U" by auto { assume "U = I" have "Q = U" proof - have f1: "BetS Q I R" using P1C \U = I\ by blast then have f2: "Col Q I R" using BetSEq Col_def by blast have f3: "Col I R Q" using f1 by (simp add: BetSEq Col_def) { assume "R \ Q" moreover { assume "(R \ Q \ R \ I) \ \ Col I Q R" moreover { assume "\p. (R \ Q \ \ Col I p I) \ Col Q I p" then have "I = Q" using f1 by (metis (no_types) BetSEq Col_def col_transitivity_2) } ultimately have "(\p pa. ((pa \ I \ \ Col pa p R) \ Col Q I pa) \ Col I pa p) \ I = Q" using f3 f2 by (metis (no_types) col_transitivity_2) } ultimately have "(\p pa. ((pa \ I \ \ Col pa p R) \ Col Q I pa) \ Col I pa p) \ I = Q" using f1 by (metis (no_types) BetSEq P9 W4 col_transitivity_2 par_strict_not_col_4) } then show ?thesis using f2 by (metis P9 W4 \U = I\ col_transitivity_2 par_strict_not_col_4) qed then have "False" using BetSEq P1C by blast } then have "U \ I" by auto then show ?thesis by (simp add: W4A V6 BetS_def) qed moreover have "BetS S Q I" proof - have "Q R TS S I" proof - have "Q R TS P I" proof - have "\ Col P Q R" using P9 col_permutation_5 par_strict_not_col_3 by blast moreover have "\ Col I Q R" proof - { assume "Col I Q R" then have "Col Q S R" proof - have f1: "\p pa pb. Col p pa pb \ \ BetS pb p pa" by (meson BetSEq Col_def) then have f2: "Col U I P" using \BetS P U I\ by blast have f3: "Col I P U" by (simp add: BetSEq Col_def \BetS P U I\) have f4: "\p. (U = Q \ Col Q p R) \ \ Col Q U p" by (metis BetSEq Col_def P1C col_transitivity_1) { assume "P \ Q" moreover { assume "(P \ Q \ U \ Q) \ Col Q P Q" then have "(P \ Q \ U \ Q) \ \ Col Q P R" using Col_cases \\ Col P Q R\ by blast moreover { assume "\p. ((U \ Q \ P \ Q) \ \ Col Q p P) \ Col Q P p" then have "U \ Q \ \ Col Q P P" by (metis col_transitivity_1) then have "\ Col U Q P" using col_transitivity_2 by blast } ultimately have "\ Col U Q P \ I \ Q" using f4 f3 by blast } ultimately have "I \ Q" using f2 f1 by (metis BetSEq P1C col_transitivity_1 col_transitivity_2) } then have "I \ Q" using BetSEq \BetS P U I\ by blast then show ?thesis by (simp add: W4 \Col I Q R\ col_transitivity_2) qed then have "False" using P9 par_strict_not_col_4 by blast } then show ?thesis by blast qed moreover have "Col U Q R" using BetSEq Bet_cases Col_def P1C by blast moreover have "Bet P U I" by (simp add: BetSEq \BetS P U I\) ultimately show ?thesis using TS_def by blast qed moreover have "Q R OS P S" proof - have "Q R Par P S" proof - have "Q \ R" using BetSEq P1 by blast moreover have "T Midpoint Q P" using BetSEq Bet_cases P1A P1E cong_3421 midpoint_def by blast moreover have "T Midpoint R S" using BetSEq P1B P1F midpoint_def not_cong_1243 by blast ultimately show ?thesis using l12_17 by blast qed then have "Q R ParStrict P S" by (simp add: P1D Par_def not_col_permutation_4) then show ?thesis using l12_6 by blast qed ultimately show ?thesis using l9_8_2 by blast qed then show ?thesis by (metis BetS_def W4 col_two_sides_bet not_col_permutation_2 ts_distincts) qed ultimately show ?thesis by auto qed } then show ?thesis using euclid_5_def by blast qed lemma tarski_s_implies_euclid_s_parallel_postulate: assumes "TarskiSParallelPostulate" shows "EuclidSParallelPostulate" by (simp add: assms euclid_5__original_euclid tarski_s_euclid_implies_euclid_5) theorem tarski_s_euclid_implies_playfair_s_postulate: assumes "TarskiSParallelPostulate" shows "PlayfairSPostulate" proof - { fix A1 A2 B1 B2 P C1 C2 assume P1: "\ Col P A1 A2 \ A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2" have P1A: "\ Col P A1 A2" by (simp add: P1) have P2: "A1 A2 Par B1 B2" by (simp add: P1) have P3: "Col P B1 B2" by (simp add: P1) have P4: "A1 A2 Par C1 C2" by (simp add: P1) have P5: "Col P C1 C2" by (simp add: P1) have P6: "A1 A2 ParStrict B1 B2" proof - have "A1 A2 Par B1 B2" by (simp add: P1) moreover have "Col B1 B2 P" using P3 not_col_permutation_2 by blast moreover have "\ Col A1 A2 P" by (simp add: P1A not_col_permutation_1) ultimately show ?thesis using par_not_col_strict by auto qed have P7: "A1 A2 ParStrict C1 C2" proof - have "A1 A2 Par C1 C2" by (simp add: P1) moreover have "Col C1 C2 P" using Col_cases P1 by blast moreover have "\ Col A1 A2 P" by (simp add: P1A not_col_permutation_1) ultimately show ?thesis using par_not_col_strict by auto qed { assume "\ Col C1 B1 B2 \ \ Col C2 B1 B2" have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" proof - have T2: "Coplanar A1 A2 P A1" using ncop_distincts by auto have T3: "Coplanar A1 A2 B1 B2" by (simp add: P1 par__coplanar) have T4: "Coplanar A1 A2 C1 C2" by (simp add: P7 pars__coplanar) have T5: "Coplanar A1 A2 P B1" using P1 col_trivial_2 ncop_distincts par__coplanar par_col2_par_bis by blast then have T6: "Coplanar A1 A2 P B2" using P3 T3 col_cop__cop by blast have T7: "Coplanar A1 A2 P C1" using P1 T4 col_cop__cop coplanar_perm_1 not_col_permutation_2 par_distincts by blast then have T8: "Coplanar A1 A2 P C2" using P5 T4 col_cop__cop by blast { assume "\ Col C1 B1 B2" moreover have "C1 \ C2" using P1 par_neq2 by auto moreover have "Col B1 B2 P" using P1 not_col_permutation_2 by blast moreover have "Col C1 C2 P" using Col_cases P5 by auto moreover have "\ Col B1 B2 C1" using Col_cases calculation(1) by auto moreover have "\ Col B1 B2 A1" using P6 par_strict_not_col_3 by auto moreover have "Coplanar B1 B2 C1 A1" using Col_cases P1A T5 T2 T6 T7 coplanar_pseudo_trans by blast ultimately have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" using cop_not_par_other_side by blast } { assume "\ Col C2 B1 B2" moreover have "C2 \ C1" using P1 par_neq2 by blast moreover have "Col B1 B2 P" using Col_cases P3 by auto moreover have "Col C2 C1 P" using Col_cases P5 by auto moreover have "\ Col B1 B2 C2" by (simp add: calculation(1) not_col_permutation_1) moreover have "\ Col B1 B2 A1" using P6 par_strict_not_col_3 by auto moreover have "Coplanar B1 B2 C2 A1" using Col_cases P1A T2 T5 T6 T8 coplanar_pseudo_trans by blast ultimately have "\ C'. Col C1 C2 C' \ B1 B2 TS A1 C'" using cop_not_par_other_side by (meson not_col_permutation_4) } then show ?thesis using \\ Col C1 B1 B2 \ \C'. Col C1 C2 C' \ B1 B2 TS A1 C'\ \\ Col C1 B1 B2 \ \ Col C2 B1 B2\ by blast qed then obtain C' where W1: "Col C1 C2 C' \ B1 B2 TS A1 C'" by auto then have W2: "\ Col A1 B1 B2" using TS_def by blast obtain B where W3: "Col B B1 B2 \ Bet A1 B C'" using TS_def W1 by blast obtain C where W4: "P Midpoint C' C" using symmetric_point_construction by blast then have W4A: "Bet A1 B C' \ Bet C P C'" using Mid_cases W3 midpoint_bet by blast then obtain D where W5: "Bet B D C \ Bet P D A1" using inner_pasch by blast have W6: "C' \ P" using P3 TS_def W1 by blast then have "A1 A2 Par C' P" by (meson P1 W1 not_col_permutation_2 par_col2_par) have W9: "A1 A2 ParStrict C' P" using Col_cases P5 P7 W1 W6 par_strict_col2_par_strict by blast then have W10: "B \ P" by (metis W6 W4A bet_out_1 out_col par_strict_not_col_3) have W11: "P \ C" using W6 W4 is_midpoint_id_2 by blast { assume "P = D" then have "False" by (metis Col_def P3 W1 W3 W4A W5 W10 W11 col_trivial_2 colx l9_18_R1) } then have "P \ D" by auto then obtain X Y where W12: "Bet P B X \ Bet P C Y \ Bet X A1 Y" using W5 assms tarski_s_parallel_postulate_def by blast then have "P \ X" using W10 bet_neq12__neq by auto then have "A1 A2 ParStrict P X" by (metis Col_cases P3 P6 W10 W12 W3 bet_col colx par_strict_col2_par_strict) then have W15: "A1 A2 OS P X" by (simp add: l12_6) have "P \ Y" using W11 W12 between_identity by blast then have "A1 A2 ParStrict P Y" by (metis Col_def W11 W12 W4A W9 col_trivial_2 par_strict_col2_par_strict) then have W16: "A1 A2 OS P Y" using l12_6 by auto have "Col A1 X Y" by (simp add: W12 bet_col col_permutation_4) then have "A1 Out X Y" using col_one_side_out W15 W16 using one_side_symmetry one_side_transitivity by blast then have "False" using W12 not_bet_and_out by blast } then have "Col C1 B1 B2 \ Col C2 B1 B2" by auto } { fix A1 A2 B1 B2 P C1 C2 assume P1: "Col P A1 A2 \ A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2" have "Col C1 B1 B2" by (smt P1 l9_10 not_col_permutation_3 not_strict_par2 par_col2_par par_comm par_id_5 par_symmetry ts_distincts) moreover have "Col C2 B1 B2" by (smt P1 l9_10 not_col_permutation_3 not_strict_par2 par_col2_par par_id_5 par_left_comm par_symmetry ts_distincts) ultimately have "Col C1 B1 B2 \ Col C2 B1 B2" by auto } then show ?thesis using playfair_s_postulate_def by (metis \\P C2 C1 B2 B1 A2 A1. \ Col P A1 A2 \ A1 A2 Par B1 B2 \ Col P B1 B2 \ A1 A2 Par C1 C2 \ Col P C1 C2 \ Col C1 B1 B2 \ Col C2 B1 B2\) qed end end diff --git a/thys/Lp/Functional_Spaces.thy b/thys/Lp/Functional_Spaces.thy --- a/thys/Lp/Functional_Spaces.thy +++ b/thys/Lp/Functional_Spaces.thy @@ -1,1249 +1,1249 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) theory Functional_Spaces imports "HOL-Analysis.Analysis" "HOL-Library.Function_Algebras" Ergodic_Theory.SG_Library_Complement begin section \Functions as a real vector space\ text \Many functional spaces are spaces of functions. To be able to use the following framework, spaces of functions thus need to be endowed with a vector space structure, coming from pointwise addition and multiplication. Some instantiations for \verb+fun+ are already given in \verb+Function_Algebras.thy+ and \verb+Lattices.thy+, we add several.\ text \\verb+minus_fun+ is already defined, in \verb+Lattices.thy+, but under the strange name \verb+fun_Compl_def+. We restate the definition so that \verb+unfolding minus_fun_def+ works. Same thing for \verb+minus_fun_def+. A better solution would be to have a coherent naming scheme in \verb+Lattices.thy+.\ lemmas uminus_fun_def = fun_Compl_def lemmas minus_fun_def = fun_diff_def lemma fun_sum_apply: fixes u::"'i \ 'a \ ('b::comm_monoid_add)" shows "(sum u I) x = sum (\i. u i x) I" by (induction I rule: infinite_finite_induct, auto) instantiation "fun" :: (type, real_vector) real_vector begin definition scaleR_fun::"real \ ('a \ 'b) \ 'a \ 'b" where "scaleR_fun = (\c f. (\x. c *\<^sub>R f x))" lemma scaleR_apply [simp, code]: "(c *\<^sub>R f) x = c *\<^sub>R (f x)" by (simp add: scaleR_fun_def) instance by (standard, auto simp add: scaleR_add_right scaleR_add_left) end lemmas divideR_apply = scaleR_apply lemma [measurable]: "0 \ borel_measurable M" unfolding zero_fun_def by auto lemma borel_measurable_const_scaleR' [measurable (raw)]: "(f::('a \ 'b::real_normed_vector)) \ borel_measurable M \ c *\<^sub>R f \ borel_measurable M" unfolding scaleR_fun_def using borel_measurable_add by auto lemma borel_measurable_add'[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "f + g \ borel_measurable M" unfolding plus_fun_def using assms by auto lemma borel_measurable_uminus'[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" shows "-f \ borel_measurable M" unfolding fun_Compl_def using assms by auto lemma borel_measurable_diff'[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "f - g \ borel_measurable M" unfolding fun_diff_def using assms by auto lemma borel_measurable_sum'[measurable (raw)]: fixes f::"'i \ 'a \ 'b::{second_countable_topology, real_normed_vector}" assumes "\i. i \ I \ f i \ borel_measurable M" shows "(\i\I. f i) \ borel_measurable M" using borel_measurable_sum[of I f, OF assms] unfolding fun_sum_apply[symmetric] by simp lemma zero_applied_to [simp]: "(0::('a \ ('b::real_vector))) x = 0" unfolding zero_fun_def by simp section \Quasinorms on function spaces\ text \A central feature of modern analysis is the use of various functional spaces, and of results of functional analysis on them. Think for instance of $L^p$ spaces, of Sobolev or Besov spaces, or variations around them. Here are several relevant facts about this point of view: \begin{itemize} \item These spaces typically depend on one or several parameters. This makes it difficult to play with type classes in a system without dependent types. \item The $L^p$ spaces are not spaces of functions (their elements are equivalence classes of functions, where two functions are identified if they coincide almost everywhere). However, in usual analysis proofs, one takes a definite representative and works with it, never going to the equivalence class point of view (which only becomes relevant when one wants to use the fact that one has a Banach space at our disposal, to apply functional analytic tools). \item It is important to describe how the spaces are related to each other, with respect to inclusions or compact inclusions. For instance, one of the most important theorems in analysis is Sobolev embedding theorem, describing when one Sobolev space is included in another one. One also needs to be able to take intersections or sums of Banach spaces, for instance to develop interpolation theory. \item Some other spaces play an important role in analysis, for instance the weak $L^1$ space. This space only has a quasi-norm (i.e., its norm satisfies the triangular inequality up to a fixed multiplicative constant). A general enough setting should also encompass this kind of space. (One could argue that one should also consider more general topologies such as Frechet spaces, to deal with Gevrey or analytic functions. This is true, but considering quasi-norms already gives a wealth of applications). \end{itemize} Given these points, it seems that the most effective way of formalizing this kind of question in Isabelle/HOL is to think of such a functional space not as an abstract space or type, but as a subset of the space of all functions or of all distributions. Functions that do not belong to the functional space under consideration will then have infinite norm. Then inclusions, intersections, and so on, become trivial to implement. Since the same object contains both the information about the norm and the space where the norm is finite, it conforms to the customary habit in mathematics of identifying the two of them, talking for instance about the $L^p$ space and the $L^p$ norm. All in all, this approach seems quite promising for ``real life analysis''. \ subsection \Definition of quasinorms\ typedef (overloaded) ('a::real_vector) quasinorm = "{(C::real, N::('a \ ennreal)). (C \ 1) \ (\ x c. N (c *\<^sub>R x) = ennreal \c\ * N(x)) \ (\ x y. N(x+y) \ C * N x + C * N y)}" morphisms Rep_quasinorm quasinorm_of proof show "(1,(\x. 0)) \ {(C::real, N::('a \ ennreal)). (C \ 1) \ (\ x c. N (c *\<^sub>R x) = ennreal \c\ * N x) \ (\ x y. N (x+y) \ C * N x + C * N y)}" by auto qed definition eNorm::"'a quasinorm \ ('a::real_vector) \ ennreal" where "eNorm N x = (snd (Rep_quasinorm N)) x" definition defect::"('a::real_vector) quasinorm \ real" where "defect N = fst (Rep_quasinorm N)" lemma eNorm_triangular_ineq: "eNorm N (x + y) \ defect N * eNorm N x + defect N * eNorm N y" unfolding eNorm_def defect_def using Rep_quasinorm[of N] by auto lemma defect_ge_1: "defect N \ 1" unfolding defect_def using Rep_quasinorm[of N] by auto lemma eNorm_cmult: "eNorm N (c *\<^sub>R x) = ennreal \c\ * eNorm N x" unfolding eNorm_def using Rep_quasinorm[of N] by auto lemma eNorm_zero [simp]: "eNorm N 0 = 0" by (metis eNorm_cmult abs_zero ennreal_0 mult_zero_left real_vector.scale_zero_left) lemma eNorm_uminus [simp]: "eNorm N (-x) = eNorm N x" using eNorm_cmult[of N "-1" x] by auto lemma eNorm_sum: "eNorm N (\i\{.. (\i\{..v. eNorm N (\i\{..n}. v i) \ (\i\{.. 'a" define w where "w = (\i. if i = n then v n + v (Suc n) else v i)" have "(\i\{..Suc n}. v i) = (\i\{..i\{..i\{..Suc n}. v i) = (\i\{..n}. w i)" by (metis lessThan_Suc_atMost sum.lessThan_Suc) then have "eNorm N (\i\{..Suc n}. v i) = eNorm N (\i\{..n}. w i)" by simp also have "... \ (\i\{..i\{.. (\i\{..i\{..i\{..i\{..Suc n}. v i) \ (\ii\{.. (\i\{.. (\i\{..i\{..i\{.. (\in = Suc m\ by auto qed text \Quasinorms are often defined by taking a meaningful formula on a vector subspace, and then extending by infinity elsewhere. Let us show that this results in a quasinorm on the whole space.\ definition quasinorm_on::"('a set) \ real \ (('a::real_vector) \ ennreal) \ bool" where "quasinorm_on F C N = ( (\x y. (x \ F \ y \ F) \ (x + y \ F) \ N (x+y) \ C * N x + C * N y) \ (\c x. x \ F \ c *\<^sub>R x \ F \ N(c *\<^sub>R x) = \c\ * N x) \ C \ 1 \ 0 \ F)" lemma quasinorm_of: fixes N::"('a::real_vector) \ ennreal" and C::real assumes "quasinorm_on UNIV C N" shows "eNorm (quasinorm_of (C,N)) x = N x" "defect (quasinorm_of (C,N)) = C" using assms unfolding eNorm_def defect_def quasinorm_on_def by (auto simp add: quasinorm_of_inverse) lemma quasinorm_onI: fixes N::"('a::real_vector) \ ennreal" and C::real and F::"'a set" assumes "\x y. x \ F \ y \ F \ x + y \ F" "\x y. x \ F \ y \ F \ N (x + y) \ C * N x + C * N y" "\c x. c \ 0 \ x \ F \ c *\<^sub>R x \ F" "\c x. c \ 0 \ x \ F \ N (c *\<^sub>R x) \ ennreal \c\ * N x" "0 \ F" "N(0) = 0" "C \ 1" shows "quasinorm_on F C N" proof - have "N(c *\<^sub>R x) = ennreal \c\ * N x" if "x \ F" for c x proof (cases "c = 0") case True then show ?thesis using \N 0 = 0\ by auto next case False have "N((1/c) *\<^sub>R (c *\<^sub>R x)) \ ennreal (abs (1/c)) * N (c *\<^sub>R x)" apply (rule \\c x. c \ 0 \ x \ F \ N(c *\<^sub>R x) \ ennreal \c\ * N x\) using False assms that by auto then have "N x \ ennreal (abs (1/c)) * N (c *\<^sub>R x)" using False by auto then have "ennreal \c\ * N x \ ennreal \c\ * ennreal (abs (1/c)) * N (c *\<^sub>R x)" by (simp add: mult.assoc mult_left_mono) also have "... = N (c *\<^sub>R x)" using ennreal_mult' abs_mult False by (metis abs_ge_zero abs_one comm_monoid_mult_class.mult_1 ennreal_1 eq_divide_eq_1 field_class.field_divide_inverse) finally show ?thesis using \\c x. c \ 0 \ x \ F \ N(c *\<^sub>R x) \ ennreal \c\ * N x\[OF False \x \ F\] by auto qed then show ?thesis unfolding quasinorm_on_def using assms by (auto, metis real_vector.scale_zero_left) qed lemma extend_quasinorm: assumes "quasinorm_on F C N" shows "quasinorm_on UNIV C (\x. if x \ F then N x else \)" proof - have *: "(if x + y \ F then N (x + y) else \) \ ennreal C * (if x \ F then N x else \) + ennreal C * (if y \ F then N y else \)" for x y proof (cases "x \ F \ y \ F") case True then show ?thesis using assms unfolding quasinorm_on_def by auto next case False moreover have "C \ 1" using assms unfolding quasinorm_on_def by auto ultimately have *: "ennreal C * (if x \ F then N x else \) + ennreal C * (if y \ F then N y else \) = \" using ennreal_mult_eq_top_iff by auto show ?thesis by (simp add: *) qed show ?thesis apply (rule quasinorm_onI) using assms * unfolding quasinorm_on_def apply (auto simp add: ennreal_top_mult mult.commute) by (metis abs_zero ennreal_0 mult_zero_right real_vector.scale_zero_right) qed subsection \The space and the zero space of a quasinorm\ text \The space of a quasinorm is the vector subspace where it is meaningful, i.e., finite.\ definition space\<^sub>N::"('a::real_vector) quasinorm \ 'a set" where "space\<^sub>N N = {f. eNorm N f < \}" lemma spaceN_iff: "x \ space\<^sub>N N \ eNorm N x < \" unfolding space\<^sub>N_def by simp lemma spaceN_cmult [simp]: assumes "x \ space\<^sub>N N" shows "c *\<^sub>R x \ space\<^sub>N N" using assms unfolding spaceN_iff using eNorm_cmult[of N c x] by (simp add: ennreal_mult_less_top) lemma spaceN_add [simp]: assumes "x \ space\<^sub>N N" "y \ space\<^sub>N N" shows "x + y \ space\<^sub>N N" proof - have "eNorm N x < \" "eNorm N y < \" using assms unfolding space\<^sub>N_def by auto then have "defect N * eNorm N x + defect N * eNorm N y < \" by (simp add: ennreal_mult_less_top) then show ?thesis unfolding space\<^sub>N_def using eNorm_triangular_ineq[of N x y] le_less_trans by blast qed lemma spaceN_diff [simp]: assumes "x \ space\<^sub>N N" "y \ space\<^sub>N N" shows "x - y \ space\<^sub>N N" using spaceN_add[OF assms(1) spaceN_cmult[OF assms(2), of "-1"]] by auto lemma spaceN_contains_zero [simp]: "0 \ space\<^sub>N N" unfolding space\<^sub>N_def by auto lemma spaceN_sum [simp]: assumes "\i. i \ I \ x i \ space\<^sub>N N" shows "(\i\I. x i) \ space\<^sub>N N" using assms by (induction I rule: infinite_finite_induct, auto) text \The zero space of a quasinorm is the vector subspace of vectors with zero norm. If one wants to get a true metric space, one should quotient the space by the zero space.\ definition zero_space\<^sub>N::"('a::real_vector) quasinorm \ 'a set" where "zero_space\<^sub>N N = {f. eNorm N f = 0}" lemma zero_spaceN_iff: "x \ zero_space\<^sub>N N \ eNorm N x = 0" unfolding zero_space\<^sub>N_def by simp lemma zero_spaceN_cmult: assumes "x \ zero_space\<^sub>N N" shows "c *\<^sub>R x \ zero_space\<^sub>N N" using assms unfolding zero_spaceN_iff using eNorm_cmult[of N c x] by simp lemma zero_spaceN_add: assumes "x \ zero_space\<^sub>N N" "y \ zero_space\<^sub>N N" shows "x + y \ zero_space\<^sub>N N" proof - have "eNorm N x = 0" "eNorm N y = 0" using assms unfolding zero_space\<^sub>N_def by auto then have "defect N * eNorm N x + defect N * eNorm N y = 0" by auto then show ?thesis unfolding zero_spaceN_iff using eNorm_triangular_ineq[of N x y] by auto qed lemma zero_spaceN_diff: assumes "x \ zero_space\<^sub>N N" "y \ zero_space\<^sub>N N" shows "x - y \ zero_space\<^sub>N N" using zero_spaceN_add[OF assms(1) zero_spaceN_cmult[OF assms(2), of "-1"]] by auto lemma zero_spaceN_subset_spaceN: "zero_space\<^sub>N N \ space\<^sub>N N" by (simp add: spaceN_iff zero_spaceN_iff subset_eq) text \On the space, the norms are finite. Hence, it is much more convenient to work there with a real valued version of the norm. We use Norm with a capital N to distinguish it from norms in a (type class) banach space.\ definition Norm::"'a quasinorm \ ('a::real_vector) \ real" where "Norm N x = enn2real (eNorm N x)" lemma Norm_nonneg [simp]: "Norm N x \ 0" unfolding Norm_def by auto lemma Norm_zero [simp]: "Norm N 0 = 0" unfolding Norm_def by auto lemma Norm_uminus [simp]: "Norm N (-x) = Norm N x" unfolding Norm_def by auto lemma eNorm_Norm: assumes "x \ space\<^sub>N N" shows "eNorm N x = ennreal (Norm N x)" using assms unfolding Norm_def by (simp add: spaceN_iff) lemma eNorm_Norm': assumes "x \ space\<^sub>N N" shows "Norm N x = 0" using assms unfolding Norm_def apply (auto simp add: spaceN_iff) using top.not_eq_extremum by fastforce lemma Norm_cmult: "Norm N (c *\<^sub>R x) = abs c * Norm N x" unfolding Norm_def unfolding eNorm_cmult by (simp add: enn2real_mult) lemma Norm_triangular_ineq: assumes "x \ space\<^sub>N N" shows "Norm N (x + y) \ defect N * Norm N x + defect N * Norm N y" proof (cases "y \ space\<^sub>N N") case True have *: "defect N * Norm N x + defect N * Norm N y \ 1 * 0 + 1 * 0" apply (rule add_mono) by (rule mult_mono'[OF defect_ge_1 Norm_nonneg], simp, simp)+ have "ennreal (Norm N (x + y)) = eNorm N (x+y)" using eNorm_Norm[OF spaceN_add[OF assms True]] by auto also have "... \ defect N * eNorm N x + defect N * eNorm N y" using eNorm_triangular_ineq[of N x y] by auto also have "... = defect N * ennreal(Norm N x) + defect N * ennreal(Norm N y)" using eNorm_Norm assms True by metis also have "... = ennreal(defect N * Norm N x + defect N * Norm N y)" using ennreal_mult ennreal_plus Norm_nonneg defect_ge_1 by (metis (no_types, opaque_lifting) ennreal_eq_0_iff less_le ennreal_ge_1 ennreal_mult' le_less_linear not_one_le_zero semiring_normalization_rules(34)) finally show ?thesis apply (subst ennreal_le_iff[symmetric]) using * by auto next case False have "x + y \ space\<^sub>N N" proof (rule ccontr) assume "\ (x + y \ space\<^sub>N N)" then have "x + y \ space\<^sub>N N" by simp have "y \ space\<^sub>N N" using spaceN_diff[OF \x + y \ space\<^sub>N N\ assms] by auto then show False using False by simp qed then have "Norm N (x+y) = 0" unfolding Norm_def using spaceN_iff top.not_eq_extremum by force moreover have "defect N * Norm N x + defect N * Norm N y \ 1 * 0 + 1 * 0" apply (rule add_mono) by (rule mult_mono'[OF defect_ge_1 Norm_nonneg], simp, simp)+ ultimately show ?thesis by simp qed lemma Norm_triangular_ineq_diff: assumes "x \ space\<^sub>N N" shows "Norm N (x - y) \ defect N * Norm N x + defect N * Norm N y" using Norm_triangular_ineq[OF assms, of "-y"] by auto lemma zero_spaceN_iff': "x \ zero_space\<^sub>N N \ (x \ space\<^sub>N N \ Norm N x = 0)" using eNorm_Norm unfolding space\<^sub>N_def zero_space\<^sub>N_def by (auto simp add: Norm_def, fastforce) lemma Norm_sum: assumes "\i. i < n \ u i \ space\<^sub>N N" shows "Norm N (\i\{.. (\i\{.. defect N * defect N ^ i * Norm N (u i)" for i by (meson Norm_nonneg defect_ge_1 dual_order.trans linear mult_nonneg_nonneg not_one_le_zero zero_le_power) have "ennreal (Norm N (\i\{..i\{.. (\i\{..i\{..i\{..i\{..i\{.. ennreal (\i\{..An example: the ambient norm in a normed vector space\ definition N_of_norm::"'a::real_normed_vector quasinorm" where "N_of_norm = quasinorm_of (1, \f. norm f)" lemma N_of_norm: "eNorm N_of_norm f = ennreal (norm f)" "Norm N_of_norm f = norm f" "defect (N_of_norm) = 1" proof - have *: "quasinorm_on UNIV 1 (\f. norm f)" by (rule quasinorm_onI, auto simp add: ennreal_mult', metis ennreal_leI ennreal_plus norm_imp_pos_and_ge norm_triangle_ineq) show "eNorm N_of_norm f = ennreal (norm f)" "defect (N_of_norm) = 1" unfolding N_of_norm_def using quasinorm_of[OF *] by auto then show "Norm N_of_norm f = norm f" unfolding Norm_def by auto qed lemma N_of_norm_space [simp]: "space\<^sub>N N_of_norm = UNIV" unfolding space\<^sub>N_def apply auto unfolding N_of_norm(1) by auto lemma N_of_norm_zero_space [simp]: "zero_space\<^sub>N N_of_norm = {0}" unfolding zero_space\<^sub>N_def apply auto unfolding N_of_norm(1) by auto subsection \An example: the space of bounded continuous functions from a topological space to a normed real vector space\ text \The Banach space of bounded continuous functions is defined in \verb+Bounded_Continuous_Function.thy+, as a type \verb+bcontfun+. We import very quickly the results proved in this file to the current framework.\ definition bcontfun\<^sub>N::"('a::topological_space \ 'b::real_normed_vector) quasinorm" where "bcontfun\<^sub>N = quasinorm_of (1, \f. if f \ bcontfun then norm(Bcontfun f) else (\::ennreal))" lemma bcontfun\<^sub>N: fixes f::"('a::topological_space \ 'b::real_normed_vector)" shows "eNorm bcontfun\<^sub>N f = (if f \ bcontfun then norm(Bcontfun f) else (\::ennreal))" "Norm bcontfun\<^sub>N f = (if f \ bcontfun then norm(Bcontfun f) else 0)" "defect (bcontfun\<^sub>N::(('a \ 'b) quasinorm)) = 1" proof - have *: "quasinorm_on bcontfun 1 (\(f::('a \ 'b)). norm(Bcontfun f))" proof (rule quasinorm_onI, auto) fix f g::"'a \ 'b" assume H: "f \ bcontfun" "g \ bcontfun" then show "f + g \ bcontfun" unfolding plus_fun_def by (simp add: plus_cont) have *: "Bcontfun(f + g) = Bcontfun f + Bcontfun g" using H by (auto simp: eq_onp_def plus_fun_def bcontfun_def intro!: plus_bcontfun.abs_eq[symmetric]) show "ennreal (norm (Bcontfun (f + g))) \ ennreal (norm (Bcontfun f)) + ennreal (norm (Bcontfun g))" unfolding * using ennreal_leI[OF norm_triangle_ineq] by auto next fix c::real and f::"'a \ 'b" assume H: "f \ bcontfun" then show "c *\<^sub>R f \ bcontfun" unfolding scaleR_fun_def by (simp add: scaleR_cont) have *: "Bcontfun(c *\<^sub>R f) = c *\<^sub>R Bcontfun f" using H by (auto simp: eq_onp_def scaleR_fun_def bcontfun_def intro!: scaleR_bcontfun.abs_eq[symmetric]) show "ennreal (norm (Bcontfun (c *\<^sub>R f))) \ ennreal \c\ * ennreal (norm (Bcontfun f))" unfolding * by (simp add: ennreal_mult'') next show "(0::'a\'b) \ bcontfun" "Bcontfun 0 = 0" unfolding zero_fun_def zero_bcontfun_def by (auto simp add: const_bcontfun) qed have **: "quasinorm_on UNIV 1 (\(f::'a\'b). if f \ bcontfun then norm(Bcontfun f) else (\::ennreal))" by (rule extend_quasinorm[OF *]) show "eNorm bcontfun\<^sub>N f = (if f \ bcontfun then norm(Bcontfun f) else (\::ennreal))" "defect (bcontfun\<^sub>N::('a \ 'b) quasinorm) = 1" using quasinorm_of[OF **] unfolding bcontfun\<^sub>N_def by auto then show "Norm bcontfun\<^sub>N f = (if f \ bcontfun then norm(Bcontfun f) else 0)" unfolding Norm_def by auto qed lemma bcontfun\<^sub>N_space: "space\<^sub>N bcontfun\<^sub>N = bcontfun" using bcontfun\<^sub>N(1) by (metis (no_types, lifting) Collect_cong bcontfun_def enn2real_top ennreal_0 ennreal_enn2real ennreal_less_top ennreal_zero_neq_top infinity_ennreal_def mem_Collect_eq space\<^sub>N_def) lemma bcontfun\<^sub>N_zero_space: "zero_space\<^sub>N bcontfun\<^sub>N = {0}" apply (auto simp add: zero_spaceN_iff) by (metis Bcontfun_inject bcontfun\<^sub>N(1) eNorm_zero ennreal_eq_zero_iff ennreal_zero_neq_top infinity_ennreal_def norm_eq_zero norm_imp_pos_and_ge) lemma bcontfun\<^sub>ND: assumes "f \ space\<^sub>N bcontfun\<^sub>N" shows "continuous_on UNIV f" "\x. norm(f x) \ Norm bcontfun\<^sub>N f" proof- have "f \ bcontfun" using assms unfolding bcontfun\<^sub>N_space by simp then show "continuous_on UNIV f" unfolding bcontfun_def by auto show "\x. norm(f x) \ Norm bcontfun\<^sub>N f" using norm_bounded bcontfun\<^sub>N(2) \f \ bcontfun\ by (metis Bcontfun_inverse) qed lemma bcontfun\<^sub>NI: assumes "continuous_on UNIV f" "\x. norm(f x) \ C" shows "f \ space\<^sub>N bcontfun\<^sub>N" "Norm bcontfun\<^sub>N f \ C" proof - have "f \ bcontfun" using assms bcontfun_normI by blast then show "f \ space\<^sub>N bcontfun\<^sub>N" unfolding bcontfun\<^sub>N_space by simp show "Norm bcontfun\<^sub>N f \ C" unfolding bcontfun\<^sub>N(2) using \f \ bcontfun\ apply auto using assms(2) by (metis apply_bcontfun_cases apply_bcontfun_inverse norm_bound) qed subsection \Continuous inclusions between functional spaces\ text \Continuous inclusions between functional spaces are now defined\ instantiation quasinorm:: (real_vector) preorder begin definition less_eq_quasinorm::"'a quasinorm \ 'a quasinorm \ bool" where "less_eq_quasinorm N1 N2 = (\C\(0::real). \f. eNorm N2 f \ C * eNorm N1 f)" definition less_quasinorm::"'a quasinorm \ 'a quasinorm \ bool" where "less_quasinorm N1 N2 = (less_eq N1 N2 \ (\ less_eq N2 N1))" instance proof - have E: "N \ N" for N::"'a quasinorm" unfolding less_eq_quasinorm_def by (rule exI[of _ 1], auto) have T: "N1 \ N3" if "N1 \ N2" "N2 \ N3" for N1 N2 N3::"'a quasinorm" proof - obtain C C' where *: "\f. eNorm N2 f \ ennreal C * eNorm N1 f" "\f. eNorm N3 f \ ennreal C' * eNorm N2 f" "C \ 0" "C' \ 0" using \N1 \ N2\ \N2 \ N3\ unfolding less_eq_quasinorm_def by metis { fix f have "eNorm N3 f \ ennreal C' * ennreal C * eNorm N1 f" by (metis *(1)[of f] *(2)[of f] mult.commute mult.left_commute mult_left_mono order_trans zero_le) also have "... = ennreal(C' * C) * eNorm N1 f" using \C \ 0\ \C' \ 0\ ennreal_mult by auto finally have "eNorm N3 f \ ennreal(C' * C) * eNorm N1 f" by simp } then show ?thesis unfolding less_eq_quasinorm_def using \C \ 0\ \C' \ 0\ zero_le_mult_iff by auto qed show "OFCLASS('a quasinorm, preorder_class)" apply standard unfolding less_quasinorm_def apply simp using E apply fast using T apply fast done qed end abbreviation quasinorm_subset :: "('a::real_vector) quasinorm \ 'a quasinorm \ bool" where "quasinorm_subset \ less" abbreviation quasinorm_subset_eq :: "('a::real_vector) quasinorm \ 'a quasinorm \ bool" where "quasinorm_subset_eq \ less_eq" notation quasinorm_subset ("'(\\<^sub>N')") and quasinorm_subset ("(_/ \\<^sub>N _)" [51, 51] 50) and quasinorm_subset_eq ("'(\\<^sub>N')") and quasinorm_subset_eq ("(_/ \\<^sub>N _)" [51, 51] 50) lemma quasinorm_subsetD: assumes "N1 \\<^sub>N N2" shows "\C\(0::real). \f. eNorm N2 f \ C * eNorm N1 f" using assms unfolding less_eq_quasinorm_def by auto lemma quasinorm_subsetI: assumes "\f. f \ space\<^sub>N N1 \ eNorm N2 f \ ennreal C * eNorm N1 f" shows "N1 \\<^sub>N N2" proof - have "eNorm N2 f \ ennreal (max C 1) * eNorm N1 f" for f proof (cases "f \ space\<^sub>N N1") case True then show ?thesis using assms[OF \f \ space\<^sub>N N1\] by (metis (no_types, opaque_lifting) dual_order.trans ennreal_leI max.cobounded2 max.commute mult.commute ordered_comm_semiring_class.comm_mult_left_mono zero_le) next case False then show ?thesis using spaceN_iff by (metis ennreal_ge_1 ennreal_mult_less_top infinity_ennreal_def max.cobounded1 max.commute not_le not_one_le_zero top.not_eq_extremum) qed then show ?thesis unfolding less_eq_quasinorm_def by (metis ennreal_max_0' max.cobounded2) qed lemma quasinorm_subsetI': assumes "\f. f \ space\<^sub>N N1 \ f \ space\<^sub>N N2" "\f. f \ space\<^sub>N N1 \ Norm N2 f \ C * Norm N1 f" shows "N1 \\<^sub>N N2" proof (rule quasinorm_subsetI) fix f assume "f \ space\<^sub>N N1" then have "f \ space\<^sub>N N2" using assms(1) by simp then have "eNorm N2 f = ennreal(Norm N2 f)" using eNorm_Norm by auto also have "... \ ennreal(C * Norm N1 f)" using assms(2)[OF \f \ space\<^sub>N N1\] ennreal_leI by blast also have "... = ennreal C * ennreal(Norm N1 f)" using ennreal_mult'' by auto also have "... = ennreal C * eNorm N1 f" using eNorm_Norm[OF \f \ space\<^sub>N N1\] by auto finally show "eNorm N2 f \ ennreal C * eNorm N1 f" by simp qed lemma quasinorm_subset_space: assumes "N1 \\<^sub>N N2" shows "space\<^sub>N N1 \ space\<^sub>N N2" using assms unfolding space\<^sub>N_def less_eq_quasinorm_def by (auto, metis ennreal_mult_eq_top_iff ennreal_neq_top less_le top.extremum_strict top.not_eq_extremum) lemma quasinorm_subset_Norm_eNorm: assumes "f \ space\<^sub>N N1 \ Norm N2 f \ C * Norm N1 f" "N1 \\<^sub>N N2" "C > 0" shows "eNorm N2 f \ ennreal C * eNorm N1 f" proof (cases "f \ space\<^sub>N N1") case True then have "f \ space\<^sub>N N2" using quasinorm_subset_space[OF \N1 \\<^sub>N N2\] by auto then show ?thesis using eNorm_Norm[OF True] eNorm_Norm assms(1)[OF True] by (metis Norm_nonneg ennreal_leI ennreal_mult'') next case False then show ?thesis using \C > 0\ by (metis ennreal_eq_zero_iff ennreal_mult_eq_top_iff infinity_ennreal_def less_imp_le neq_top_trans not_le spaceN_iff) qed lemma quasinorm_subset_zero_space: assumes "N1 \\<^sub>N N2" shows "zero_space\<^sub>N N1 \ zero_space\<^sub>N N2" using assms unfolding zero_space\<^sub>N_def less_eq_quasinorm_def by (auto, metis le_zero_eq mult_zero_right) text \We would like to define the equivalence relation associated to the above order, i.e., the equivalence between norms. This is not equality, so we do not have a true order, but nevertheless this is handy, and not standard in a preorder in Isabelle. The file Library/Preorder.thy defines such an equivalence relation, but including it breaks some proofs so we go the naive way.\ definition quasinorm_equivalent::"('a::real_vector) quasinorm \ 'a quasinorm \ bool" (infix "=\<^sub>N" 60) where "quasinorm_equivalent N1 N2 = ((N1 \\<^sub>N N2) \ (N2 \\<^sub>N N1))" lemma quasinorm_equivalent_sym [sym]: assumes "N1 =\<^sub>N N2" shows "N2 =\<^sub>N N1" using assms unfolding quasinorm_equivalent_def by auto lemma quasinorm_equivalent_trans [trans]: assumes "N1 =\<^sub>N N2" "N2 =\<^sub>N N3" shows "N1 =\<^sub>N N3" using assms order_trans unfolding quasinorm_equivalent_def by blast subsection \The intersection and the sum of two functional spaces\ text \In this paragraph, we define the intersection and the sum of two functional spaces. In terms of the order introduced above, this corresponds to the minimum and the maximum. More important, these are the first two examples of interpolation spaces between two functional spaces, and they are central as all the other ones are built using them.\ definition quasinorm_intersection::"('a::real_vector) quasinorm \ 'a quasinorm \ 'a quasinorm" (infix "\\<^sub>N" 70) where "quasinorm_intersection N1 N2 = quasinorm_of (max (defect N1) (defect N2), \f. eNorm N1 f + eNorm N2 f)" lemma quasinorm_intersection: "eNorm (N1 \\<^sub>N N2) f = eNorm N1 f + eNorm N2 f" "defect (N1 \\<^sub>N N2) = max (defect N1) (defect N2)" proof - have T: "eNorm N1 (x + y) + eNorm N2 (x + y) \ ennreal (max (defect N1) (defect N2)) * (eNorm N1 x + eNorm N2 x) + ennreal (max (defect N1) (defect N2)) * (eNorm N1 y + eNorm N2 y)" for x y proof - have "eNorm N1 (x + y) \ ennreal (max (defect N1) (defect N2)) * eNorm N1 x + ennreal (max (defect N1) (defect N2)) * eNorm N1 y" using eNorm_triangular_ineq[of N1 x y] by (metis (no_types) max_def distrib_left ennreal_leI mult_right_mono order_trans zero_le) moreover have "eNorm N2 (x + y) \ ennreal (max (defect N1) (defect N2)) * eNorm N2 x + ennreal (max (defect N1) (defect N2)) * eNorm N2 y" using eNorm_triangular_ineq[of N2 x y] by (metis (no_types) max_def max.commute distrib_left ennreal_leI mult_right_mono order_trans zero_le) ultimately have "eNorm N1 (x + y) + eNorm N2 (x + y) \ ennreal (max (defect N1) (defect N2)) * (eNorm N1 x + eNorm N1 y + (eNorm N2 x + eNorm N2 y))" by (simp add: add_mono_thms_linordered_semiring(1) distrib_left) then show ?thesis by (simp add: ab_semigroup_add_class.add_ac(1) add.left_commute distrib_left) qed have H: "eNorm N1 (c *\<^sub>R x) + eNorm N2 (c *\<^sub>R x) \ ennreal \c\ * (eNorm N1 x + eNorm N2 x)" for c x by (simp add: eNorm_cmult[of N1 c x] eNorm_cmult[of N2 c x] distrib_left) have *: "quasinorm_on UNIV (max (defect N1) (defect N2)) (\f. eNorm N1 f + eNorm N2 f)" apply (rule quasinorm_onI) using T H defect_ge_1[of N1] defect_ge_1[of N2] by auto show "defect (N1 \\<^sub>N N2) = max (defect N1) (defect N2)" "eNorm (N1 \\<^sub>N N2) f = eNorm N1 f + eNorm N2 f" unfolding quasinorm_intersection_def using quasinorm_of[OF *] by auto qed lemma quasinorm_intersection_commute: "N1 \\<^sub>N N2 = N2 \\<^sub>N N1" unfolding quasinorm_intersection_def max.commute[of "defect N1"] add.commute[of "eNorm N1 _"] by simp lemma quasinorm_intersection_space: "space\<^sub>N (N1 \\<^sub>N N2) = space\<^sub>N N1 \ space\<^sub>N N2" apply auto unfolding quasinorm_intersection(1) spaceN_iff by auto lemma quasinorm_intersection_zero_space: "zero_space\<^sub>N (N1 \\<^sub>N N2) = zero_space\<^sub>N N1 \ zero_space\<^sub>N N2" apply auto unfolding quasinorm_intersection(1) zero_spaceN_iff by (auto simp add: add_eq_0_iff_both_eq_0) lemma quasinorm_intersection_subset: "N1 \\<^sub>N N2 \\<^sub>N N1" "N1 \\<^sub>N N2 \\<^sub>N N2" by (rule quasinorm_subsetI[of _ _ 1], auto simp add: quasinorm_intersection(1))+ lemma quasinorm_intersection_minimum: assumes "N \\<^sub>N N1" "N \\<^sub>N N2" shows "N \\<^sub>N N1 \\<^sub>N N2" proof - obtain C1 C2::real where *: "\f. eNorm N1 f \ C1 * eNorm N f" "\f. eNorm N2 f \ C2 * eNorm N f" "C1 \ 0" "C2 \ 0" using quasinorm_subsetD[OF assms(1)] quasinorm_subsetD[OF assms(2)] by blast have **: "eNorm (N1 \\<^sub>N N2) f \ (C1 + C2) * eNorm N f" for f unfolding quasinorm_intersection(1) using add_mono[OF *(1) *(2)] by (simp add: distrib_right *) show ?thesis apply (rule quasinorm_subsetI) using ** by auto qed lemma quasinorm_intersection_assoc: "(N1 \\<^sub>N N2) \\<^sub>N N3 =\<^sub>N N1 \\<^sub>N (N2 \\<^sub>N N3)" unfolding quasinorm_equivalent_def by (meson order_trans quasinorm_intersection_minimum quasinorm_intersection_subset) definition quasinorm_sum::"('a::real_vector) quasinorm \ 'a quasinorm \ 'a quasinorm" (infix "+\<^sub>N" 70) where "quasinorm_sum N1 N2 = quasinorm_of (max (defect N1) (defect N2), \f. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})" lemma quasinorm_sum: "eNorm (N1 +\<^sub>N N2) f = Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" "defect (N1 +\<^sub>N N2) = max (defect N1) (defect N2)" proof - define N where "N = (\f. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})" have T: "N (f+g) \ ennreal (max (defect N1) (defect N2)) * N f + ennreal (max (defect N1) (defect N2)) * N g" for f g proof - have "\u. (\n. u n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) \ u \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto) then obtain uf where uf: "\n. uf n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" "uf \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by blast have "\f1 f2. \n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) \ f = f1 n + f2 n" - apply (rule SMT.choices(1)) using uf(1) by blast + using uf(1) by (subst choice_iff[symmetric])+ blast then obtain f1 f2 where F: "\n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "\n. f = f1 n + f2 n" by blast have "\u. (\n. u n \ {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}) \ u \ Inf {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}" by (rule Inf_as_limit, auto, rule exI[of _ "g"], rule exI[of _ 0], auto) then obtain ug where ug: "\n. ug n \ {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}" "ug \ Inf {eNorm N1 g1 + eNorm N2 g2| g1 g2. g = g1 + g2}" by blast have "\g1 g2. \n. ug n = eNorm N1 (g1 n) + eNorm N2 (g2 n) \ g = g1 n + g2 n" - apply (rule SMT.choices(1)) using ug(1) by blast + using ug(1) by (subst choice_iff[symmetric])+ blast then obtain g1 g2 where G: "\n. ug n = eNorm N1 (g1 n) + eNorm N2 (g2 n)" "\n. g = g1 n + g2 n" by blast define h1 where "h1 = (\n. f1 n + g1 n)" define h2 where "h2 = (\n. f2 n + g2 n)" have *: "f + g = h1 n + h2 n" for n unfolding h1_def h2_def using F(2) G(2) by (auto simp add: algebra_simps) have "N (f+g) \ ennreal (max (defect N1) (defect N2)) * (uf n + ug n)" for n proof - have "N (f+g) \ eNorm N1 (h1 n) + eNorm N2 (h2 n)" unfolding N_def apply (rule Inf_lower, auto, rule exI[of _ "h1 n"], rule exI[of _ "h2 n"]) using * by auto also have "... \ ennreal (defect N1) * eNorm N1 (f1 n) + ennreal (defect N1) * eNorm N1 (g1 n) + (ennreal (defect N2) * eNorm N2 (f2 n) + ennreal (defect N2) * eNorm N2 (g2 n))" unfolding h1_def h2_def apply (rule add_mono) using eNorm_triangular_ineq by auto also have "... \ (ennreal (max (defect N1) (defect N2)) * eNorm N1 (f1 n) + ennreal (max (defect N1) (defect N2)) * eNorm N1 (g1 n)) + (ennreal (max (defect N1) (defect N2)) * eNorm N2 (f2 n) + ennreal (max (defect N1) (defect N2)) * eNorm N2 (g2 n))" by (auto intro!: add_mono mult_mono ennreal_leI) also have "... = ennreal (max (defect N1) (defect N2)) * (uf n + ug n)" unfolding F(1) G(1) by (auto simp add: algebra_simps) finally show ?thesis by simp qed moreover have "... \ ennreal (max (defect N1) (defect N2)) * (N f + N g)" unfolding N_def by (auto intro!: tendsto_intros simp add: uf(2) ug(2)) ultimately have "N (f+g) \ ennreal (max (defect N1) (defect N2)) * (N f + N g)" using LIMSEQ_le_const by blast then show ?thesis by (auto simp add: algebra_simps) qed have H: "N (c *\<^sub>R f) \ ennreal \c\ * N f" for c f proof - have "\u. (\n. u n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) \ u \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto) then obtain uf where uf: "\n. uf n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" "uf \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by blast have "\f1 f2. \n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) \ f = f1 n + f2 n" - apply (rule SMT.choices(1)) using uf(1) by blast + using uf(1) by (subst choice_iff[symmetric])+ blast then obtain f1 f2 where F: "\n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "\n. f = f1 n + f2 n" by blast have "N (c *\<^sub>R f) \ \c\ * uf n" for n proof - have "N (c *\<^sub>R f) \ eNorm N1 (c *\<^sub>R f1 n) + eNorm N2 (c *\<^sub>R f2 n)" unfolding N_def apply (rule Inf_lower, auto, rule exI[of _ "c *\<^sub>R f1 n"], rule exI[of _ "c *\<^sub>R f2 n"]) using F(2)[of n] scaleR_add_right by auto also have "... = \c\ * (eNorm N1 (f1 n) + eNorm N2 (f2 n))" by (auto simp add: algebra_simps eNorm_cmult) finally show ?thesis using F(1) by simp qed moreover have "... \ \c\ * N f" unfolding N_def by (auto intro!: tendsto_intros simp add: uf(2)) ultimately show ?thesis using LIMSEQ_le_const by blast qed have "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. 0 = f1 + f2} \ 0" by (rule Inf_lower, auto, rule exI[of _ 0], auto) then have Z: "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. 0 = f1 + f2} = 0" by auto have *: "quasinorm_on UNIV (max (defect N1) (defect N2)) (\f. Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2})" apply (rule quasinorm_onI) using T H Z defect_ge_1[of N1] defect_ge_1[of N2] unfolding N_def by auto show "defect (N1 +\<^sub>N N2) = max (defect N1) (defect N2)" "eNorm (N1 +\<^sub>N N2) f = Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" unfolding quasinorm_sum_def using quasinorm_of[OF *] by auto qed lemma quasinorm_sum_limit: "\f1 f2. (\n. f = f1 n + f2 n) \ (\n. eNorm N1 (f1 n) + eNorm N2 (f2 n)) \ eNorm (N1 +\<^sub>N N2) f" proof - have "\u. (\n. u n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}) \ u \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by (rule Inf_as_limit, auto, rule exI[of _ "f"], rule exI[of _ 0], auto) then obtain uf where uf: "\n. uf n \ {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" "uf \ Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f = f1 + f2}" by blast have "\f1 f2. \n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n) \ f = f1 n + f2 n" - apply (rule SMT.choices(1)) using uf(1) by blast + using uf(1) by (subst choice_iff[symmetric])+ blast then obtain f1 f2 where F: "\n. uf n = eNorm N1 (f1 n) + eNorm N2 (f2 n)" "\n. f = f1 n + f2 n" by blast have "(\n. eNorm N1 (f1 n) + eNorm N2 (f2 n)) \ eNorm (N1 +\<^sub>N N2) f" using F(1) uf(2) unfolding quasinorm_sum(1) by presburger then show ?thesis using F(2) by auto qed lemma quasinorm_sum_space: "space\<^sub>N (N1 +\<^sub>N N2) = {f + g|f g. f \ space\<^sub>N N1 \ g \ space\<^sub>N N2}" proof (auto) fix x assume "x \ space\<^sub>N (N1 +\<^sub>N N2)" then have "Inf {eNorm N1 f + eNorm N2 g| f g. x = f + g} < \" unfolding quasinorm_sum(1) spaceN_iff. then have "\z \ {eNorm N1 f + eNorm N2 g| f g. x = f + g}. z < \" by (simp add: Inf_less_iff) then show "\f g. x = f + g \ f \ space\<^sub>N N1 \ g \ space\<^sub>N N2" using spaceN_iff by force next fix f g assume H: "f \ space\<^sub>N N1" "g \ space\<^sub>N N2" have "Inf {eNorm N1 u + eNorm N2 v| u v. f + g = u + v} \ eNorm N1 f + eNorm N2 g" by (rule Inf_lower, auto) also have "... < \" using spaceN_iff H by auto finally show "f + g \ space\<^sub>N (N1 +\<^sub>N N2)" unfolding spaceN_iff quasinorm_sum(1). qed lemma quasinorm_sum_zerospace: "{f + g |f g. f \ zero_space\<^sub>N N1 \ g \ zero_space\<^sub>N N2} \ zero_space\<^sub>N (N1 +\<^sub>N N2)" proof (auto, unfold zero_spaceN_iff) fix f g assume H: "eNorm N1 f = 0" "eNorm N2 g = 0" have "Inf {eNorm N1 f1 + eNorm N2 f2| f1 f2. f + g = f1 + f2} \ 0" by (rule Inf_lower, auto, rule exI[of _ f], auto simp add: H) then show "eNorm (N1 +\<^sub>N N2) (f + g) = 0" unfolding quasinorm_sum(1) by auto qed lemma quasinorm_sum_subset: "N1 \\<^sub>N N1 +\<^sub>N N2" "N2 \\<^sub>N N1 +\<^sub>N N2" by (rule quasinorm_subsetI[of _ _ 1], auto simp add: quasinorm_sum(1), rule Inf_lower, auto, metis add.commute add.left_neutral eNorm_zero)+ lemma quasinorm_sum_maximum: assumes "N1 \\<^sub>N N" "N2 \\<^sub>N N" shows "N1 +\<^sub>N N2 \\<^sub>N N" proof - obtain C1 C2::real where *: "\f. eNorm N f \ C1 * eNorm N1 f" "\f. eNorm N f \ C2 * eNorm N2 f" "C1 \ 0" "C2 \ 0" using quasinorm_subsetD[OF assms(1)] quasinorm_subsetD[OF assms(2)] by blast have **: "eNorm N f \ (defect N * max C1 C2) * eNorm (N1 +\<^sub>N N2) f" for f proof - obtain f1 f2 where F: "\n. f = f1 n + f2 n" "(\n. eNorm N1 (f1 n) + eNorm N2 (f2 n)) \ eNorm (N1 +\<^sub>N N2) f" using quasinorm_sum_limit by blast have "eNorm N f \ ennreal (defect N * max C1 C2) * (eNorm N1 (f1 n) + eNorm N2 (f2 n))" for n proof - have "eNorm N f \ ennreal(defect N) * eNorm N (f1 n) + ennreal(defect N) * eNorm N (f2 n)" unfolding \f = f1 n + f2 n\ using eNorm_triangular_ineq by auto also have "... \ ennreal(defect N) * (C1 * eNorm N1 (f1 n)) + ennreal(defect N) * (C2 * eNorm N2 (f2 n))" apply (rule add_mono) by (rule mult_mono, simp, simp add: *, simp, simp)+ also have "... \ ennreal(defect N) * (max C1 C2 * eNorm N1 (f1 n)) + ennreal(defect N) * (max C1 C2 * eNorm N2 (f2 n))" by (auto intro!:add_mono mult_mono ennreal_leI) also have "... = ennreal (defect N * max C1 C2) * (eNorm N1 (f1 n) + eNorm N2 (f2 n))" apply (subst ennreal_mult') using defect_ge_1 order_trans zero_le_one apply blast by (auto simp add: algebra_simps) finally show ?thesis by simp qed moreover have "... \ (defect N * max C1 C2) * eNorm (N1 +\<^sub>N N2) f" by (auto intro!:tendsto_intros F(2)) ultimately show ?thesis using LIMSEQ_le_const by blast qed then show ?thesis using quasinorm_subsetI by force qed lemma quasinorm_sum_assoc: "(N1 +\<^sub>N N2) +\<^sub>N N3 =\<^sub>N N1 +\<^sub>N (N2 +\<^sub>N N3)" unfolding quasinorm_equivalent_def by (meson order_trans quasinorm_sum_maximum quasinorm_sum_subset) subsection \Topology\ definition topology\<^sub>N::"('a::real_vector) quasinorm \ 'a topology" where "topology\<^sub>N N = topology (\U. \x\U. \e>0. \y. eNorm N (y-x) < e \ y \ U)" lemma istopology_topology\<^sub>N: "istopology (\U. \x\U. \e>0. \y. eNorm N (y-x) < e \ y \ U)" unfolding istopology_def by (auto, metis dual_order.strict_trans less_linear, meson) lemma openin_topology\<^sub>N: "openin (topology\<^sub>N N) U \ (\x\U. \e>0. \y. eNorm N (y-x) < e \ y \ U)" unfolding topology\<^sub>N_def using istopology_topology\<^sub>N[of N] by (simp add: topology_inverse') lemma openin_topology\<^sub>N_I: assumes "\x. x \ U \ \e>0. \y. eNorm N (y-x) < e \ y \ U" shows "openin (topology\<^sub>N N) U" using assms unfolding openin_topology\<^sub>N by auto lemma openin_topology\<^sub>N_D: assumes "openin (topology\<^sub>N N) U" "x \ U" shows "\e>0. \y. eNorm N (y-x) < e \ y \ U" using assms unfolding openin_topology\<^sub>N by auto text \One should then use this topology to define limits and so on. This is not something specific to quasinorms, but to all topologies defined in this way, not using type classes. However, there is no such body of material (yet?) in Isabelle-HOL, where topology is essentially done with type classes. So, we do not go any further for now. One exception is the notion of completeness, as it is so important in functional analysis. We give a naive definition, which will be sufficient for the proof of completeness of several spaces. Usually, the most convenient criterion to prove completeness of a normed vector space is in terms of converging series. This criterion is the only nontrivial thing we prove here. We will apply it to prove the completeness of $L^p$ spaces.\ definition cauchy_ine\<^sub>N::"('a::real_vector) quasinorm \ (nat \ 'a) \ bool" where "cauchy_ine\<^sub>N N u = (\e>0. \M. \n\M. \m\M. eNorm N (u n - u m) < e)" definition tendsto_ine\<^sub>N::"('a::real_vector) quasinorm \ (nat \ 'a) \ 'a => bool" where "tendsto_ine\<^sub>N N u x = (\n. eNorm N (u n - x)) \ 0" definition complete\<^sub>N::"('a::real_vector) quasinorm \ bool" where "complete\<^sub>N N = (\u. cauchy_ine\<^sub>N N u \ (\x. tendsto_ine\<^sub>N N u x))" text \The above definitions are in terms of eNorms, but usually the nice definitions only make sense on the space of the norm, and are expressed in terms of Norms. We formulate the same definitions with norms, they will be more convenient for the proofs.\ definition cauchy_in\<^sub>N::"('a::real_vector) quasinorm \ (nat \ 'a) \ bool" where "cauchy_in\<^sub>N N u = (\e>0. \M. \n\M. \m\M. Norm N (u n - u m) < e)" definition tendsto_in\<^sub>N::"('a::real_vector) quasinorm \ (nat \ 'a) \ 'a => bool" where "tendsto_in\<^sub>N N u x = (\n. Norm N (u n - x)) \ 0" lemma cauchy_ine\<^sub>N_I: assumes "\e. e > 0 \ (\M. \n\M. \m\M. eNorm N (u n - u m) < e)" shows "cauchy_ine\<^sub>N N u" using assms unfolding cauchy_ine\<^sub>N_def by auto lemma cauchy_in\<^sub>N_I: assumes "\e. e > 0 \ (\M. \n\M. \m\M. Norm N (u n - u m) < e)" shows "cauchy_in\<^sub>N N u" using assms unfolding cauchy_in\<^sub>N_def by auto lemma cauchy_ine_in: assumes "\n. u n \ space\<^sub>N N" shows "cauchy_ine\<^sub>N N u \ cauchy_in\<^sub>N N u" proof assume "cauchy_in\<^sub>N N u" show "cauchy_ine\<^sub>N N u" proof (rule cauchy_ine\<^sub>N_I) fix e::ennreal assume "e > 0" define e2 where "e2 = min e 1" then obtain r where "e2 = ennreal r" "r > 0" unfolding e2_def using \e > 0\ by (metis ennreal_eq_1 ennreal_less_zero_iff le_ennreal_iff le_numeral_extra(1) min_def zero_less_one) then obtain M where *: "\n\M. \m\M. Norm N (u n - u m) < r" using \cauchy_in\<^sub>N N u\ \r > 0\ unfolding cauchy_in\<^sub>N_def by auto then have "\n\M. \m\M. eNorm N (u n - u m) < r" by (auto simp add: assms eNorm_Norm \0 < r\ ennreal_lessI) then have "\n\M. \m\M. eNorm N (u n - u m) < e" unfolding \e2 = ennreal r\[symmetric] e2_def by auto then show "\M. \n\M. \m\M. eNorm N (u n - u m) < e" by auto qed next assume "cauchy_ine\<^sub>N N u" show "cauchy_in\<^sub>N N u" proof (rule cauchy_in\<^sub>N_I) fix e::real assume "e > 0" then obtain M where *: "\n\M. \m\M. eNorm N (u n - u m) < e" using \cauchy_ine\<^sub>N N u\ \e > 0\ ennreal_less_zero_iff unfolding cauchy_ine\<^sub>N_def by blast then have "\n\M. \m\M. Norm N (u n - u m) < e" by (auto, metis Norm_def \0 < e\ eNorm_Norm eNorm_Norm' enn2real_nonneg ennreal_less_iff) then show "\M. \n\M. \m\M. Norm N (u n - u m) < e" by auto qed qed lemma tendsto_ine_in: assumes "\n. u n \ space\<^sub>N N" "x \ space\<^sub>N N" shows "tendsto_ine\<^sub>N N u x \ tendsto_in\<^sub>N N u x" proof - have *: "eNorm N (u n - x) = Norm N (u n - x)" for n using assms eNorm_Norm spaceN_diff by blast show ?thesis unfolding tendsto_in\<^sub>N_def tendsto_ine\<^sub>N_def * apply (auto) apply (metis (full_types) Norm_nonneg ennreal_0 eventually_sequentiallyI order_refl tendsto_ennreal_iff) using tendsto_ennrealI by fastforce qed lemma complete\<^sub>N_I: assumes "\u. cauchy_in\<^sub>N N u \ (\n. u n \ space\<^sub>N N) \ (\x\ space\<^sub>N N. tendsto_in\<^sub>N N u x)" shows "complete\<^sub>N N" proof - have "\x. tendsto_ine\<^sub>N N u x" if "cauchy_ine\<^sub>N N u" for u proof - obtain M::nat where *: "\n m. n \ M \ m \ M \ eNorm N (u n - u m) < 1" using \cauchy_ine\<^sub>N N u\ ennreal_zero_less_one unfolding cauchy_ine\<^sub>N_def by presburger define v where "v = (\n. u (n+M) - u M)" have "eNorm N (v n) < 1" for n unfolding v_def using * by auto then have "v n \ space\<^sub>N N" for n using spaceN_iff[of _ N] by (metis dual_order.strict_trans ennreal_1 ennreal_less_top infinity_ennreal_def) have "cauchy_ine\<^sub>N N v" proof (rule cauchy_ine\<^sub>N_I) fix e::ennreal assume "e > 0" then obtain P::nat where *: "\n m. n \ P \ m \ P \ eNorm N (u n - u m) < e" using \cauchy_ine\<^sub>N N u\ unfolding cauchy_ine\<^sub>N_def by presburger have "eNorm N (v n - v m) < e" if "n \ P" "m \ P" for m n unfolding v_def by (auto, rule *, insert that, auto) then show "\M. \n\M. \m\M. eNorm N (v n - v m) < e" by auto qed then have "cauchy_in\<^sub>N N v" using cauchy_ine_in[OF \\n. v n \ space\<^sub>N N\] by auto then obtain y where "tendsto_in\<^sub>N N v y" "y \ space\<^sub>N N" using assms \\n. v n \ space\<^sub>N N\ by auto then have *: "tendsto_ine\<^sub>N N v y" using tendsto_ine_in \\n. v n \ space\<^sub>N N\ by auto have "tendsto_ine\<^sub>N N u (y + u M)" unfolding tendsto_ine\<^sub>N_def apply (rule LIMSEQ_offset[of _ M]) using * unfolding v_def tendsto_ine\<^sub>N_def by (auto simp add: algebra_simps) then show ?thesis by auto qed then show ?thesis unfolding complete\<^sub>N_def by auto qed lemma cauchy_tendsto_in_subseq: assumes "\n. u n \ space\<^sub>N N" "cauchy_in\<^sub>N N u" "strict_mono r" "tendsto_in\<^sub>N N (u o r) x" shows "tendsto_in\<^sub>N N u x" proof - have "\M. \n\M. Norm N (u n - x) < e" if "e > 0" for e proof - define f where "f = e / (2 * defect N)" have "f > 0" unfolding f_def using \e > 0\ defect_ge_1[of N] by (auto simp add: divide_simps) obtain M1 where M1: "\m n. m \ M1 \ n \ M1 \ Norm N (u n - u m) < f" using \cauchy_in\<^sub>N N u\ unfolding cauchy_in\<^sub>N_def using \f > 0\ by meson obtain M2 where M2: "\n. n \ M2 \ Norm N ((u o r) n - x) < f" using \tendsto_in\<^sub>N N (u o r) x\ \f > 0\ unfolding tendsto_in\<^sub>N_def order_tendsto_iff eventually_sequentially by blast define M where "M = max M1 M2" have "Norm N (u n - x) < e" if "n \ M" for n proof - have "Norm N (u n - x) = Norm N ((u n - u (r M)) + (u (r M) - x))" by auto also have "... \ defect N * Norm N (u n - u (r M)) + defect N * Norm N (u (r M) - x)" apply (rule Norm_triangular_ineq) using \\n. u n \ space\<^sub>N N\ by simp also have "... < defect N * f + defect N * f" apply (auto intro!: add_strict_mono mult_mono simp only:) using defect_ge_1[of N] \n \ M\ seq_suble[OF \strict_mono r\, of M] M1 M2 o_def unfolding M_def by auto finally show ?thesis unfolding f_def using \e > 0\ defect_ge_1[of N] by (auto simp add: divide_simps) qed then show ?thesis by auto qed then show ?thesis unfolding tendsto_in\<^sub>N_def order_tendsto_iff eventually_sequentially using Norm_nonneg less_le_trans by blast qed proposition complete\<^sub>N_I': assumes "\n. c n > 0" "\u. (\n. u n \ space\<^sub>N N) \ (\n. Norm N (u n) \ c n) \ \x\ space\<^sub>N N. tendsto_in\<^sub>N N (\n. (\i\{0..N N" proof (rule complete\<^sub>N_I) fix v assume "cauchy_in\<^sub>N N v" "\n. v n \ space\<^sub>N N" have *: "\y. (\m\y. \p\y. Norm N (v m - v p) < c (Suc n)) \ x < y" if "\m\x. \p\x. Norm N (v m - v p) < c n" for x n proof - obtain M where i: "\m\M. \p\M. Norm N (v m - v p) < c (Suc n)" using \cauchy_in\<^sub>N N v\ \c (Suc n) > 0\ unfolding cauchy_in\<^sub>N_def by (meson zero_less_power) then show ?thesis apply (intro exI[of _ "max M (x+1)"]) by auto qed have "\r. \n. (\m\r n. \p\r n. Norm N (v m - v p) < c n) \ r n < r (Suc n)" apply (intro dependent_nat_choice) using \cauchy_in\<^sub>N N v\ \\n. c n > 0\ * unfolding cauchy_in\<^sub>N_def by auto then obtain r where r: "strict_mono r" "\n. \m\r n. \p\r n. Norm N (v m - v p) < c n" by (auto simp: strict_mono_Suc_iff) define u where "u = (\n. v (r (Suc n)) - v (r n))" have "u n \ space\<^sub>N N" for n unfolding u_def using \\n. v n \ space\<^sub>N N\ by simp moreover have "Norm N (u n) \ c n" for n unfolding u_def using r by (simp add: less_imp_le strict_mono_def) ultimately obtain y where y: "y \ space\<^sub>N N" "tendsto_in\<^sub>N N (\n. (\i\{0.. space\<^sub>N N" unfolding x_def using \y \ space\<^sub>N N\ \\n. v n \ space\<^sub>N N\ by simp have "Norm N (v (r n) - x) = Norm N ((\i\{0..i\{0..n. Norm N (v (r n) - x)) \ 0" using y(2) unfolding tendsto_in\<^sub>N_def by auto then have "tendsto_in\<^sub>N N (v o r) x" unfolding tendsto_in\<^sub>N_def comp_def by force then have "tendsto_in\<^sub>N N v x" using \\n. v n \ space\<^sub>N N\ by (intro cauchy_tendsto_in_subseq[OF _ \cauchy_in\<^sub>N N v\ \strict_mono r\], auto) then show "\x\space\<^sub>N N. tendsto_in\<^sub>N N v x" using \x \ space\<^sub>N N\ by blast qed text \Next, we show when the two examples of norms we have introduced before, the ambient norm in a Banach space, and the norm on bounded continuous functions, are complete. We just have to translate in our setting the already known completeness of these spaces.\ lemma complete_N_of_norm: "complete\<^sub>N (N_of_norm::'a::banach quasinorm)" proof (rule complete\<^sub>N_I) fix u::"nat \ 'a" assume "cauchy_in\<^sub>N N_of_norm u" then have "Cauchy u" unfolding Cauchy_def cauchy_in\<^sub>N_def N_of_norm(2) by (simp add: dist_norm) then obtain x where "u \ x" using convergent_eq_Cauchy by blast then have "tendsto_in\<^sub>N N_of_norm u x" unfolding tendsto_in\<^sub>N_def N_of_norm(2) using Lim_null tendsto_norm_zero_iff by fastforce moreover have "x \ space\<^sub>N N_of_norm" by auto ultimately show "\x\space\<^sub>N N_of_norm. tendsto_in\<^sub>N N_of_norm u x" by auto qed text \In the next statement, the assumption that \verb+'a+ is a metric space is not necessary, a topological space would be enough, but a statement about uniform convergence is not available in this setting. TODO: fix it. \ lemma complete_bcontfunN: "complete\<^sub>N (bcontfun\<^sub>N::('a::metric_space \ 'b::banach) quasinorm)" proof (rule complete\<^sub>N_I) fix u::"nat \ ('a \ 'b)" assume H: "cauchy_in\<^sub>N bcontfun\<^sub>N u" "\n. u n \ space\<^sub>N bcontfun\<^sub>N" then have H2: "u n \ bcontfun" for n using bcontfun\<^sub>N_space by auto then have **: "Bcontfun(u n - u m) = Bcontfun (u n) - Bcontfun (u m)" for m n unfolding minus_fun_def minus_bcontfun_def by (simp add: Bcontfun_inverse) have *: "Norm bcontfun\<^sub>N (u n - u m) = norm (Bcontfun (u n - u m))" for n m unfolding bcontfun\<^sub>N(2) using H(2) bcontfun\<^sub>N_space by auto have "Cauchy (\n. Bcontfun (u n))" using H(1) unfolding Cauchy_def cauchy_in\<^sub>N_def dist_norm * ** by simp then obtain v where v: "(\n. Bcontfun (u n)) \ v" using convergent_eq_Cauchy by blast have v_space: "apply_bcontfun v \ space\<^sub>N bcontfun\<^sub>N" unfolding bcontfun\<^sub>N_space by (simp add: apply_bcontfun) have ***: "Norm bcontfun\<^sub>N (u n - v) = norm(Bcontfun (u n) - v)" for n proof - have "Norm bcontfun\<^sub>N (u n - v) = norm (Bcontfun(u n - v))" unfolding bcontfun\<^sub>N(2) using H(2) bcontfun\<^sub>N_space v_space by auto moreover have "Bcontfun(u n - v) = Bcontfun (u n) - v" unfolding minus_fun_def minus_bcontfun_def by (simp add: Bcontfun_inverse H2) ultimately show ?thesis by simp qed have "tendsto_in\<^sub>N bcontfun\<^sub>N u v" unfolding tendsto_in\<^sub>N_def *** using v Lim_null tendsto_norm_zero_iff by fastforce then show "\v\space\<^sub>N bcontfun\<^sub>N. tendsto_in\<^sub>N bcontfun\<^sub>N u v" using v_space by auto qed end diff --git a/thys/Multi_Party_Computation/Uniform_Sampling.thy b/thys/Multi_Party_Computation/Uniform_Sampling.thy --- a/thys/Multi_Party_Computation/Uniform_Sampling.thy +++ b/thys/Multi_Party_Computation/Uniform_Sampling.thy @@ -1,279 +1,279 @@ section \Uniform Sampling\ text\Here we prove different one time pad lemmas based on uniform sampling we require throughout our proofs.\ theory Uniform_Sampling imports CryptHOL.Cyclic_Group_SPMF "HOL-Number_Theory.Cong" CryptHOL.List_Bits begin text \If q is a prime we can sample from the units.\ definition sample_uniform_units :: "nat \ nat spmf" where "sample_uniform_units q = spmf_of_set ({..< q} - {0})" lemma set_spmf_sampl_uni_units [simp]: "set_spmf (sample_uniform_units q) = {..< q} - {0}" by(simp add: sample_uniform_units_def) lemma lossless_sample_uniform_units: assumes "q > 1" shows "lossless_spmf (sample_uniform_units q)" apply(simp add: sample_uniform_units_def) using assms by auto text \General lemma for mapping using uniform sampling from units.\ lemma one_time_pad_units: assumes inj_on: "inj_on f ({..s. f s) (spmf_of_set ({..s. f s) ` ({..General lemma for mapping using uniform sampling.\ lemma one_time_pad: assumes inj_on: "inj_on f {..s. f s) (spmf_of_set {..s. f s) ` {..The addition map case.\ lemma inj_add: assumes x: "x < q" and x': "x' < q" and map: "((y :: nat) + x) mod q = (y + x') mod q" shows "x = x'" proof- have aa: "((y :: nat) + x) mod q = (y + x') mod q \ x mod q = x' mod q" proof- have 4: "((y:: nat) + x) mod q = (y + x') mod q \ [((y:: nat) + x) = (y + x')] (mod q)" by(simp add: cong_def) have 5: "[((y:: nat) + x) = (y + x')] (mod q) \ [x = x'] (mod q)" by (simp add: cong_add_lcancel_nat) have 6: "[x = x'] (mod q) \ x mod q = x' mod q" by(simp add: cong_def) then show ?thesis by(simp add: map 4 5 6) qed also have bb: "x mod q = x' mod q \ x = x'" by(simp add: x x') ultimately show ?thesis by(simp add: map) qed lemma inj_uni_samp_add: "inj_on (\(b :: nat). (y + b) mod q ) {..(b :: nat). (y + b) mod q ) {..(b :: nat). (y + b) mod q) ` {..< q} = {..< q}" apply(rule endo_inj_surj) using inj by auto lemma samp_uni_plus_one_time_pad: shows "map_spmf (\b. (y + b) mod q) (sample_uniform q) = (sample_uniform q)" using inj_uni_samp_add surj_uni_samp one_time_pad by simp text \The multiplicaton map case.\ lemma inj_mult: assumes coprime: "coprime x (q::nat)" and y: "y < q" and y': "y' < q" and map: "x * y mod q = x * y' mod q" shows "y = y'" proof- have "x*y mod q = x*y' mod q \ y mod q = y' mod q" proof- have "x*y mod q = x*y' mod q \ [x*y = x*y'] (mod q)" by(simp add: cong_def) also have "[x*y = x*y'] (mod q) = [y = y'] (mod q)" by(simp add: cong_mult_lcancel_nat coprime) also have "[y = y'] (mod q) \ y mod q = y' mod q" by(simp add: cong_def) ultimately show ?thesis by(simp add: map) qed also have "y mod q = y' mod q \ y = y'" by(simp add: y y') ultimately show ?thesis by(simp add: map) qed lemma inj_on_mult: assumes coprime: "coprime x (q::nat)" shows "inj_on (\ b. x*b mod q) {.. b. x*b mod q) {.. b. x*b mod q) ` {..< q} = {..< q}" apply(rule endo_inj_surj) using coprime inj by auto lemma mult_one_time_pad: assumes coprime: "coprime x q" shows "map_spmf (\ b. x*b mod q) (sample_uniform q) = (sample_uniform q)" using inj_on_mult surj_on_mult one_time_pad coprime by simp text \The multiplication map for sampling from units.\ lemma inj_on_mult_units: assumes 1: "coprime x (q::nat)" shows "inj_on (\ b. x*b mod q) ({.. b. x*b mod q) ({.. b. x*b mod q) ` ({..b. x * b mod q) ` ({.. {..b. x * b mod q) ` ({.. {.. (nat \ nat) \ nat set \ nat" where "\x0 x1 x2. (\v3. v3 \ x2 \ x1 v3 \ x0) = (n x0 x1 x2 \ x2 \ x1 (n x0 x1 x2) \ x0)" by moura then have subset: "\N f Na. n Na f N \ N \ f (n Na f N) \ Na \ f ` N \ Na" by (meson image_subsetI) have mem_insert: "x * n ({..n. x * n mod q) ({.. {.. x * n ({..n. x * n mod q) ({.. insert 0 {..n. x * n mod q) ({.. insert 0 {..n. x * n mod q) ({.. {..n. x * n mod q) ({.. q) = (0 = q) \ (n ({..n. x * n mod q) ({.. {.. n ({..n. x * n mod q) ({.. {0}) \ n ({..n. x * n mod q) ({.. {.. x * n ({..n. x * n mod q) ({.. {.. x * n ({..n. x * n mod q) ({..n. x * n mod q) ({.. insert 0 {.. x * n ({..n. x * n mod q) ({.. {0}" then have "(\n. x * n mod q) ` ({.. {..n. x * n mod q) ` ({.. {.. (0 \ q) = (0 = q)" using mem_insert by (metis antisym_conv1 lessThan_iff mod_less_divisor singletonD) } ultimately have "(\n. x * n mod q) ` ({.. {.. n ({..n. x * n mod q) ({.. {.. x * n ({..n. x * n mod q) ({.. {..n. x * n mod q) ` ({.. {..b. x * b mod q) ({.. b. x*b mod q) (sample_uniform_units q) = sample_uniform_units q" using inj_on_mult_units surj_on_mult_units one_time_pad_units coprime by simp text \Addition and multiplication map.\ lemma samp_uni_add_mult: assumes coprime: "coprime x (q::nat)" and xa: "xa < q" and ya: "ya < q" and map: "(y + x * xa) mod q = (y + x * ya) mod q" shows "xa = ya" proof- have "(y + x * xa) mod q = (y + x * ya) mod q \ xa mod q = ya mod q" proof- have "(y + x * xa) mod q = (y + x * ya) mod q \ [y + x*xa = y + x *ya] (mod q)" using cong_def by blast also have "[y + x*xa = y + x *ya] (mod q) \ [xa = ya] (mod q)" by(simp add: cong_add_lcancel_nat)(simp add: coprime cong_mult_lcancel_nat) ultimately show ?thesis by(simp add: cong_def map) qed also have "xa mod q = ya mod q \ xa = ya" by(simp add: xa ya) ultimately show ?thesis by(simp add: map) qed lemma inj_on_add_mult: assumes coprime: "coprime x (q::nat)" shows "inj_on (\ b. (y + x*b) mod q) {.. b. (y + x*b) mod q) {.. b. (y + x*b) mod q) ` {..< q} = {..< q}" apply(rule endo_inj_surj) using coprime inj by auto lemma add_mult_one_time_pad: assumes coprime: "coprime x q" shows "map_spmf (\ b. (y + x*b) mod q) (sample_uniform q) = (sample_uniform q)" using inj_on_add_mult surj_on_add_mult one_time_pad coprime by simp text \Subtraction Map.\ lemma inj_minus: assumes x: "(x :: nat) < q" and ya: "ya < q" and map: "(y + q - x) mod q = (y + q - ya) mod q" shows "x = ya" proof- have "(y + q - x) mod q = (y + q - ya) mod q \ x mod q = ya mod q" proof- have "(y + q - x) mod q = (y + q - ya) mod q \ [y + q - x = y + q - ya] (mod q)" using cong_def by blast moreover have "[y + q - x = y + q - ya] (mod q) \ [q - x = q - ya] (mod q)" using x ya cong_add_lcancel_nat by fastforce moreover have "[y + q - x = y + q - ya] (mod q) \ [q + x = q + ya] (mod q)" by (metis add_diff_inverse_nat calculation(2) cong_add_lcancel_nat cong_add_rcancel_nat cong_sym less_imp_le_nat not_le x ya) ultimately show ?thesis by (simp add: cong_def map) qed moreover have "x mod q = ya mod q \ x = ya" by(simp add: x ya) ultimately show ?thesis by(simp add: map) qed lemma inj_on_minus: "inj_on (\(b :: nat). (y + (q - b)) mod q ) {..(b :: nat). (y + (q - b)) mod q ) {..(b :: nat). (y + (q - b)) mod q) ` {..< q} = {..< q}" apply(rule endo_inj_surj) using inj by auto lemma samp_uni_minus_one_time_pad: shows "map_spmf(\ b. (y + (q - b)) mod q) (sample_uniform q) = (sample_uniform q)" using inj_on_minus surj_on_minus one_time_pad by simp lemma not_coin_flip: "map_spmf (\ a. \ a) coin_spmf = coin_spmf" proof- have "inj_on Not {True, False}" by simp also have "Not ` {True, False} = {True, False}" by auto ultimately show ?thesis using one_time_pad by (simp add: UNIV_bool) qed lemma xor_uni_samp: "map_spmf(\ b. y \ b) (coin_spmf) = map_spmf(\ b. b) (coin_spmf)" (is "?lhs = ?rhs") proof- have rhs: "?rhs = spmf_of_set {True, False}" by (simp add: UNIV_bool insert_commute) also have "map_spmf(\ b. y \ b) (spmf_of_set {True, False}) = spmf_of_set((\ b. y \ b) ` {True, False})" by (simp add: xor_def) also have "(\ b. xor y b) ` {True, False} = {True, False}" using xor_def by auto finally show ?thesis using rhs by(simp) qed end \ No newline at end of file diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy @@ -1,1191 +1,1191 @@ (* Title: Intruder_Deduction.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Dolev-Yao Intruder Model\ theory Intruder_Deduction imports Messages More_Unification begin subsection \Syntax for the Intruder Deduction Relations\ consts INTRUDER_SYNTH::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\\<^sub>c" 50) consts INTRUDER_DEDUCT::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\" 50) subsection \Intruder Model Locale\ text \ The intruder model is parameterized over arbitrary function symbols (e.g, cryptographic operators) and variables. It requires three functions: - \arity\ that assigns an arity to each function symbol. - \public\ that partitions the function symbols into those that will be available to the intruder and those that will not. - \Ana\, the analysis interface, that defines how messages can be decomposed (e.g., decryption). \ locale intruder_model = fixes arity :: "'fun \ nat" and public :: "'fun \ bool" and Ana :: "('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" assumes Ana_keys_fv: "\t K R. Ana t = (K,R) \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" and Ana_keys_wf: "\t k K R f T. Ana t = (K,R) \ (\g S. Fun g S \ t \ length S = arity g) \ k \ set K \ Fun f T \ k \ length T = arity f" and Ana_var[simp]: "\x. Ana (Var x) = ([],[])" and Ana_fun_subterm: "\f T K R. Ana (Fun f T) = (K,R) \ set R \ set T" and Ana_subst: "\t \ K R. \Ana t = (K,R); K \ [] \ R \ []\ \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" begin lemma Ana_subterm: assumes "Ana t = (K,T)" shows "set T \ subterms t" using assms by (cases t) (simp add: psubsetI, metis Ana_fun_subterm Fun_gt_params UN_I term.order_refl params_subterms psubsetI subset_antisym subset_trans) lemma Ana_subterm': "s \ set (snd (Ana t)) \ s \ t" using Ana_subterm by (cases "Ana t") auto lemma Ana_vars: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" by (rule Ana_keys_fv[OF assms]) (use Ana_subterm[OF assms] subtermeq_vars_subset in auto) abbreviation \ where "\ \ UNIV::'var set" abbreviation \n ("\\<^sup>_") where "\\<^sup>n \ {f::'fun. arity f = n}" abbreviation \npub ("\\<^sub>p\<^sub>u\<^sub>b\<^sup>_") where "\\<^sub>p\<^sub>u\<^sub>b\<^sup>n \ {f. public f} \ \\<^sup>n" abbreviation \npriv ("\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>_") where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n \ {f. \public f} \ \\<^sup>n" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ (\n. \\<^sub>p\<^sub>u\<^sub>b\<^sup>n)" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ (\n. \\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n)" abbreviation \ where "\ \ (\n. \\<^sup>n)" abbreviation \ where "\ \ \\<^sup>0" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ {f. public f} \ \" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ {f. \public f} \ \" abbreviation \\<^sub>f where "\\<^sub>f \ \ - \" abbreviation \\<^sub>f\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ \\<^sub>f \ \\<^sub>p\<^sub>u\<^sub>b" abbreviation \\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ \\<^sub>f \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v" lemma disjoint_fun_syms: "\\<^sub>f \ \ = {}" by auto lemma id_union_univ: "\\<^sub>f \ \ = UNIV" "\ = UNIV" by auto lemma const_arity_eq_zero[dest]: "c \ \ \ arity c = 0" by simp lemma const_pub_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>u\<^sub>b \ arity c = 0 \ public c" by simp lemma const_priv_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ arity c = 0 \ \public c" by simp lemma fun_arity_gt_zero[dest]: "f \ \\<^sub>f \ arity f > 0" by fastforce lemma pub_fun_public[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ public f" by fastforce lemma pub_fun_arity_gt_zero[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ arity f > 0" by fastforce lemma \\<^sub>f_unfold: "\\<^sub>f = {f::'fun. arity f > 0}" by auto lemma \_unfold: "\ = {f::'fun. arity f = 0}" by auto lemma \pub_unfold: "\\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f = 0 \ public f}" by auto lemma \priv_unfold: "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f = 0 \ \public f}" by auto lemma \npub_unfold: "(\\<^sub>p\<^sub>u\<^sub>b\<^sup>n) = {f::'fun. arity f = n \ public f}" by auto lemma \npriv_unfold: "(\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n) = {f::'fun. arity f = n \ \public f}" by auto lemma \fpub_unfold: "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f > 0 \ public f}" by auto lemma \fpriv_unfold: "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f > 0 \ \public f}" by auto lemma \n_m_eq: "\(\\<^sup>n) \ {}; (\\<^sup>n) = (\\<^sup>m)\ \ n = m" by auto subsection \Term Well-formedness\ definition "wf\<^sub>t\<^sub>r\<^sub>m t \ \f T. Fun f T \ t \ length T = arity f" abbreviation "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T \ \t \ T. wf\<^sub>t\<^sub>r\<^sub>m t" lemma Ana_keys_wf': "Ana t = (K,T) \ wf\<^sub>t\<^sub>r\<^sub>m t \ k \ set K \ wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by metis lemma wf_trm_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp lemma wf_trm_subst_range_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp lemma wf_trm_subst_range_iff: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by force lemma wf_trm_subst_rangeD: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m (\ x)" by (metis wf_trm_subst_range_iff) lemma wf_trm_subst_rangeI[intro]: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis wf_trm_subst_range_iff) lemma wf_trmI[intro]: assumes "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" "length T = arity f" shows "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using assms unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma wf_trm_subterm: "\wf\<^sub>t\<^sub>r\<^sub>m t; s \ t\ \ wf\<^sub>t\<^sub>r\<^sub>m s" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trm_subtermeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "wf\<^sub>t\<^sub>r\<^sub>m s" proof (cases "s = t") case False thus "wf\<^sub>t\<^sub>r\<^sub>m s" using assms(2) wf_trm_subterm[OF assms(1)] by simp qed (metis assms(1)) lemma wf_trm_param: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "t \ set T" shows "wf\<^sub>t\<^sub>r\<^sub>m t" by (meson assms subtermeqI'' wf_trm_subtermeq) lemma wf_trm_param_idx: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and "i < length T" shows "wf\<^sub>t\<^sub>r\<^sub>m (T ! i)" using wf_trm_param[OF assms(1), of "T ! i"] assms(2) by fastforce lemma wf_trm_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof show "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof (induction t) case (Fun f T) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" by (meson wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: wf_trm_subst_rangeD[OF assms]) show "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" proof (induction t) case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set (map (\s. s \ \) T)" for t by (metis that wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans eval_term.simps(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: assms) qed lemma wf_trm_subst_singleton: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ Var(v := t'))" proof - have "wf\<^sub>t\<^sub>r\<^sub>m ((Var(v := t')) w)" for w using assms(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using assms(1) wf_trm_subst[of "Var(v := t')" t, OF wf_trm_subst_rangeI] by simp qed lemma wf_trm_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" using assms proof (induction t) case (Fun f T) have "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" when "t \ set T" for t using that wf_trm_param[of f "map (\t. t \ \) T"] Fun.prems by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" when "t \ set T" for t using that Fun.IH by simp moreover have "length T = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed simp lemma wf_trm_subst_rm_vars': "wf\<^sub>t\<^sub>r\<^sub>m (\ v) \ wf\<^sub>t\<^sub>r\<^sub>m (rm_vars X \ v)" by auto lemma wf_trms_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" by (metis (no_types, lifting) assms imageE wf_trm_subst) lemma wf_trms_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t rm_vars X \)" using assms wf_trm_subst_rm_vars by blast lemma wf_trms_subst_rm_vars': assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (rm_vars X \))" using assms by force lemma wf_trms_subst_compose: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms subst_img_comp_subset' wf_trm_subst by blast lemma wf_trm_subst_compose: fixes \::"('fun, 'v) subst" assumes "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" shows "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" using wf_trm_subst[of \ "\ x", OF wf_trm_subst_rangeI[OF assms(2)]] assms(1) subst_subst_compose[of "Var x" \ \] by auto lemma wf_trms_Var_range: assumes "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms by fastforce lemma wf_trms_subst_compose_Var_range: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms wf_trms_subst_compose wf_trms_Var_range by metis+ lemma wf_trm_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trms_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using wf_trm_subst_inv by fast lemma wf_trm_subterms: "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms t)" using wf_trm_subterm by blast lemma wf_trms_subterms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms\<^sub>s\<^sub>e\<^sub>t M)" using wf_trm_subterms by blast lemma wf_trm_arity: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T) \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma wf_trm_subterm_arity: "wf\<^sub>t\<^sub>r\<^sub>m t \ Fun f T \ t \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma unify_list_wf_trm: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" and "\(v,t) \ set B. wf\<^sub>t\<^sub>r\<^sub>m t" shows "\(v,t) \ set U. wf\<^sub>t\<^sub>r\<^sub>m t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B U) thus ?case by auto next case (2 f T g S E B U) have wf_fun: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" using "2.prems"(2) by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f T) (Fun g S) = Some E'" and [simp]: "f = g" "length T = length S" "E' = zip T S" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "t \ Fun f T" "t' \ Fun g S" when "(t,t') \ set E'" for t t' using that by (metis zip_arg_subterm(1), metis zip_arg_subterm(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" when "(t,t') \ set E'" for t t' using wf_trm_subterm wf_fun \f = g\ that by blast+ thus ?case using "2.IH"[OF * ** _ "2.prems"(3)] "2.prems"(2) by fastforce next case (3 v t E B) hence *: "\(w,x) \ set ((v, t) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ _ *] by metis qed next case (4 f T v E B U) hence *: "\(w,x) \ set ((v, Fun f T) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by auto have "v \ fv (Fun f T)" using "4.prems"(1) by force hence "Unification.unify (subst_list (subst v (Fun f T)) E) ((v, Fun f T)#B) = Some U" using "4.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f T)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)\] "4.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?case using "4.IH"[OF \v \ fv (Fun f T)\ _ _ *] by metis qed lemma mgu_wf_trm: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - from assms obtain \' where "subst_of \' = \" "\(v,t) \ set \'. wf\<^sub>t\<^sub>r\<^sub>m t" using unify_list_wf_trm[of "[(s,t)]" "[]"] by (auto simp: mgu_def split: option.splits) thus ?thesis proof (induction \' arbitrary: \ v rule: List.rev_induct) case (snoc x \' \ v) define \ where "\ = subst_of \'" hence "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" for v using snoc.prems(2) snoc.IH[of \] by fastforce moreover obtain w t where x: "x = (w,t)" by (metis surj_pair) hence \: "\ = Var(w := t) \\<^sub>s \" using snoc.prems(1) by (simp add: subst_def \_def) moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" using snoc.prems(2) x by auto ultimately show ?case using wf_trm_subst[of _ t] unfolding subst_compose_def by auto qed (simp add: wf\<^sub>t\<^sub>r\<^sub>m_def) qed lemma mgu_wf_trms: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF assms] by simp subsection \Definitions: Intruder Deduction Relations\ text \ A standard Dolev-Yao intruder. \ inductive intruder_deduct::"('fun,'var) terms \ ('fun,'var) term \ bool" where Axiom[simp]: "t \ M \ intruder_deduct M t" | Compose[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_deduct M t\ \ intruder_deduct M (Fun f T)" | Decompose: "\intruder_deduct M t; Ana t = (K, T); \k. k \ set K \ intruder_deduct M k; t\<^sub>i \ set T\ \ intruder_deduct M t\<^sub>i" text \ A variant of the intruder relation which limits the intruder to composition only. \ inductive intruder_synth::"('fun,'var) terms \ ('fun,'var) term \ bool" where AxiomC[simp]: "t \ M \ intruder_synth M t" | ComposeC[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_synth M t\ \ intruder_synth M (Fun f T)" adhoc_overloading INTRUDER_DEDUCT intruder_deduct adhoc_overloading INTRUDER_SYNTH intruder_synth lemma intruder_deduct_induct[consumes 1, case_names Axiom Compose Decompose]: assumes "M \ t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \ t; \t. t \ set T \ P M t\ \ P M (Fun f T)" "\t K T t\<^sub>i. \M \ t; P M t; Ana t = (K, T); \k. k \ set K \ M \ k; \k. k \ set K \ P M k; t\<^sub>i \ set T\ \ P M t\<^sub>i" shows "P M t" using assms by (induct rule: intruder_deduct.induct) blast+ lemma intruder_synth_induct[consumes 1, case_names AxiomC ComposeC]: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \\<^sub>c t; \t. t \ set T \ P M t\ \ P M (Fun f T)" shows "P M t" using assms by (induct rule: intruder_synth.induct) auto subsection \Definitions: Analyzed Knowledge and Public Ground Well-formed Terms (PGWTs)\ definition analyzed::"('fun,'var) terms \ bool" where "analyzed M \ \t. M \ t \ M \\<^sub>c t" definition analyzed_in where "analyzed_in t M \ \K R. (Ana t = (K,R) \ (\k \ set K. M \\<^sub>c k)) \ (\r \ set R. M \\<^sub>c r)" definition decomp_closure::"('fun,'var) terms \ ('fun,'var) terms \ bool" where "decomp_closure M M' \ \t. M \ t \ (\t' \ M. t \ t') \ t \ M'" inductive public_ground_wf_term::"('fun,'var) term \ bool" where PGWT[simp]: "\public f; arity f = length T; \t. t \ set T \ public_ground_wf_term t\ \ public_ground_wf_term (Fun f T)" abbreviation "public_ground_wf_terms \ {t. public_ground_wf_term t}" lemma public_const_deduct: assumes "c \ \\<^sub>p\<^sub>u\<^sub>b" shows "M \ Fun c []" "M \\<^sub>c Fun c []" proof - have "arity c = 0" "public c" using const_arity_eq_zero \c \ \\<^sub>p\<^sub>u\<^sub>b\ by auto thus "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_synth.ComposeC[OF _ \public c\, of "[]"] intruder_deduct.Compose[OF _ \public c\, of "[]"] by auto qed lemma public_const_deduct'[simp]: assumes "arity c = 0" "public c" shows "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_deduct.Compose[of "[]" c] intruder_synth.ComposeC[of "[]" c] assms by simp_all lemma private_fun_deduct_in_ik: assumes t: "M \ t" "Fun f T \ subterms t" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" using t proof (induction t rule: intruder_deduct.induct) case Decompose thus ?case by (meson Ana_subterm psubsetD term.order_trans) qed (auto simp add: f in_subterms_Union) lemma private_fun_deduct_in_ik': assumes t: "M \ Fun f T" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" by (rule private_fun_deduct_in_ik[OF t term.order_refl f]) lemma pgwt_public: "\public_ground_wf_term t; Fun f T \ t\ \ public f" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_ground: "public_ground_wf_term t \ fv t = {}" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_fun: "public_ground_wf_term t \ \f T. t = Fun f T" using pgwt_ground[of t] by (cases t) auto lemma pgwt_arity: "\public_ground_wf_term t; Fun f T \ t\ \ arity f = length T" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_wellformed: "public_ground_wf_term t \ wf\<^sub>t\<^sub>r\<^sub>m t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_deducible: "public_ground_wf_term t \ M \\<^sub>c t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_is_empty_synth: "public_ground_wf_term t \ {} \\<^sub>c t" proof - { fix M::"('fun,'var) term set" assume "M \\<^sub>c t" "M = {}" hence "public_ground_wf_term t" by (induct t rule: intruder_synth.induct) auto } thus ?thesis using pgwt_deducible by auto qed lemma ideduct_synth_subst_apply: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "{} \\<^sub>c t" "\v. M \\<^sub>c \ v" shows "M \\<^sub>c t \ \" proof - { fix M'::"('fun,'var) term set" assume "M' \\<^sub>c t" "M' = {}" hence "M \\<^sub>c t \ \" proof (induction t rule: intruder_synth.induct) case (ComposeC T f M') hence "length (map (\t. t \ \) T) = arity f" "\x. x \ set (map (\t. t \ \) T) \ M \\<^sub>c x" by auto thus ?case using intruder_synth.ComposeC[of "map (\t. t \ \) T" f M] \public f\ by fastforce qed simp } thus ?thesis using assms by metis qed subsection \Lemmata: Monotonicity, Deduction of Private Constants, etc.\ context begin lemma ideduct_mono: "\M \ t; M \ M'\ \ M' \ t" proof (induction rule: intruder_deduct.induct) case (Decompose M t K T t\<^sub>i) have "\k. k \ set K \ M' \ k" using Decompose.IH \M \ M'\ by simp moreover have "M' \ t" using Decompose.IH \M \ M'\ by simp ultimately show ?case using Decompose.hyps intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_mono: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \\<^sub>c t; M \ M'\ \ M' \\<^sub>c t" by (induct rule: intruder_synth.induct) auto context begin \ \Used by \inductive_set\\ private lemma ideduct_mono_set[mono_set]: "M \ N \ M \ t \ N \ t" "M \ N \ M \\<^sub>c t \ N \\<^sub>c t" using ideduct_mono ideduct_synth_mono by (blast, blast) end lemma ideduct_reduce: "\M \ M' \ t; \t'. t' \ M' \ M \ t'\ \ M \ t" proof (induction rule: intruder_deduct_induct) case Decompose thus ?case using intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_reduce: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \ M' \\<^sub>c t; \t'. t' \ M' \ M \\<^sub>c t'\ \ M \\<^sub>c t" by (induct rule: intruder_synth_induct) auto lemma ideduct_mono_eq: assumes "\t. M \ t \ M' \ t" shows "M \ N \ t \ M' \ N \ t" proof show "M \ N \ t \ M' \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M") case True hence "M \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M' t "M' \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M' \ N" t K T] by auto qed show "M' \ N \ t \ M \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M'") case True hence "M' \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M t "M \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M \ N" t K T] by auto qed qed lemma deduct_synth_subterm: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "s \ subterms t" "\m \ M. \s \ subterms m. M \\<^sub>c s" shows "M \\<^sub>c s" using assms by (induct t rule: intruder_synth.induct) auto lemma deduct_if_synth[intro, dest]: "M \\<^sub>c t \ M \ t" by (induct rule: intruder_synth.induct) auto private lemma ideduct_ik_eq: assumes "\t \ M. M' \ t" shows "M' \ t \ M' \ M \ t" by (meson assms ideduct_mono ideduct_reduce sup_ge1) private lemma synth_if_deduct_empty: "{} \ t \ {} \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K M m) then obtain f T where "t = Fun f T" "m \ set T" using Ana_fun_subterm Ana_var by (cases t) fastforce+ with Decompose.IH(1) show ?case by (induction rule: intruder_synth_induct) auto qed auto private lemma ideduct_deduct_synth_mono_eq: assumes "\t. M \ t \ M' \\<^sub>c t" "M \ M'" and "\t. M' \ N \ t \ M' \ N \ D \\<^sub>c t" shows "M \ N \ t \ M' \ N \ D \\<^sub>c t" proof - have "\m \ M'. M \ m" using assms(1) by auto hence "\t. M \ t \ M' \ t" by (metis assms(1,2) deduct_if_synth ideduct_reduce sup.absorb2) hence "\t. M' \ N \ t \ M \ N \ t" by (meson ideduct_mono_eq) thus ?thesis by (meson assms(3)) qed lemma ideduct_subst: "M \ t \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof (induction t rule: intruder_deduct_induct) case (Compose T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" by auto thus ?case using intruder_deduct.Compose[OF _ Compose.hyps(2), of "map (\t. t \ \) T"] by auto next case (Decompose t K M' m') hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ k" "m' \ \ \ set (M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(2)] by fastforce+ thus ?case using intruder_deduct.Decompose[OF Decompose.IH(1)] by metis qed simp lemma ideduct_synth_subst: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" and \::"('fun,'var) subst" shows "M \\<^sub>c t \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof (induction t rule: intruder_synth_induct) case (ComposeC T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" by auto thus ?case using intruder_synth.ComposeC[OF _ ComposeC.hyps(2), of "map (\t. t \ \) T"] by auto qed simp lemma ideduct_vars: assumes "M \ t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) thus ?case using Ana_vars(2) fv_subset by blast qed auto lemma ideduct_synth_vars: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_fun_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "f \ funs_term t" "\public f" shows "f \ \(funs_term ` M)" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_const_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c Fun c []" "\public c" shows "Fun c [] \ M" using intruder_synth.cases[OF assms(1)] assms(2) by fast lemma ideduct_synth_ik_replace: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "\t \ M. N \\<^sub>c t" and "M \\<^sub>c t" shows "N \\<^sub>c t" using assms(2,1) by (induct t rule: intruder_synth.induct) auto end subsection \Lemmata: Analyzed Intruder Knowledge Closure\ lemma deducts_eq_if_analyzed: "analyzed M \ M \ t \ M \\<^sub>c t" unfolding analyzed_def by auto lemma closure_is_superset: "decomp_closure M M' \ M \ M'" unfolding decomp_closure_def by force lemma deduct_if_closure_deduct: "\M' \ t; decomp_closure M M'\ \ M \ t" proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[OF _ \Ana t = (K,T)\ _ \t\<^sub>i \ set T\] by simp qed (auto simp add: decomp_closure_def) lemma deduct_if_closure_synth: "\decomp_closure M M'; M' \\<^sub>c t\ \ M \ t" using deduct_if_closure_deduct by blast lemma decomp_closure_subterms_composable: assumes "decomp_closure M M'" and "M' \\<^sub>c t'" "M' \ t" "t \ t'" shows "M' \\<^sub>c t" using \M' \\<^sub>c t'\ assms proof (induction t' rule: intruder_synth.induct) case (AxiomC t' M') have "M \ t" using \M' \ t\ deduct_if_closure_deduct AxiomC.prems(1) by blast moreover { have "\s \ M. t' \ s" using \t' \ M'\ AxiomC.prems(1) unfolding decomp_closure_def by blast hence "\s \ M. t \ s" using \t \ t'\ term.order_trans by auto } ultimately have "t \ M'" using AxiomC.prems(1) unfolding decomp_closure_def by blast thus ?case by simp next case (ComposeC T f M') let ?t' = "Fun f T" { assume "t = ?t'" have "M' \\<^sub>c t" using \M' \\<^sub>c ?t'\ \t = ?t'\ by simp } moreover { assume "t \ ?t'" have "\x \ set T. t \ x" using \t \ ?t'\ \t \ ?t'\ by simp hence "M' \\<^sub>c t" using ComposeC.IH ComposeC.prems(1,3) ComposeC.hyps(3) by blast } ultimately show ?case using cases_simp[of "t = ?t'" "M' \\<^sub>c t"] by simp qed lemma decomp_closure_analyzed: assumes "decomp_closure M M'" shows "analyzed M'" proof - { fix t assume "M' \ t" have "M' \\<^sub>c t" using \M' \ t\ assms proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) hence "M' \ t\<^sub>i" using Decompose.hyps intruder_deduct.Decompose by blast moreover have "t\<^sub>i \ t" using Decompose.hyps(4) Ana_subterm[OF Decompose.hyps(2)] by blast moreover have "M' \\<^sub>c t" using Decompose.IH(1) Decompose.prems by blast ultimately show "M' \\<^sub>c t\<^sub>i" using decomp_closure_subterms_composable Decompose.prems by blast qed auto } moreover have "\t. M \\<^sub>c t \ M \ t" by auto ultimately show ?thesis by (auto simp add: decomp_closure_def analyzed_def) qed lemma analyzed_if_all_analyzed_in: assumes M: "\t \ M. analyzed_in t M" shows "analyzed M" proof (unfold analyzed_def, intro allI iffI) fix t assume t: "M \ t" thus "M \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) { assume "t \ M" hence ?case using M Decompose.IH(2) Decompose.hyps(2,4) unfolding analyzed_in_def by fastforce } moreover { fix f S assume "t = Fun f S" "\s. s \ set S \ M \\<^sub>c s" hence ?case using Ana_fun_subterm[of f S] Decompose.hyps(2,4) by blast } ultimately show ?case using intruder_synth.cases[OF Decompose.IH(1), of ?case] by blast qed simp_all qed auto lemma analyzed_is_all_analyzed_in: "(\t \ M. analyzed_in t M) \ analyzed M" proof show "analyzed M \ \t \ M. analyzed_in t M" unfolding analyzed_in_def analyzed_def by (auto intro: intruder_deduct.Decompose[OF intruder_deduct.Axiom]) qed (rule analyzed_if_all_analyzed_in) lemma ik_has_synth_ik_closure: fixes M :: "('fun,'var) terms" shows "\M'. (\t. M \ t \ M' \\<^sub>c t) \ decomp_closure M M' \ (finite M \ finite M')" proof - let ?M' = "{t. M \ t \ (\t' \ M. t \ t')}" have M'_closes: "decomp_closure M ?M'" unfolding decomp_closure_def by simp hence "M \ ?M'" using closure_is_superset by simp have "\t. ?M' \\<^sub>c t \ M \ t" using deduct_if_closure_synth[OF M'_closes] by blast moreover have "\t. M \ t \ ?M' \ t" using ideduct_mono[OF _ \M \ ?M'\] by simp moreover have "analyzed ?M'" using decomp_closure_analyzed[OF M'_closes] . ultimately have "\t. M \ t \ ?M' \\<^sub>c t" unfolding analyzed_def by blast moreover have "finite M \ finite ?M'" by auto ultimately show ?thesis using M'_closes by blast qed lemma deducts_eq_if_empty_ik: "{} \ t \ {} \\<^sub>c t" using analyzed_is_all_analyzed_in[of "{}"] deducts_eq_if_analyzed[of "{}" t] by blast subsection \Intruder Variants: Numbered and Composition-Restricted Intruder Deduction Relations\ text \ A variant of the intruder relation which restricts composition to only those terms that satisfy a given predicate Q. \ inductive intruder_deduct_restricted:: "('fun,'var) terms \ (('fun,'var) term \ bool) \ ('fun,'var) term \ bool" ("\_;_\ \\<^sub>r _" 50) where AxiomR[simp]: "t \ M \ \M; Q\ \\<^sub>r t" | ComposeR[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; Q (Fun f T)\ \ \M; Q\ \\<^sub>r Fun f T" | DecomposeR: "\\M; Q\ \\<^sub>r t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; t\<^sub>i \ set T\ \ \M; Q\ \\<^sub>r t\<^sub>i" text \ A variant of the intruder relation equipped with a number representing the height of the derivation tree (i.e., \\M; k\ \\<^sub>n t\ iff k is the maximum number of applications of the compose and decompose rules in any path of the derivation tree for \M \ t\). \ inductive intruder_deduct_num:: "('fun,'var) terms \ nat \ ('fun,'var) term \ bool" ("\_; _\ \\<^sub>n _" 50) where AxiomN[simp]: "t \ M \ \M; 0\ \\<^sub>n t" | ComposeN[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t\ \ \M; Suc (Max (insert 0 (steps ` set T)))\ \\<^sub>n Fun f T" | DecomposeN: "\\M; n\ \\<^sub>n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T\ \ \M; Suc (Max (insert n (steps ` set K)))\ \\<^sub>n t\<^sub>i" lemma intruder_deduct_restricted_induct[consumes 1, case_names AxiomR ComposeR DecomposeR]: assumes "\M; Q\ \\<^sub>r t" "\t. t \ M \ P M Q t" "\T f. \length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; \t. t \ set T \ P M Q t; Q (Fun f T) \ \ P M Q (Fun f T)" "\t K T t\<^sub>i. \\M; Q\ \\<^sub>r t; P M Q t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; \k. k \ set K \ P M Q k; t\<^sub>i \ set T\ \ P M Q t\<^sub>i" shows "P M Q t" using assms by (induct t rule: intruder_deduct_restricted.induct) blast+ lemma intruder_deduct_num_induct[consumes 1, case_names AxiomN ComposeN DecomposeN]: assumes "\M; n\ \\<^sub>n t" "\t. t \ M \ P M 0 t" "\T f steps. \length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t; \t. t \ set T \ P M (steps t) t\ \ P M (Suc (Max (insert 0 (steps ` set T)))) (Fun f T)" "\t K T t\<^sub>i steps n. \\M; n\ \\<^sub>n t; P M n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T; \k. k \ set K \ P M (steps k) k\ \ P M (Suc (Max (insert n (steps ` set K)))) t\<^sub>i" shows "P M n t" using assms by (induct rule: intruder_deduct_num.induct) blast+ lemma ideduct_restricted_mono: "\\M; P\ \\<^sub>r t; M \ M'\ \ \M'; P\ \\<^sub>r t" proof (induction rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) have "\k. k \ set K \ \M'; P\ \\<^sub>r k" using DecomposeR.IH \M \ M'\ by simp moreover have "\M'; P\ \\<^sub>r t" using DecomposeR.IH \M \ M'\ by simp ultimately show ?case using DecomposeR intruder_deduct_restricted.DecomposeR[OF _ DecomposeR.hyps(2) _ DecomposeR.hyps(4)] by blast qed auto subsection \Lemmata: Intruder Deduction Equivalences\ lemma deduct_if_restricted_deduct: "\M;P\ \\<^sub>r m \ M \ m" proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose by blast qed simp_all lemma restricted_deduct_if_restricted_ik: assumes "\M;P\ \\<^sub>r m" "\m \ M. P m" and P: "\t t'. P t \ t' \ t \ P t'" shows "P m" using assms(1) proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) obtain f S where "t = Fun f S" using Ana_var \t\<^sub>i \ set T\ \Ana t = (K, T)\ by (cases t) auto thus ?case using DecomposeR assms(2) P Ana_subterm by blast qed (simp_all add: assms(2)) lemma deduct_restricted_if_synth: assumes P: "P m" "\t t'. P t \ t' \ t \ P t'" and m: "M \\<^sub>c m" shows "\M; P\ \\<^sub>r m" using m P(1) proof (induction m rule: intruder_synth_induct) case (ComposeC T f) hence "\M; P\ \\<^sub>r t" when t: "t \ set T" for t using t P(2) subtermeqI''[of _ T f] by fastforce thus ?case using intruder_deduct_restricted.ComposeR[OF ComposeC.hyps(1,2)] ComposeC.prems(1) by metis qed simp lemma deduct_zero_in_ik: assumes "\M; 0\ \\<^sub>n t" shows "t \ M" proof - { fix k assume "\M; k\ \\<^sub>n t" hence "k > 0 \ t \ M" by (induct t) auto } thus ?thesis using assms by auto qed lemma deduct_if_deduct_num: "\M; k\ \\<^sub>n t \ M \ t" by (induct t rule: intruder_deduct_num.induct) (metis intruder_deduct.Axiom, metis intruder_deduct.Compose, metis intruder_deduct.Decompose) lemma deduct_num_if_deduct: "M \ t \ \k. \M; k\ \\<^sub>n t" proof (induction t rule: intruder_deduct_induct) case (Compose T f) - then obtain steps where *: "\t \ set T. \M; steps t\ \\<^sub>n t" by moura + then obtain steps where *: "\t \ set T. \M; steps t\ \\<^sub>n t" by atomize_elim metis then obtain n where "\t \ set T. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` set T"] by auto thus ?case using ComposeN[OF Compose.hyps(1,2), of M steps] * by force next case (Decompose t K T t\<^sub>i) hence "\u. u \ insert t (set K) \ \k. \M; k\ \\<^sub>n u" by auto - then obtain steps where *: "\M; steps t\ \\<^sub>n t" "\t \ set K. \M; steps t\ \\<^sub>n t" by moura + then obtain steps where *: "\M; steps t\ \\<^sub>n t" "\t \ set K. \M; steps t\ \\<^sub>n t" by (metis insert_iff) then obtain n where "steps t \ n" "\t \ set K. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` insert t (set K)"] by auto thus ?case using DecomposeN[OF _ Decompose.hyps(2) _ Decompose.hyps(4), of M _ steps] * by force qed (metis AxiomN) lemma deduct_normalize: assumes M: "\m \ M. \f T. Fun f T \ m \ P f T" and t: "\M; k\ \\<^sub>n t" "Fun f T \ t" "\P f T" shows "\l \ k. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using t proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case using M by auto next case (ComposeN T' f' steps) thus ?case proof (cases "Fun f' T' = Fun f T") case True hence "\M; Suc (Max (insert 0 (steps ` set T')))\ \\<^sub>n Fun f T" "T = T'" using intruder_deduct_num.ComposeN[OF ComposeN.hyps] by auto moreover have "\t. t \ set T \ \M; steps t\ \\<^sub>n t" using True ComposeN.hyps(3) by auto moreover have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto ultimately show ?thesis by auto next case False then obtain t' where t': "t' \ set T'" "Fun f T \ t'" using ComposeN by auto hence "\l \ steps t'. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using ComposeN.IH[OF _ _ ComposeN.prems(2)] by auto moreover have "steps t' < Suc (Max (insert 0 (steps ` set T')))" using Max_less_iff[of "insert 0 (steps ` set T')" "Suc (Max (insert 0 (steps ` set T')))"] using t'(1) by auto ultimately show ?thesis using ComposeN.hyps(3)[OF t'(1)] by (meson Suc_le_eq le_Suc_eq le_trans) qed next case (DecomposeN t K T' t\<^sub>i steps n) hence *: "Fun f T \ t" using term.order_trans[of "Fun f T" t\<^sub>i t] Ana_subterm[of t K T'] by blast have "\l \ n. (\M; l\ \\<^sub>n Fun f T) \ (\t' \ set T. \j < l. \M; j\ \\<^sub>n t')" using DecomposeN.IH(1)[OF * DecomposeN.prems(2)] by auto moreover have "n < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto ultimately show ?case using DecomposeN.hyps(4) by (meson Suc_le_eq le_Suc_eq le_trans) qed lemma deduct_inv: assumes "\M; n\ \\<^sub>n t" shows "t \ M \ (\f T. t = Fun f T \ public f \ length T = arity f \ (\t \ set T. \l < n. \M; l\ \\<^sub>n t)) \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ t \ set (snd (Ana m)))" (is "?P t n \ ?Q t n \ ?R t n") using assms proof (induction n arbitrary: t rule: nat_less_induct) case (1 n t) thus ?case proof (cases n) case 0 hence "t \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n t" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ ?P x m \ ?Q x m \ ?R x m" using "1.prems" "1.IH" by blast+ hence "?P t (Suc n') \ ?Q t (Suc n') \ ?R t (Suc n')" proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case by simp next case (ComposeN T f steps) have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto thus ?case using ComposeN.hyps by metis next case (DecomposeN t K T t\<^sub>i steps n) have 0: "n < Suc (Max (insert n (steps ` set K)))" "\k. k \ set K \ steps k < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto have IH1: "?P t j \ ?Q t j \ ?R t j" when jt: "j < n" "\M; j\ \\<^sub>n t" for j t using jt DecomposeN.prems(1) 0(1) by simp have IH2: "?P t n \ ?Q t n \ ?R t n" using DecomposeN.IH(1) IH1 by simp have 1: "\k \ set (fst (Ana t)). \l < Suc (Max (insert n (steps ` set K))). \M; l\ \\<^sub>n k" using DecomposeN.hyps(1,2,3) 0(2) by auto have 2: "t\<^sub>i \ set (snd (Ana t))" using DecomposeN.hyps(2,4) by fastforce have 3: "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" when "t \ set (snd (Ana m))" "m \\<^sub>s\<^sub>e\<^sub>t M" for m using that(1) Ana_subterm[of m _ "snd (Ana m)"] in_subterms_subset_Union[OF that(2)] by (metis (no_types, lifting) prod.collapse psubsetD subsetCE subsetD) have 4: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?R t n" using that 0(1) 1 2 3 DecomposeN.hyps(1) by (metis (no_types, lifting)) have 5: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?P t n" using that 0(1) 1 2 DecomposeN.hyps(1) by blast have 6: ?case when *: "?Q t n" proof - obtain g S where g: "t = Fun g S" "public g" "length S = arity g" "\t \ set S. \l < n. \M; l\ \\<^sub>n t" - using * by moura + using * by atomize_elim auto then obtain l where l: "l < n" "\M; l\ \\<^sub>n t\<^sub>i" using 0(1) DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] by blast have **: "l < Suc (Max (insert n (steps ` set K)))" using l(1) 0(1) by simp show ?thesis using IH1[OF l] less_trans[OF _ **] by fastforce qed show ?case using IH2 4 5 6 by argo qed thus ?thesis using Suc by fast qed qed lemma deduct_inv': assumes "M \ Fun f ts" shows "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. M \ t)" proof - obtain k where k: "intruder_deduct_num M k (Fun f ts)" using deduct_num_if_deduct[OF assms] by fast have "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. \l. intruder_deduct_num M l t)" using deduct_inv[OF k] Ana_subterm'[of "Fun f ts"] in_subterms_subset_Union by blast thus ?thesis using deduct_if_deduct_num by blast qed lemma restricted_deduct_if_deduct: assumes M: "\m \ M. \f T. Fun f T \ m \ P (Fun f T)" and P_subterm: "\f T t. M \ Fun f T \ P (Fun f T) \ t \ set T \ P t" and P_Ana_key: "\t K T k. M \ t \ P t \ Ana t = (K, T) \ M \ k \ k \ set K \ P k" and m: "M \ m" "P m" shows "\M; P\ \\<^sub>r m" proof - { fix k assume "\M; k\ \\<^sub>n m" hence ?thesis using m(2) proof (induction k arbitrary: m rule: nat_less_induct) case (1 n m) thus ?case proof (cases n) case 0 hence "m \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n m" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ P x \ \M;P\ \\<^sub>r x" using "1.prems" "1.IH" by blast+ thus ?thesis using "1.prems"(2) proof (induction m rule: intruder_deduct_num_induct) case (ComposeN T f steps) have *: "steps t < Suc (Max (insert 0 (steps ` set T)))" when "t \ set T" for t using Max_less_iff[of "insert 0 (steps ` set T)"] that by blast have **: "P t" when "t \ set T" for t using P_subterm ComposeN.prems(2) that Fun_param_is_subterm[OF that] intruder_deduct.Compose[OF ComposeN.hyps(1,2)] deduct_if_deduct_num[OF ComposeN.hyps(3)] by blast have "\M; P\ \\<^sub>r t" when "t \ set T" for t using ComposeN.prems(1) ComposeN.hyps(3)[OF that] *[OF that] **[OF that] by blast thus ?case by (metis intruder_deduct_restricted.ComposeR[OF ComposeN.hyps(1,2)] ComposeN.prems(2)) next case (DecomposeN t K T t\<^sub>i steps l) show ?case proof (cases "P t") case True hence "\k. k \ set K \ P k" using P_Ana_key DecomposeN.hyps(1,2,3) deduct_if_deduct_num by blast moreover have "\k m x. k \ set K \ m < steps k \ \M; m\ \\<^sub>n x \ P x \ \M;P\ \\<^sub>r x" proof - fix k m x assume *: "k \ set K" "m < steps k" "\M; m\ \\<^sub>n x" "P x" have "steps k \ insert l (steps ` set K)" using *(1) by simp hence "m < Suc (Max (insert l (steps ` set K)))" using less_trans[OF *(2), of "Suc (Max (insert l (steps ` set K)))"] Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by auto thus "\M;P\ \\<^sub>r x" using DecomposeN.prems(1) *(3,4) by simp qed ultimately have "\k. k \ set K \ \M; P\ \\<^sub>r k" using DecomposeN.IH(2) by auto moreover have "\M; P\ \\<^sub>r t" using True DecomposeN.prems(1) DecomposeN.hyps(1) le_imp_less_Suc Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by blast ultimately show ?thesis using intruder_deduct_restricted.DecomposeR[OF _ DecomposeN.hyps(2) _ DecomposeN.hyps(4)] by metis next case False - obtain g S where gS: "t = Fun g S" using DecomposeN.hyps(2,4) by (cases t) moura+ + obtain g S where gS: "t = Fun g S" using DecomposeN.hyps(2,4) by (cases t) auto hence *: "Fun g S \ t" "\P (Fun g S)" using False by force+ have "\jM; j\ \\<^sub>n t\<^sub>i" using gS DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] deduct_normalize[of M "\f T. P (Fun f T)", OF M DecomposeN.hyps(1) *] by force hence "\jM; j\ \\<^sub>n t\<^sub>i" using Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] less_trans[of _ l "Suc (Max (insert l (steps ` set K)))"] by blast thus ?thesis using DecomposeN.prems(1,2) by meson qed qed auto qed qed } thus ?thesis using deduct_num_if_deduct m(1) by metis qed lemma restricted_deduct_if_deduct': assumes "\m \ M. P m" and "\t t'. P t \ t' \ t \ P t'" and "\t K T k. P t \ Ana t = (K, T) \ k \ set K \ P k" and "M \ m" "P m" shows "\M; P\ \\<^sub>r m" using restricted_deduct_if_deduct[of M P m] assms by blast lemma private_const_deduct: assumes c: "\public c" "M \ (Fun c []::('fun,'var) term)" shows "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. M \ m \ (\k \ set (fst (Ana m)). M \ m) \ Fun c [] \ set (snd (Ana m)))" proof - obtain n where "\M; n\ \\<^sub>n Fun c []" - using c(2) deduct_num_if_deduct by moura + using c(2) deduct_num_if_deduct by atomize_elim auto hence "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ Fun c [] \ set (snd (Ana m)))" using deduct_inv[of M n "Fun c []"] c(1) by fast thus ?thesis using deduct_if_deduct_num[of M] by blast qed lemma private_fun_deduct_in_ik'': assumes t: "M \ Fun f T" "Fun c [] \ set T" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun f T \ set (snd (Ana m))" and c: "\public c" "Fun c [] \ M" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun c [] \ set (snd (Ana m))" shows "Fun f T \ M" proof - have *: "\n. \M; n\ \\<^sub>n Fun c []" using private_const_deduct[OF c(1)] c(2,3) deduct_if_deduct_num by blast obtain n where n: "\M; n\ \\<^sub>n Fun f T" using t(1) deduct_num_if_deduct by blast show ?thesis using deduct_inv[OF n] t(2,3) * by blast qed end subsection \Executable Definitions for Code Generation\ fun intruder_synth' where "intruder_synth' pu ar M (Var x) = (Var x \ M)" | "intruder_synth' pu ar M (Fun f T) = ( Fun f T \ M \ (pu f \ length T = ar f \ list_all (intruder_synth' pu ar M) T))" definition "wf\<^sub>t\<^sub>r\<^sub>m' ar t \ (\s \ subterms t. is_Fun s \ ar (the_Fun s) = length (args s))" definition "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' ar M \ (\t \ M. wf\<^sub>t\<^sub>r\<^sub>m' ar t)" definition "analyzed_in' An pu ar t M \ (case An t of (K,T) \ (\k \ set K. intruder_synth' pu ar M k) \ (\s \ set T. intruder_synth' pu ar M s))" lemma (in intruder_model) intruder_synth'_induct[consumes 1, case_names Var Fun]: assumes "intruder_synth' public arity M t" "\x. intruder_synth' public arity M (Var x) \ P (Var x)" "\f T. (\z. z \ set T \ intruder_synth' public arity M z \ P z) \ intruder_synth' public arity M (Fun f T) \ P (Fun f T) " shows "P t" using assms by (induct public arity M t rule: intruder_synth'.induct) auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m' arity t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def wf\<^sub>t\<^sub>r\<^sub>m'_def by auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M = wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M" using wf\<^sub>t\<^sub>r\<^sub>m_code unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def by auto lemma (in intruder_model) intruder_synth_code[code_unfold]: "intruder_synth M t = intruder_synth' public arity M t" (is "?A \ ?B") proof show "?A \ ?B" proof (induction t rule: intruder_synth_induct) case (AxiomC t) thus ?case by (cases t) auto qed (fastforce simp add: list_all_iff) show "?B \ ?A" proof (induction t rule: intruder_synth'_induct) case (Fun f T) thus ?case proof (cases "Fun f T \ M") case False hence "public f" "length T = arity f" "list_all (intruder_synth' public arity M) T" using Fun.hyps by fastforce+ thus ?thesis using Fun.IH intruder_synth.ComposeC[of T f M] Ball_set[of T] by blast qed simp qed simp qed lemma (in intruder_model) analyzed_in_code[code_unfold]: "analyzed_in t M = analyzed_in' Ana public arity t M" using intruder_synth_code[of M] unfolding analyzed_in_def analyzed_in'_def by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Stateful_Strands.thy b/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Stateful_Strands.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Stateful_Strands.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Stateful_Strands.thy @@ -1,1080 +1,1080 @@ (* Title: Labeled_Stateful_Strands.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Labeled Stateful Strands\ theory Labeled_Stateful_Strands imports Stateful_Strands Labeled_Strands begin subsection \Definitions\ text\Syntax for stateful strand labels\ abbreviation Star_step ("\\, _\") where "\\, (s::('a,'b) stateful_strand_step)\ \ (\, s)" abbreviation LabelN_step ("\_, _\") where "\(l::'a), (s::('b,'c) stateful_strand_step)\ \ (ln l, s)" text\Database projection\ definition dbproj where "dbproj l D \ filter (\d. fst d = l) D" text\The type of labeled stateful strands\ type_synonym ('a,'b,'c) labeled_stateful_strand_step = "'c strand_label \ ('a,'b) stateful_strand_step" type_synonym ('a,'b,'c) labeled_stateful_strand = "('a,'b,'c) labeled_stateful_strand_step list" text\Dual strands\ fun dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b,'c) labeled_stateful_strand_step \ ('a,'b,'c) labeled_stateful_strand_step" where "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,send\t\) = (l,receive\t\)" | "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,receive\t\) = (l,send\t\)" | "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = x" definition dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t::"('a,'b,'c) labeled_stateful_strand \ ('a,'b,'c) labeled_stateful_strand" where "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ map dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p" text\Substitution application\ fun subst_apply_labeled_stateful_strand_step:: "('a,'b,'c) labeled_stateful_strand_step \ ('a,'b) subst \ ('a,'b,'c) labeled_stateful_strand_step" (infix "\\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p" 51) where "(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (l,s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" definition subst_apply_labeled_stateful_strand:: "('a,'b,'c) labeled_stateful_strand \ ('a,'b) subst \ ('a,'b,'c) labeled_stateful_strand" (infix "\\<^sub>l\<^sub>s\<^sub>s\<^sub>t" 51) where "S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) S" text\Definitions lifted from stateful strands\ abbreviation wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation db\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ db\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ db'\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation trms_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "trms_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t n S \ trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n S)" abbreviation vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation vars_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "vars_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t n S \ vars\<^sub>s\<^sub>s\<^sub>t (proj_unl n S)" abbreviation bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel S)" abbreviation fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel S)" text\Labeled set-operations\ fun setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,insert\t,s\) = {(i,t,s)}" | "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,delete\t,s\) = {(i,t,s)}" | "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\_: t \ s\) = {(i,t,s)}" | "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\_\\\: _ \\: F'\) = ((\(t,s). (i,t,s)) ` set F')" | "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = {}" definition setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ \(setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)" subsection \Minor Lemmata\ lemma in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\l ts. (l,receive\ts\) \ set A \ t \ set ts)" unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by force lemma ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_concat: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (concat xs) = \(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` set xs)" by (induct xs) auto lemma ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons[simp]: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,send\ts\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?A) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,receive\ts\)#A) = set ts \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?B) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\ac: t \ s\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?C) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,insert\t,s\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?D) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,delete\t,s\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?E) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\ac: t \ s\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?F) "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\X\\\: F \\: G\)#A) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?G) proof - note 0 = ik\<^sub>s\<^sub>s\<^sub>t_append[of _ "unlabel A"] note 1 = in_ik\<^sub>s\<^sub>s\<^sub>t_iff show ?A using 0[of "[send\ts\]"] 1[of _ "[send\ts\]"] by auto show ?B using 0[of "[receive\ts\]"] 1[of _ "[receive\ts\]"] by auto show ?C using 0[of "[\ac: t \ s\]"] 1[of _ "[\ac: t \ s\]"] by auto show ?D using 0[of "[insert\t,s\]"] 1[of _ "[insert\t,s\]"] by auto show ?E using 0[of "[delete\t,s\]"] 1[of _ "[delete\t,s\]"] by auto show ?F using 0[of "[\ac: t \ s\]"] 1[of _ "[\ac: t \ s\]"] by auto show ?G using 0[of "[\X\\\: F \\: G\]"] 1[of _ "[\X\\\: F \\: G\]"] by auto qed lemma subst_lsstp_fst_eq: "fst (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fst a" by (cases a) auto lemma subst_lsst_map_fst_eq: "map fst (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = map fst S" using subst_lsstp_fst_eq unfolding subst_apply_labeled_stateful_strand_def by auto lemma subst_lsst_nil[simp]: "[] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_labeled_stateful_strand_def) lemma subst_lsst_cons: "a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#(A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_labeled_stateful_strand_def) lemma subst_lsstp_id_subst: "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p Var = a" proof - obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?thesis unfolding a by (cases b) auto qed lemma subst_lsst_id_subst: "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t Var = A" by (induct A) (simp, metis subst_lsstp_id_subst subst_lsst_cons) lemma subst_lsst_singleton: "[(l,s)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = [(l,s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)]" by (simp add: subst_apply_labeled_stateful_strand_def) lemma subst_lsst_append: "A@B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@(B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_labeled_stateful_strand_def) lemma subst_lsst_append_inv: assumes "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B1@B2" shows "\A1 A2. A = A1@A2 \ A1 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B1 \ A2 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B2" using assms proof (induction A arbitrary: B1 B2) case (Cons a A) note prems = Cons.prems note IH = Cons.IH show ?case proof (cases B1) case Nil then obtain b B3 where "B2 = b#B3" "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = b" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B3" using prems subst_lsst_cons by fastforce thus ?thesis by (simp add: Nil subst_apply_labeled_stateful_strand_def) next case (Cons b B3) hence "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = b" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B3@B2" using prems by (simp_all add: subst_lsst_cons) thus ?thesis by (metis Cons_eq_appendI Cons IH subst_lsst_cons) qed qed (metis append_is_Nil_conv subst_lsst_nil) lemma subst_lsst_memI[intro]: "x \ set A \ x \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis image_eqI set_map subst_apply_labeled_stateful_strand_def) lemma subst_lsstpD: "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,send\ts\) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = (n,send\ts'\)" (is "?A \ ?A'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,receive\ts\) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = (n,receive\ts'\)" (is "?B \ ?B'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,\c: t \ s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,\c: t' \ s'\)" (is "?C \ ?C'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,insert\t,s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,insert\t',s'\)" (is "?D \ ?D'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,delete\t,s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,delete\t',s'\)" (is "?E \ ?E'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,\c: t \ s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,\c: t' \ s'\)" (is "?F \ ?F'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,\X\\\: F \\: G\) \ \F' G'. F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ a = (n,\X\\\: F' \\: G'\)" (is "?G \ ?G'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,\t != s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,\t' != s'\)" (is "?H \ ?H'") "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = (n,\t not in s\) \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = (n,\t' not in s'\)" (is "?I \ ?I'") proof - obtain m b where a: "a = (m,b)" by (metis surj_pair) show "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" "?F \ ?F'" "?G \ ?G'" "?H \ ?H'" "?I \ ?I'" by (cases b; auto simp add: a subst_apply_pairs_def; fail)+ qed lemma subst_lsst_memD: "(n,receive\ts\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \us. (n,receive\us\) \ set S \ ts = us \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "(n,send\ts\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \us. (n,send\us\) \ set S \ ts = us \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "(n,\ac: t \ s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,\ac: u \ v\) \ set S \ t = u \ \ \ s = v \ \" "(n,insert\t, s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,insert\u, v\) \ set S \ t = u \ \ \ s = v \ \" "(n,delete\t, s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,delete\u, v\) \ set S \ t = u \ \ \ s = v \ \" "(n,\ac: t \ s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,\ac: u \ v\) \ set S \ t = u \ \ \ s = v \ \" "(n,\X\\\: F \\: G\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \F' G'. (n,\X\\\: F' \\: G'\) \ set S \ F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "(n,\t != s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,\u != v\) \ set S \ t = u \ \ \ s = v \ \" "(n,\t not in s\) \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \u v. (n,\u not in v\) \ set S \ t = u \ \ \ s = v \ \" proof (induction S) case (Cons b S) obtain m a where a: "b = (m,a)" by (metis surj_pair) have *: "x \ set (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "x \ set (b#S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "x \ b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" for x using that by (simp add: subst_apply_labeled_stateful_strand_def) { case 1 thus ?case using Cons.IH(1)[OF *] a by (cases a) auto } { case 2 thus ?case using Cons.IH(2)[OF *] a by (cases a) auto } { case 3 thus ?case using Cons.IH(3)[OF *] a by (cases a) auto } { case 4 thus ?case using Cons.IH(4)[OF *] a by (cases a) auto } { case 5 thus ?case using Cons.IH(5)[OF *] a by (cases a) auto } { case 6 thus ?case using Cons.IH(6)[OF *] a by (cases a) auto } { case 7 thus ?case using Cons.IH(7)[OF *] a by (cases a) auto } { case 8 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "(n,\t != s\) = b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using subst_lsstpD(8)[of b \ n t s] by auto qed (use 8 Cons.IH(8)[OF *] a in auto) qed (use 8 Cons.IH(8)[OF *] a in simp_all) } { case 9 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "(n,\t not in s\) = b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using subst_lsstpD(9)[of b \ n t s] by auto qed (use 9 Cons.IH(9)[OF *] a in auto) qed (use 9 Cons.IH(9)[OF *] a in simp_all) } qed simp_all lemma subst_lsst_unlabel_cons: "unlabel ((l,b)#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#(unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by (simp add: subst_apply_labeled_stateful_strand_def) lemma subst_lsst_unlabel: "unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = unlabel A \\<^sub>s\<^sub>s\<^sub>t \" proof (induction A) case (Cons a A) then obtain l b where "a = (l,b)" by (metis surj_pair) thus ?case using Cons by (simp add: subst_apply_labeled_stateful_strand_def subst_apply_stateful_strand_def) qed simp lemma subst_lsst_unlabel_member[intro]: assumes "x \ set (unlabel A)" shows "x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof - - obtain l where x: "(l,x) \ set A" using assms unfolding unlabel_def by moura + obtain l where x: "(l,x) \ set A" using assms unfolding unlabel_def by atomize_elim auto thus ?thesis using subst_lsst_memI by (metis unlabel_def in_set_zipE subst_apply_labeled_stateful_strand_step.simps zip_map_fst_snd) qed lemma subst_lsst_prefix: assumes "prefix B (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\C. C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B \ prefix C A" using assms proof (induction A rule: List.rev_induct) case (snoc a A) thus ?case proof (cases "B = A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case False thus ?thesis using snoc by (auto simp add: subst_lsst_append[of A] subst_lsst_cons) qed auto qed simp lemma subst_lsst_tl: "tl (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = tl S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (metis map_tl subst_apply_labeled_stateful_strand_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tl: "tl (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (tl S)" by (metis map_tl dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_fst_eq: "fst (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = fst a" proof - obtain l b where "a = (l,b)" by (metis surj_pair) thus ?thesis by (cases b) auto qed lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_map_fst_eq: "map fst (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) = map fst S" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_fst_eq unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = []" by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons[simp]: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,send\ts\)#A) = (l,receive\ts\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,receive\ts\)#A) = (l,send\ts\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\a: t \ s\)#A) = (l,\a: t \ s\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,insert\t,s\)#A) = (l,insert\t,s\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,delete\t,s\)#A) = (l,delete\t,s\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\a: t \ s\)#A) = (l,\a: t \ s\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\X\\\: F \\: G\)#A) = (l,\X\\\: F \\: G\)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append[simp]: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" proof - - obtain l x where s: "s = (l,x)" by moura + obtain l x where s: "s = (l,x)" by atomize_elim auto thus ?thesis by (cases x) (auto simp add: subst_apply_labeled_stateful_strand_def) qed lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof (induction S) case (Cons s S) thus ?case using Cons dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[of s \] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) qed (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>s\<^sub>t \" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst list.simps(9) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_append: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (metis (no_types) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons list.map(1) map_append subst_apply_labeled_stateful_strand_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD: assumes "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "\b. (l,b) \ set A \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using assms proof (induction A) case (Cons c A) hence "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p c = (l,a)" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force thus ?case proof assume "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" thus ?case using Cons.IH by auto next assume a: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p c = (l,a)" obtain i b where b: "c = (i,b)" by (metis surj_pair) thus ?case using a by (cases b) auto qed qed simp lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD': assumes a: "a \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtains b where "b \ set A" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "fst a = fst b" proof - obtain l a' where a': "a = (l,a')" by (metis surj_pair) then obtain b' where b': "(l,b') \ set A" "(l,a') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,b') \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of A \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD[of l a' "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] unfolding subst_apply_labeled_stateful_strand_def by auto show thesis using that[OF b'(1) b'(2)[unfolded dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[of "(l,b')" \] a'[symmetric]], unfolded a'] by auto qed lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv: assumes "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l, a) = (k, b)" shows "l = k" and "a = receive\t\ \ b = send\t\" and "a = send\t\ \ b = receive\t\" and "(\t. a = receive\t\ \ a = send\t\) \ b = a" proof - show "l = k" using assms by (cases a) auto show "a = receive\t\ \ b = send\t\" using assms by (cases a) auto show "a = send\t\ \ b = receive\t\" using assms by (cases a) auto show "(\t. a = receive\t\ \ a = send\t\) \ b = a" using assms by (cases a) auto qed lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = A" proof (induction A) case (Cons a A) obtain l b where "a = (l,b)" by (metis surj_pair) thus ?case using Cons by (cases b) auto qed simp lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_cong: assumes "unlabel S = unlabel S'" shows "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S')" using assms proof (induction S arbitrary: S') case (Cons x S S') obtain y S'' where y: "S' = y#S''" using Cons.prems unfolding unlabel_def by force hence IH: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S'')" using Cons by (simp add: unlabel_def) have "snd x = snd y" using Cons y by simp then obtain lx ly a where a: "x = (lx,a)" "y = (ly,a)" by (metis prod.exhaust_sel) have "snd (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) = snd (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p y)" unfolding a by (cases a) simp_all thus ?case using IH unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def y by force qed (simp add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) auto qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) auto qed simp lemma bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) simp+ qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (metis unlabel_Cons(1) vars\<^sub>s\<^sub>s\<^sub>t_Cons) lemma fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (metis unlabel_Cons(1) fv\<^sub>s\<^sub>s\<^sub>t_Cons) lemma bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (metis unlabel_Cons(1) bvars\<^sub>s\<^sub>s\<^sub>t_Cons) lemma bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (metis subst_lsst_unlabel bvars\<^sub>s\<^sub>s\<^sub>t_subst) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_member: assumes "(l,x) \ set A" and "\is_Receive x" "\is_Send x" shows "(l,x) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using assms proof (induction A) case (Cons a A) thus ?case using assms(2,3) by (cases x) (auto simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) qed simp lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_member: assumes "x \ set (unlabel A)" and "\is_Receive x" "\is_Send x" shows "x \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using assms dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_member[of _ _ A] by (meson unlabel_in unlabel_mem_has_label) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff: "(l,send\ts\) \ set A \ (l,receive\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,receive\ts\) \ set A \ (l,send\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,\c: t \ s\) \ set A \ (l,\c: t \ s\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,insert\t,s\) \ set A \ (l,insert\t,s\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,delete\t,s\) \ set A \ (l,delete\t,s\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,\c: t \ s\) \ set A \ (l,\c: t \ s\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "(l,\X\\\: F \\: G\) \ set A \ (l,\X\\\: F \\: G\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) obtain j b where a: "a = (j,b)" by (metis surj_pair) { case 1 thus ?case by (cases b) (simp_all add: Cons.IH(1) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case by (cases b) (simp_all add: Cons.IH(2) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 3 thus ?case by (cases b) (simp_all add: Cons.IH(3) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 4 thus ?case by (cases b) (simp_all add: Cons.IH(4) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 5 thus ?case by (cases b) (simp_all add: Cons.IH(5) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 6 thus ?case by (cases b) (simp_all add: Cons.IH(6) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } { case 7 thus ?case by (cases b) (simp_all add: Cons.IH(7) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } qed (simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff: "send\ts\ \ set (unlabel A) \ receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "receive\ts\ \ set (unlabel A) \ send\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "\c: t \ s\ \ set (unlabel A) \ \c: t \ s\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "insert\t,s\ \ set (unlabel A) \ insert\t,s\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "delete\t,s\ \ set (unlabel A) \ delete\t,s\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "\c: t \ s\ \ set (unlabel A) \ \c: t \ s\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "\X\\\: F \\: G\ \ set (unlabel A) \ \X\\\: F \\: G\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1,2)[of _ ts A] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(3,6)[of _ c t s A] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4,5)[of _ t s A] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(7)[of _ X F G A] by (meson unlabel_in unlabel_mem_has_label)+ lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all: "list_all is_Receive (unlabel A) \ list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Send (unlabel A) \ list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Equality (unlabel A) \ list_all is_Equality (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Insert (unlabel A) \ list_all is_Insert (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Delete (unlabel A) \ list_all is_Delete (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_InSet (unlabel A) \ list_all is_InSet (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_NegChecks (unlabel A) \ list_all is_NegChecks (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Assignment (unlabel A) \ list_all is_Assignment (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Check (unlabel A) \ list_all is_Check (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Update (unlabel A) \ list_all is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" "list_all is_Check_or_Assignment (unlabel A) \ list_all is_Check_or_Assignment (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" proof (induct A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) { case 1 thus ?case using Cons.hyps(1) a by (cases b) auto } { case 2 thus ?case using Cons.hyps(2) a by (cases b) auto } { case 3 thus ?case using Cons.hyps(3) a by (cases b) auto } { case 4 thus ?case using Cons.hyps(4) a by (cases b) auto } { case 5 thus ?case using Cons.hyps(5) a by (cases b) auto } { case 6 thus ?case using Cons.hyps(6) a by (cases b) auto } { case 7 thus ?case using Cons.hyps(7) a by (cases b) auto } { case 8 thus ?case using Cons.hyps(8) a by (cases b) auto } { case 9 thus ?case using Cons.hyps(9) a by (cases b) auto } { case 10 thus ?case using Cons.hyps(10) a by (cases b) auto } { case 11 thus ?case using Cons.hyps(11) a by (cases b) auto } qed simp_all lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all_same: "list_all is_Equality (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Insert (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Delete (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_InSet (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_NegChecks (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Assignment (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Check (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Update (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" "list_all is_Check_or_Assignment (unlabel A) \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = A" proof (induct A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) { case 1 thus ?case using Cons.hyps(1) a by (cases b) auto } { case 2 thus ?case using Cons.hyps(2) a by (cases b) auto } { case 3 thus ?case using Cons.hyps(3) a by (cases b) auto } { case 4 thus ?case using Cons.hyps(4) a by (cases b) auto } { case 5 thus ?case using Cons.hyps(5) a by (cases b) auto } { case 6 thus ?case using Cons.hyps(6) a by (cases b) auto } { case 7 thus ?case using Cons.hyps(7) a by (cases b) auto } { case 8 thus ?case using Cons.hyps(8) a by (cases b) auto } { case 9 thus ?case using Cons.hyps(9) a by (cases b) auto } qed simp_all lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain: assumes "s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" shows "\l B s'. (l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,s') \ prefix (B@[(l,s')]) A" using assms proof (induction A rule: List.rev_induct) case (snoc a A) obtain i b where a: "a = (i,b)" by (metis surj_pair) show ?case using snoc proof (cases "s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))") case False thus ?thesis using a snoc.prems unlabel_append[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t [a]"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] by (cases b) (force simp add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)+ qed auto qed simp lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst: assumes "s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" shows "\l B s'. (l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s') \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s') \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" proof - obtain B l s' where B: "(l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,s')" "prefix (B@[(l,s')]) (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" - using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain[OF assms] by moura + using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain[OF assms] by atomize_elim auto obtain C where C: "C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B@[(l,s')]" - using subst_lsst_prefix[OF B(2)] by moura + using subst_lsst_prefix[OF B(2)] by atomize_elim auto obtain D u where D: "C = D@[(l,u)]" "D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = B" "[(l,u)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = [(l, s')]" using subst_lsst_prefix[OF B(2)] subst_lsst_append_inv[OF C(1)] by (auto simp add: subst_apply_labeled_stateful_strand_def) show ?thesis using B D subst_lsst_cons subst_lsst_singleton by (metis (no_types, lifting) nth_append_length) qed lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) auto qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst_cons: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis subst_lsst_unlabel trms\<^sub>s\<^sub>s\<^sub>t_subst_cons unlabel_Cons(1)) lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst: assumes "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by (metis trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms] subst_lsst_unlabel) lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S. \X. set X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ t = s \ rm_vars (set X) \" using assms proof (induction S) case (Cons a S) obtain l b where a: "a = (l,b)" by (metis surj_pair) hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst_cons by fast thus ?case proof assume *: "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF *] a by auto next assume *: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" show ?thesis using Cons.IH[OF *] a by auto qed qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'': fixes t::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" shows "\s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S. \X. set X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ t = s \ rm_vars (set X) \ \\<^sub>s \" proof - - obtain s where s: "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t = s \ \" using assms by moura + obtain s where s: "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t = s \ \" using assms by atomize_elim auto show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'[OF s(1)] s(2) by auto qed lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual_subst_cons: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" proof - obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?thesis using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons[of a A \] by (cases b) auto qed lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_funs_term: "\(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S)))) = \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t (unlabel S)))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by fast lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t: "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons by (cases b) auto qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_unlabel_append: "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) I D = db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I D)" by (metis db\<^sub>s\<^sub>s\<^sub>t_append unlabel_append) lemma db\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: "db'\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ D = db'\<^sub>s\<^sub>s\<^sub>t (unlabel (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D" proof (induction T arbitrary: D) case (Cons x T) - obtain l s where "x = (l,s)" by moura + obtain l s where "x = (l,s)" by atomize_elim auto thus ?case using Cons by (cases s) (simp_all add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) qed (simp add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) lemma labeled_list_insert_eq_cases: "d \ set (unlabel D) \ List.insert d (unlabel D) = unlabel (List.insert (i,d) D)" "(i,d) \ set D \ List.insert d (unlabel D) = unlabel (List.insert (i,d) D)" unfolding unlabel_def by (metis (no_types, opaque_lifting) List.insert_def image_eqI list.simps(9) set_map snd_conv, metis in_set_insert set_zip_rightD zip_map_fst_snd) lemma labeled_list_insert_eq_ex_cases: "List.insert d (unlabel D) = unlabel (List.insert (i,d) D) \ (\j. (j,d) \ set D \ List.insert d (unlabel D) = unlabel (List.insert (j,d) D))" using labeled_list_insert_eq_cases unfolding unlabel_def by (metis in_set_impl_in_set_zip2 length_map zip_map_fst_snd) lemma in_proj_set: assumes "\l,r\ \ set A" shows "\l,r\ \ set (proj l A)" using assms unfolding proj_def by force lemma proj_subst: "proj l (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = proj l A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof (induction A) case (Cons a A) obtain l b where "a = (l,b)" by (metis surj_pair) thus ?case using Cons unfolding proj_def subst_apply_labeled_stateful_strand_def by force qed simp lemma proj_set_subset[simp]: "set (proj n A) \ set A" unfolding proj_def by auto lemma proj_proj_set_subset[simp]: "set (proj n (proj m A)) \ set (proj n A)" "set (proj n (proj m A)) \ set (proj m A)" "set (proj_unl n (proj m A)) \ set (proj_unl n A)" "set (proj_unl n (proj m A)) \ set (proj_unl m A)" unfolding unlabel_def proj_def by auto lemma proj_mem_iff: "(ln i, d) \ set D \ (ln i, d) \ set (proj i D)" "(\, d) \ set D \ (\, d) \ set (proj i D)" unfolding proj_def by auto lemma proj_list_insert: "proj i (List.insert (ln i,d) D) = List.insert (ln i,d) (proj i D)" "proj i (List.insert (\,d) D) = List.insert (\,d) (proj i D)" "i \ j \ proj i (List.insert (ln j,d) D) = proj i D" unfolding List.insert_def proj_def by auto lemma proj_filter: "proj i [d\D. d \ set Di] = [d\proj i D. d \ set Di]" by (simp_all add: proj_def conj_commute) lemma proj_list_Cons: "proj i ((ln i,d)#D) = (ln i,d)#proj i D" "proj i ((\,d)#D) = (\,d)#proj i D" "i \ j \ proj i ((ln j,d)#D) = proj i D" unfolding List.insert_def proj_def by auto lemma proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: "proj l (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A)" proof (induction A) case (Cons a A) obtain k b where "a = (k,b)" by (metis surj_pair) thus ?case using Cons unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def by (cases b) auto qed simp lemma proj_instance_ex: assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ P \" and b: "b \ set (proj l B)" shows "\a \ set (proj l A). \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ P \" proof - obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "P \" using B b proj_set_subset by fast obtain k b' where b': "b = (k, b')" "k = (ln l) \ k = \" using b proj_in_setD by metis obtain a' where a': "a = (k, a')" using b'(1) a(2) by (cases a) simp_all show ?thesis using a a' b'(2) unfolding proj_def by auto qed lemma proj_dbproj: "dbproj (ln i) (proj i D) = dbproj (ln i) D" "dbproj \ (proj i D) = dbproj \ D" "i \ j \ dbproj (ln j) (proj i D) = []" unfolding proj_def dbproj_def by (induct D) auto lemma dbproj_Cons: "dbproj i ((i,d)#D) = (i,d)#dbproj i D" "i \ j \ dbproj j ((i,d)#D) = dbproj j D" unfolding dbproj_def by auto lemma dbproj_subset[simp]: "set (unlabel (dbproj i D)) \ set (unlabel D)" unfolding unlabel_def dbproj_def by auto lemma dbproj_subseq: assumes "Di \ set (subseqs (dbproj k D))" shows "dbproj k Di = Di" (is ?A) and "i \ k \ dbproj i Di = []" (is "i \ k \ ?B") proof - have *: "set Di \ set (dbproj k D)" using subseqs_powset[of "dbproj k D"] assms by auto thus ?A by (metis dbproj_def filter_True filter_set member_filter subsetCE) have "\j d. (j,d) \ set Di \ j = k" using * unfolding dbproj_def by auto moreover have "\j d. (j,d) \ set (dbproj i Di) \ j = i" unfolding dbproj_def by auto moreover have "\j d. (j,d) \ set (dbproj i Di) \ (j,d) \ set Di" unfolding dbproj_def by auto ultimately show "i \ k \ ?B" by (metis set_empty subrelI subset_empty) qed lemma dbproj_subseq_subset: assumes "Di \ set (subseqs (dbproj i D))" shows "set Di \ set D" using assms unfolding dbproj_def by (metis Pow_iff filter_set image_eqI member_filter subseqs_powset subsetCE subsetI) lemma dbproj_subseq_in_subseqs: assumes "Di \ set (subseqs (dbproj i D))" shows "Di \ set (subseqs D)" using assms in_set_subseqs subseq_filter_left subseq_order.dual_order.trans unfolding dbproj_def by blast lemma proj_subseq: assumes "Di \ set (subseqs (dbproj (ln j) D))" "j \ i" shows "[d\proj i D. d \ set Di] = proj i D" proof - have "set Di \ set (dbproj (ln j) D)" using subseqs_powset[of "dbproj (ln j) D"] assms by auto hence "\k d. (k,d) \ set Di \ k = ln j" unfolding dbproj_def by auto moreover have "\k d. (k,d) \ set (proj i D) \ k \ ln j" using assms(2) unfolding proj_def by auto ultimately have "\d. d \ set (proj i D) \ d \ set Di" by auto thus ?thesis by simp qed lemma unlabel_subseqsD: assumes "A \ set (subseqs (unlabel B))" shows "\C \ set (subseqs B). unlabel C = A" using assms map_subseqs unfolding unlabel_def by (metis imageE set_map) lemma unlabel_filter_eq: assumes "\(j, p) \ set A \ B. \(k, q) \ set A \ B. p = q \ j = k" (is "?P (set A)") shows "[d\unlabel A. d \ snd ` B] = unlabel [d\A. d \ B]" using assms unfolding unlabel_def proof (induction A) case (Cons a A) have "set A \ set (a#A)" "{a} \ set (a#A)" by auto hence *: "?P (set A)" "?P {a}" using Cons.prems by fast+ hence IH: "[d\map snd A . d \ snd ` B] = map snd [d\A . d \ B]" using Cons.IH by auto { assume "snd a \ snd ` B" - then obtain b where b: "b \ B" "snd a = snd b" by moura + then obtain b where b: "b \ B" "snd a = snd b" by atomize_elim auto hence "fst a = fst b" using *(2) by auto hence "a \ B" using b by (metis surjective_pairing) } hence **: "a \ B \ snd a \ snd ` B" by metis show ?case by (cases "a \ B") (simp add: ** IH)+ qed simp lemma subseqs_mem_dbproj: assumes "Di \ set (subseqs D)" "list_all (\d. fst d = i) Di" shows "Di \ set (subseqs (dbproj i D))" using assms proof (induction D arbitrary: Di) case (Cons di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) show ?case proof (cases "Di \ set (subseqs D)") case True hence "Di \ set (subseqs (dbproj i D))" using Cons.IH Cons.prems by auto thus ?thesis using subseqs_Cons unfolding dbproj_def by auto next case False then obtain Di' where Di': "Di = di#Di'" using Cons.prems(1) by (metis (mono_tags, lifting) Un_iff imageE set_append set_map subseqs.simps(2)) hence "Di' \ set (subseqs D)" using Cons.prems(1) False by (metis (no_types, lifting) UnE imageE list.inject set_append set_map subseqs.simps(2)) hence "Di' \ set (subseqs (dbproj i D))" using Cons.IH Cons.prems Di' by auto moreover have "i = j" using Di' di Cons.prems(2) by auto hence "dbproj i (di#D) = di#dbproj i D" unfolding dbproj_def by (simp add: di) ultimately show ?thesis using Di' by (metis (no_types, lifting) UnCI image_eqI set_append set_map subseqs.simps(2)) qed qed simp lemma unlabel_subst: "unlabel S \\<^sub>s\<^sub>s\<^sub>t \ = unlabel (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding unlabel_def subst_apply_stateful_strand_def subst_apply_labeled_stateful_strand_def by auto lemma subterms_subst_lsst: assumes "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\f. \ x = Fun f []) \ (\y. \ x = Var y)" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \" using subterms_subst''[OF assms(1)] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms(2)] unlabel_subst[of S \] by simp lemma subterms_subst_lsst_ik: assumes "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \" using subterms_subst''[OF assms(1)] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel S" \] unlabel_subst[of S \] by simp lemma labeled_stateful_strand_subst_comp: assumes "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = {}" shows "S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using assms proof (induction S) case (Cons s S) obtain l x where s: "s = (l,x)" by (metis surj_pair) hence IH: "S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using Cons by auto have "x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = (x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using s Cons.prems stateful_strand_step_subst_comp[of \ x \] by auto thus ?case using s IH by (simp add: subst_apply_labeled_stateful_strand_def) qed simp lemma sst_vars_proj_subset[simp]: "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "vars\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "proj_unl n A"] unfolding unlabel_def proj_def by auto lemma trms\<^sub>s\<^sub>s\<^sub>t_proj_subset[simp]: "trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A)" (is ?A) "trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \ trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" (is ?B) "trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \ trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m A)" (is ?C) proof - show ?A unfolding unlabel_def proj_def by auto show ?B using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_proj_set_subset(4)] by metis show ?C using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_proj_set_subset(3)] by metis qed lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset: "trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))" (is ?A) "trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))" (is ?B) using trms\<^sub>s\<^sub>s\<^sub>t_mono[of "proj_unl n A" "proj_unl n (A@B)"] unfolding unlabel_def proj_def by auto lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_suffix_subset: "trms\<^sub>s\<^sub>s\<^sub>t (unlabel B) \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))" "trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n B) \ trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))" using trms\<^sub>s\<^sub>s\<^sub>t_mono[of "proj_unl n B" "proj_unl n (A@B)"] unfolding unlabel_def proj_def by auto lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>pD: assumes p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" shows "fst p = fst a" (is ?P) and "is_Update (snd a) \ is_InSet (snd a) \ is_NegChecks (snd a)" (is ?Q) proof - obtain l k p' a' where a: "p = (l,p')" "a = (k,a')" by (metis surj_pair) show ?P using p a by (cases a') auto show ?Q using p a by (cases a') auto qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons[simp]: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (x#S) = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_proj_subset: "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m A)" unfolding unlabel_def proj_def proof (induction A) case (Cons a A) - obtain l b where lb: "a = (l,b)" by moura + obtain l b where lb: "a = (l,b)" by atomize_elim auto { case 1 thus ?case using Cons.IH(1) unfolding lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using Cons.IH(2) unfolding lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } { case 3 thus ?case using Cons.IH(3) unfolding lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } qed simp_all lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset: "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))" "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))" unfolding unlabel_def proj_def proof (induction A) case (Cons a A) - obtain l b where lb: "a = (l,b)" by moura + obtain l b where lb: "a = (l,b)" by atomize_elim auto { case 1 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } qed (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_suffix_subset: "setops\<^sub>s\<^sub>s\<^sub>t (unlabel B) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))" "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n B) \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))" unfolding unlabel_def proj_def proof (induction A) case (Cons a A) - obtain l b where lb: "a = (l,b)" by moura + obtain l b where lb: "a = (l,b)" by atomize_elim auto { case 1 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) } qed simp_all lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_subset: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj m (proj n A)) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A)" unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_prefix_subset: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B)" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n (A@B))" unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_suffix_subset: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B)" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n (A@B))" unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono: "set M \ set N \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t M \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t N" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label: "\list_ex (has_LabelN l) A \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l' A)" by (rule trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(2)[of l A l']]) lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label: "\list_ex (has_LabelN l) A \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \ setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l' A)" by (rule setops\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(2)[of l A l']]) lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_subset_if_no_label: "\list_ex (has_LabelN l) A \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l' A)" by (rule setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(1)[of l A l']]) lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,send\ts\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,receive\ts\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\ac: s \ t\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,insert\t,s\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {(l,t \ \,s \ \)}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,delete\t,s\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {(l,t \ \,s \ \)}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\ac: t \ s\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = {(l,t \ \,s \ \)}" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\X\\\: F \\: F'\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = ((\(t,s). (l,t \ rm_vars (set X) \,s \ rm_vars (set X) \)) ` set F')" (is "?A = ?B") proof - have "?A = (\(t,s). (l,t,s)) ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" by auto thus "?A = ?B" unfolding subst_apply_pairs_def by auto qed simp_all lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ subst_domain \ = {}" shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (\p. (fst a,snd p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof - obtain l a' where a: "a = (l,a')" by (metis surj_pair) show ?thesis proof (cases a') case (NegChecks X F G) hence *: "rm_vars (set X) \ = \" using a assms rm_vars_apply'[of \ "set X"] by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (\p. (fst a, p)) ` set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using * NegChecks a by auto moreover have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\p. (fst a, p)) ` set G" using NegChecks a by simp hence "(\p. (fst a,snd p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\p. (fst a, p \\<^sub>p \)) ` set G" by (metis (mono_tags, lifting) image_cong image_image snd_conv) hence "(\p. (fst a,snd p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\p. (fst a, p)) ` (set G \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" unfolding case_prod_unfold by auto ultimately show ?thesis by (simp add: subst_apply_pairs_def) qed (use a in simp_all) qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst': assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ subst_domain \ = {}" shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (\(i,p). (i,p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF assms] setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>pD(1) unfolding case_prod_unfold by (metis (mono_tags, lifting) image_cong) lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: assumes "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (\p. (fst p,snd p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" using assms proof (induction S) case (Cons a S) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ subst_domain \ = {}" using Cons.prems by auto hence IH: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (\p. (fst p,snd p \\<^sub>p \)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" using Cons.IH by auto show ?case using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'[OF *] IH unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold subst_lsst_cons by auto qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst: assumes p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. fst p = fst q \ snd p = snd q \\<^sub>p rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a))) \" (is "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. ?P q") proof - obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?thesis proof (cases b) case (NegChecks X F F') hence "p \ (\(t, s). (l, t \ rm_vars (set X) \, s \ rm_vars (set X) \)) ` set F'" using p a setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of l X F F' \] by blast then obtain s t where st: "(t,s) \ set F'" "p = (l, t \ rm_vars (set X) \, s \ rm_vars (set X) \)" by auto hence "(l,t,s) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst p = fst (l,t,s)" "snd p = snd (l,t,s) \\<^sub>p rm_vars (set X) \" using a NegChecks by fastforce+ moreover have "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a) = X" using NegChecks a by auto ultimately show ?thesis by blast qed (use p a in auto) qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst: assumes "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. fst p = fst q \ (\X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. snd p = snd q \\<^sub>p rm_vars X \)" (is "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. ?P A q") using assms proof (induction A) case (Cons a A) note 0 = unlabel_Cons(2)[of a A] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of "snd a" "unlabel A"] show ?case proof (cases "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)") case False hence "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons[of "a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] subst_lsst_cons[of a A \] by auto moreover have "(set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a))) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" using 0 by simp ultimately have "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. ?P (a#A) q" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst[of p a \] by blast thus ?thesis by auto qed (use Cons.IH 0 in auto) qed simp lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where "a = (l,b)" by (metis surj_pair) thus ?case using Cons unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases b) auto qed simp end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Strands.thy b/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Strands.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Strands.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Labeled_Strands.thy @@ -1,360 +1,360 @@ (* Title: Labeled_Strands.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, The University of Sheffield SPDX-License-Identifier: BSD-3-Clause *) section \Labeled Strands\ theory Labeled_Strands imports Strands_and_Constraints begin subsection \Definitions: Labeled Strands and Constraints\ datatype 'l strand_label = LabelN (the_LabelN: "'l") ("ln _") | LabelS ("\") text \Labeled strands are strands whose steps are equipped with labels\ type_synonym ('a,'b,'c) labeled_strand_step = "'c strand_label \ ('a,'b) strand_step" type_synonym ('a,'b,'c) labeled_strand = "('a,'b,'c) labeled_strand_step list" abbreviation has_LabelN where "has_LabelN n x \ fst x = ln n" abbreviation has_LabelS where "has_LabelS x \ fst x = \" definition unlabel where "unlabel S \ map snd S" definition proj where "proj n S \ filter (\s. has_LabelN n s \ has_LabelS s) S" abbreviation proj_unl where "proj_unl n S \ unlabel (proj n S)" abbreviation wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t where "wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t (unlabel S)" abbreviation subst_apply_labeled_strand_step (infix "\\<^sub>l\<^sub>s\<^sub>t\<^sub>p" 51) where "x \\<^sub>l\<^sub>s\<^sub>t\<^sub>p \ \ (case x of (l, s) \ (l, s \\<^sub>s\<^sub>t\<^sub>p \))" abbreviation subst_apply_labeled_strand (infix "\\<^sub>l\<^sub>s\<^sub>t" 51) where "S \\<^sub>l\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>l\<^sub>s\<^sub>t\<^sub>p \) S" abbreviation trms\<^sub>l\<^sub>s\<^sub>t where "trms\<^sub>l\<^sub>s\<^sub>t S \ trms\<^sub>s\<^sub>t (unlabel S)" abbreviation trms_proj\<^sub>l\<^sub>s\<^sub>t where "trms_proj\<^sub>l\<^sub>s\<^sub>t n S \ trms\<^sub>s\<^sub>t (proj_unl n S)" abbreviation vars\<^sub>l\<^sub>s\<^sub>t where "vars\<^sub>l\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t (unlabel S)" abbreviation vars_proj\<^sub>l\<^sub>s\<^sub>t where "vars_proj\<^sub>l\<^sub>s\<^sub>t n S \ vars\<^sub>s\<^sub>t (proj_unl n S)" abbreviation bvars\<^sub>l\<^sub>s\<^sub>t where "bvars\<^sub>l\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t (unlabel S)" abbreviation fv\<^sub>l\<^sub>s\<^sub>t where "fv\<^sub>l\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t (unlabel S)" abbreviation wf\<^sub>l\<^sub>s\<^sub>t where "wf\<^sub>l\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (unlabel S)" subsection \Lemmata: Projections\ lemma has_LabelS_proj_iff_not_has_LabelN: "list_all has_LabelS (proj l A) \ \list_ex (has_LabelN l) A" by (induct A) (auto simp add: proj_def) lemma proj_subset_if_no_label: assumes "\list_ex (has_LabelN l) A" shows "set (proj l A) \ set (proj l' A)" and "set (proj_unl l A) \ set (proj_unl l' A)" using assms by (induct A) (auto simp add: unlabel_def proj_def) lemma proj_in_setD: assumes a: "a \ set (proj l A)" obtains k b where "a = (k, b)" "k = (ln l) \ k = \" using that a unfolding proj_def by (cases a) auto lemma proj_set_mono: assumes "set A \ set B" shows "set (proj n A) \ set (proj n B)" and "set (proj_unl n A) \ set (proj_unl n B)" using assms unfolding proj_def unlabel_def by auto lemma unlabel_nil[simp]: "unlabel [] = []" by (simp add: unlabel_def) lemma unlabel_mono: "set A \ set B \ set (unlabel A) \ set (unlabel B)" by (auto simp add: unlabel_def) lemma unlabel_in: "(l,x) \ set A \ x \ set (unlabel A)" unfolding unlabel_def by force lemma unlabel_mem_has_label: "x \ set (unlabel A) \ \l. (l,x) \ set A" unfolding unlabel_def by auto lemma proj_ident: assumes "list_all (\s. has_LabelN l s \ has_LabelS s) S" shows "proj l S = S" using assms unfolding proj_def list_all_iff by fastforce lemma proj_elims_label: assumes "k \ l" shows "\list_ex (has_LabelN l) (proj k S)" using assms unfolding proj_def list_ex_iff by force lemma proj_nil[simp]: "proj n [] = []" "proj_unl n [] = []" unfolding unlabel_def proj_def by auto lemma singleton_lst_proj[simp]: "proj_unl l [(ln l, a)] = [a]" "l \ l' \ proj_unl l' [(ln l, a)] = []" "proj_unl l [(\, a)] = [a]" "unlabel [(l'', a)] = [a]" unfolding proj_def unlabel_def by simp_all lemma unlabel_nil_only_if_nil[simp]: "unlabel A = [] \ A = []" unfolding unlabel_def by auto lemma unlabel_Cons[simp]: "unlabel ((l,a)#A) = a#unlabel A" "unlabel (b#A) = snd b#unlabel A" unfolding unlabel_def by simp_all lemma unlabel_append[simp]: "unlabel (A@B) = unlabel A@unlabel B" unfolding unlabel_def by auto lemma proj_Cons[simp]: "proj n ((ln n,a)#A) = (ln n,a)#proj n A" "proj n ((\,a)#A) = (\,a)#proj n A" "m \ n \ proj n ((ln m,a)#A) = proj n A" "l = (ln n) \ proj n ((l,a)#A) = (l,a)#proj n A" "l = \ \ proj n ((l,a)#A) = (l,a)#proj n A" "fst b \ \ \ fst b \ (ln n) \ proj n (b#A) = proj n A" unfolding proj_def by auto lemma proj_append[simp]: "proj l (A'@B') = proj l A'@proj l B'" "proj_unl l (A@B) = proj_unl l A@proj_unl l B" unfolding proj_def unlabel_def by auto lemma proj_unl_cons[simp]: "proj_unl l ((ln l, a)#A) = a#proj_unl l A" "l \ l' \ proj_unl l' ((ln l, a)#A) = proj_unl l' A" "proj_unl l ((\, a)#A) = a#proj_unl l A" unfolding proj_def unlabel_def by simp_all lemma trms_unlabel_proj[simp]: "trms\<^sub>s\<^sub>t\<^sub>p (snd (ln l, x)) \ trms_proj\<^sub>l\<^sub>s\<^sub>t l [(ln l, x)]" by auto lemma trms_unlabel_star[simp]: "trms\<^sub>s\<^sub>t\<^sub>p (snd (\, x)) \ trms_proj\<^sub>l\<^sub>s\<^sub>t l [(\, x)]" by auto lemma trms\<^sub>l\<^sub>s\<^sub>t_union[simp]: "trms\<^sub>l\<^sub>s\<^sub>t A = (\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" proof (induction A) case (Cons a A) - obtain l s where ls: "a = (l,s)" by moura + obtain l s where ls: "a = (l,s)" by atomize_elim auto have "trms\<^sub>l\<^sub>s\<^sub>t [a] = (\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a])" proof - have *: "trms\<^sub>l\<^sub>s\<^sub>t [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by simp show ?thesis proof (cases l) case (LabelN n) hence "trms_proj\<^sub>l\<^sub>s\<^sub>t n [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by simp moreover have "\m. n \ m \ trms_proj\<^sub>l\<^sub>s\<^sub>t m [a] = {}" using ls LabelN by auto ultimately show ?thesis using * ls by fastforce next case LabelS hence "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by auto thus ?thesis using * ls by fastforce qed qed moreover have "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l (a#A) = trms_proj\<^sub>l\<^sub>s\<^sub>t l [a] \ trms_proj\<^sub>l\<^sub>s\<^sub>t l A" unfolding unlabel_def proj_def by auto hence "(\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l (a#A)) = (\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a]) \ (\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" by auto ultimately show ?case using Cons.IH ls by auto qed simp lemma trms\<^sub>l\<^sub>s\<^sub>t_append[simp]: "trms\<^sub>l\<^sub>s\<^sub>t (A@B) = trms\<^sub>l\<^sub>s\<^sub>t A \ trms\<^sub>l\<^sub>s\<^sub>t B" by (metis trms\<^sub>s\<^sub>t_append unlabel_append) lemma trms_proj\<^sub>l\<^sub>s\<^sub>t_append[simp]: "trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B) = trms_proj\<^sub>l\<^sub>s\<^sub>t l A \ trms_proj\<^sub>l\<^sub>s\<^sub>t l B" by (metis (no_types, lifting) filter_append proj_def trms\<^sub>l\<^sub>s\<^sub>t_append) lemma trms_proj\<^sub>l\<^sub>s\<^sub>t_subset[simp]: "trms_proj\<^sub>l\<^sub>s\<^sub>t l A \ trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B)" "trms_proj\<^sub>l\<^sub>s\<^sub>t l B \ trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B)" using trms_proj\<^sub>l\<^sub>s\<^sub>t_append[of l] by blast+ lemma trms\<^sub>l\<^sub>s\<^sub>t_subset[simp]: "trms\<^sub>l\<^sub>s\<^sub>t A \ trms\<^sub>l\<^sub>s\<^sub>t (A@B)" "trms\<^sub>l\<^sub>s\<^sub>t B \ trms\<^sub>l\<^sub>s\<^sub>t (A@B)" proof (induction A) case (Cons a A) - obtain l s where *: "a = (l,s)" by moura + obtain l s where *: "a = (l,s)" by atomize_elim auto { case 1 thus ?case using Cons * by auto } { case 2 thus ?case using Cons * by auto } qed simp_all lemma vars\<^sub>l\<^sub>s\<^sub>t_union: "vars\<^sub>l\<^sub>s\<^sub>t A = (\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l A)" proof (induction A) case (Cons a A) - obtain l s where ls: "a = (l,s)" by moura + obtain l s where ls: "a = (l,s)" by atomize_elim auto have "vars\<^sub>l\<^sub>s\<^sub>t [a] = (\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a])" proof - have *: "vars\<^sub>l\<^sub>s\<^sub>t [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by auto show ?thesis proof (cases l) case (LabelN n) hence "vars_proj\<^sub>l\<^sub>s\<^sub>t n [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by simp moreover have "\m. n \ m \ vars_proj\<^sub>l\<^sub>s\<^sub>t m [a] = {}" using ls LabelN by auto ultimately show ?thesis using * ls by fast next case LabelS hence "\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by auto thus ?thesis using * ls by fast qed qed moreover have "\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l (a#A) = vars_proj\<^sub>l\<^sub>s\<^sub>t l [a] \ vars_proj\<^sub>l\<^sub>s\<^sub>t l A" unfolding unlabel_def proj_def by auto hence "(\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l (a#A)) = (\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a]) \ (\l. vars_proj\<^sub>l\<^sub>s\<^sub>t l A)" using strand_vars_split(1) by auto ultimately show ?case using Cons.IH ls strand_vars_split(1) by auto qed simp lemma unlabel_Cons_inv: "unlabel A = b#B \ \A'. (\n. A = (ln n, b)#A') \ A = (\, b)#A'" proof - assume *: "unlabel A = b#B" - then obtain l A' where "A = (l,b)#A'" unfolding unlabel_def by moura + then obtain l A' where "A = (l,b)#A'" unfolding unlabel_def by atomize_elim auto thus "\A'. (\l. A = (ln l, b)#A') \ A = (\, b)#A'" by (metis strand_label.exhaust) qed lemma unlabel_snoc_inv: "unlabel A = B@[b] \ \A'. (\n. A = A'@[(ln n, b)]) \ A = A'@[(\, b)]" proof - assume *: "unlabel A = B@[b]" then obtain A' l where "A = A'@[(l,b)]" unfolding unlabel_def by (induct A rule: List.rev_induct) auto thus "\A'. (\n. A = A'@[(ln n, b)]) \ A = A'@[(\, b)]" by (cases l) auto qed lemma proj_idem[simp]: "proj l (proj l A) = proj l A" unfolding proj_def by auto lemma proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set: "ik\<^sub>s\<^sub>t (proj_unl n A) = {t. \ts. ((ln n, Receive ts) \ set A \ (\, Receive ts) \ set A) \ t \ set ts} " using ik\<^sub>s\<^sub>t_is_rcv_set unfolding unlabel_def proj_def by force lemma unlabel_ik\<^sub>s\<^sub>t_is_rcv_set: "ik\<^sub>s\<^sub>t (unlabel A) = {t | l t ts. (l, Receive ts) \ set A \ t \ set ts}" using ik\<^sub>s\<^sub>t_is_rcv_set unfolding unlabel_def by force lemma proj_ik_union_is_unlabel_ik: "ik\<^sub>s\<^sub>t (unlabel A) = (\l. ik\<^sub>s\<^sub>t (proj_unl l A))" proof show "(\l. ik\<^sub>s\<^sub>t (proj_unl l A)) \ ik\<^sub>s\<^sub>t (unlabel A)" using unlabel_ik\<^sub>s\<^sub>t_is_rcv_set[of A] proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set[of _ A] by auto show "ik\<^sub>s\<^sub>t (unlabel A) \ (\l. ik\<^sub>s\<^sub>t (proj_unl l A))" proof fix t assume "t \ ik\<^sub>s\<^sub>t (unlabel A)" then obtain l ts where "(l, Receive ts) \ set A" "t \ set ts" using ik\<^sub>s\<^sub>t_is_rcv_set unlabel_mem_has_label[of _ A] - by moura + by atomize_elim blast thus "t \ (\l. ik\<^sub>s\<^sub>t (proj_unl l A))" using proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set[of _ A] by (cases l) auto qed qed lemma proj_ik_append[simp]: "ik\<^sub>s\<^sub>t (proj_unl l (A@B)) = ik\<^sub>s\<^sub>t (proj_unl l A) \ ik\<^sub>s\<^sub>t (proj_unl l B)" using proj_append(2)[of l A B] ik_append by auto lemma proj_ik_append_subst_all: "ik\<^sub>s\<^sub>t (proj_unl l (A@B)) \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t I) \ (ik\<^sub>s\<^sub>t (proj_unl l B) \\<^sub>s\<^sub>e\<^sub>t I)" using proj_ik_append[of l] by auto lemma ik_proj_subset[simp]: "ik\<^sub>s\<^sub>t (proj_unl n A) \ trms_proj\<^sub>l\<^sub>s\<^sub>t n A" by auto lemma prefix_unlabel: "prefix A B \ prefix (unlabel A) (unlabel B)" unfolding prefix_def unlabel_def by auto lemma prefix_proj: "prefix A B \ prefix (proj n A) (proj n B)" "prefix A B \ prefix (proj_unl n A) (proj_unl n B)" unfolding prefix_def proj_def by auto lemma suffix_unlabel: "suffix A B \ suffix (unlabel A) (unlabel B)" unfolding suffix_def unlabel_def by auto lemma suffix_proj: "suffix A B \ suffix (proj n A) (proj n B)" "suffix A B \ suffix (proj_unl n A) (proj_unl n B)" unfolding suffix_def proj_def by auto subsection \Lemmata: Well-formedness\ lemma wfvarsoccs\<^sub>s\<^sub>t_proj_union: "wfvarsoccs\<^sub>s\<^sub>t (unlabel A) = (\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A))" proof (induction A) case (Cons a A) - obtain l s where ls: "a = (l,s)" by moura + obtain l s where ls: "a = (l,s)" by atomize_elim auto have "wfvarsoccs\<^sub>s\<^sub>t (unlabel [a]) = (\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]))" proof - have *: "wfvarsoccs\<^sub>s\<^sub>t (unlabel [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by auto show ?thesis proof (cases l) case (LabelN n) hence "wfvarsoccs\<^sub>s\<^sub>t (proj_unl n [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by simp moreover have "\m. n \ m \ wfvarsoccs\<^sub>s\<^sub>t (proj_unl m [a]) = {}" using ls LabelN by auto ultimately show ?thesis using * ls by fast next case LabelS hence "\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by auto thus ?thesis using * ls by fast qed qed moreover have "wfvarsoccs\<^sub>s\<^sub>t (proj_unl l (a#A)) = wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]) \ wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)" for l unfolding unlabel_def proj_def by auto hence "(\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l (a#A))) = (\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a])) \ (\l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A))" using strand_vars_split(1) by auto ultimately show ?case using Cons.IH ls strand_vars_split(1) by auto qed simp lemma wf_if_wf_proj: assumes "\l. wf\<^sub>s\<^sub>t V (proj_unl l A)" shows "wf\<^sub>s\<^sub>t V (unlabel A)" using assms proof (induction A arbitrary: V rule: List.rev_induct) case (snoc a A) hence IH: "wf\<^sub>s\<^sub>t V (unlabel A)" using proj_append(2)[of _ A] by auto obtain b l where b: "a = (ln l, b) \ a = (\, b)" by (cases a, metis strand_label.exhaust) hence *: "wf\<^sub>s\<^sub>t V (proj_unl l A@[b])" by (metis snoc.prems proj_append(2) singleton_lst_proj(1) proj_unl_cons(1,3)) thus ?case using IH b snoc.prems proj_append(2)[of l A "[a]"] unlabel_append[of A "[a]"] proof (cases b) case (Receive ts) have "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \ V" proof fix x assume "x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" hence "x \ V \ wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)" using wf_append_exec[OF *] b Receive by auto thus "x \ wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \ V" using wfvarsoccs\<^sub>s\<^sub>t_proj_union[of A] by auto qed hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t (unlabel A) \ V" using vars_snd_rcv_strand_subset2(4)[of "unlabel A"] by blast hence "wf\<^sub>s\<^sub>t V (unlabel A@[Receive ts])" by (rule wf_rcv_append'''[OF IH]) thus ?thesis using b Receive unlabel_append[of A "[a]"] by auto next case (Equality ac s t) have "fv t \ wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \ V" when "ac = Assign" proof fix x assume "x \ fv t" hence "x \ V \ wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)" using wf_append_exec[OF *] b Equality that by auto thus "x \ wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \ V" using wfvarsoccs\<^sub>s\<^sub>t_proj_union[of A] by auto qed hence "fv t \ wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t A \ V" when "ac = Assign" using vars_snd_rcv_strand_subset2(4)[of "unlabel A"] that by blast hence "wf\<^sub>s\<^sub>t V (unlabel A@[Equality ac s t])" by (cases ac) (metis wf_eq_append'''[OF IH], metis wf_eq_check_append''[OF IH]) thus ?thesis using b Equality unlabel_append[of A "[a]"] by auto qed auto qed simp end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy b/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy @@ -1,994 +1,994 @@ (* Title: Lazy_Intruder.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Lazy Intruder\ theory Lazy_Intruder imports Strands_and_Constraints Intruder_Deduction begin context intruder_model begin subsection \Definition of the Lazy Intruder\ text \The lazy intruder constraint reduction system, defined as a relation on constraint states\ inductive_set LI_rel:: "((('fun,'var) strand \ (('fun,'var) subst)) \ ('fun,'var) strand \ (('fun,'var) subst)) set" and LI_rel' (infix "\" 50) and LI_rel_trancl (infix "\\<^sup>+" 50) and LI_rel_rtrancl (infix "\\<^sup>*" 50) where "A \ B \ (A,B) \ LI_rel" | "A \\<^sup>+ B \ (A,B) \ LI_rel\<^sup>+" | "A \\<^sup>* B \ (A,B) \ LI_rel\<^sup>*" | Compose: "\simple S; length T = arity f; public f\ \ (S@Send [Fun f T]#S',\) \ (S@(map Send1 T)@S',\)" | Unify: "\simple S; Fun f T' \ ik\<^sub>s\<^sub>t S; Some \ = mgu (Fun f T) (Fun f T')\ \ (S@Send [Fun f T]#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" | Equality: "\simple S; Some \ = mgu t t'\ \ (S@Equality _ t t'#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" text \A "pre-processing step" to be applied before constraint reduction. It transforms constraints such that exactly one message is transmitted in each message transmission step. It is sound and complete and preserves the various well-formedness properties required by the lazy intruder.\ fun LI_preproc where "LI_preproc [] = []" | "LI_preproc (Send ts#S) = map Send1 ts@LI_preproc S" | "LI_preproc (Receive ts#S) = map Receive1 ts@LI_preproc S" | "LI_preproc (x#S) = x#LI_preproc S" definition LI_preproc_prop where "LI_preproc_prop S \ \ts. Send ts \ set S \ Receive ts \ set S \ (\t. ts = [t])" subsection \Lemmata: Preprocessing \ lemma LI_preproc_preproc_prop: "LI_preproc_prop (LI_preproc S)" by (induct S rule: LI_preproc.induct) (auto simp add: LI_preproc_prop_def) lemma LI_preproc_sem_eq: "\M; S\\<^sub>c \ \ \M; LI_preproc S\\<^sub>c \" (is "?A \ ?B") proof show "?A \ ?B" proof (induction S rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; LI_preproc S\\<^sub>c \" "\M; map Send1 ts\\<^sub>c \" using strand_sem_Send_map(5) by auto moreover have "ik\<^sub>s\<^sub>t (map Send1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = {}" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by fastforce ultimately show ?case using strand_sem_append(1) by simp next case (ConsRcv M ts S) hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; LI_preproc S\\<^sub>c \" "\M; map Receive1 ts\\<^sub>c \" using strand_sem_Receive_map(3) by auto moreover have "ik\<^sub>s\<^sub>t (map Receive1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = set ts \\<^sub>s\<^sub>e\<^sub>t \" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by force ultimately show ?case using strand_sem_append(1) by (simp add: Un_commute) qed simp_all show "?B \ ?A" proof (induction S arbitrary: M rule: LI_preproc.induct) case (2 ts S) have "ik\<^sub>s\<^sub>t (map Send1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = {}" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by fastforce hence "\M; S\\<^sub>c \" "\M; map Send1 ts\\<^sub>c \" using 2 strand_sem_append(1) by auto thus ?case using strand_sem_Send_map(5) by simp next case (3 ts S) have "ik\<^sub>s\<^sub>t (map Receive1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = set ts \\<^sub>s\<^sub>e\<^sub>t \" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by force hence "\M \ (set ts \\<^sub>s\<^sub>e\<^sub>t \); S\\<^sub>c \" "\M; map Receive1 ts\\<^sub>c \" using 3 strand_sem_append(1) by auto thus ?case using strand_sem_Receive_map(3) by (simp add: Un_commute) qed simp_all qed lemma LI_preproc_sem_eq': "(\ \\<^sub>c \S, \\) \ (\ \\<^sub>c \LI_preproc S, \\)" using LI_preproc_sem_eq unfolding constr_sem_c_def by simp lemma LI_preproc_vars_eq: "fv\<^sub>s\<^sub>t (LI_preproc S) = fv\<^sub>s\<^sub>t S" "bvars\<^sub>s\<^sub>t (LI_preproc S) = bvars\<^sub>s\<^sub>t S" "vars\<^sub>s\<^sub>t (LI_preproc S) = vars\<^sub>s\<^sub>t S" by (induct S rule: LI_preproc.induct) auto lemma LI_preproc_trms_eq: "trms\<^sub>s\<^sub>t (LI_preproc S) = trms\<^sub>s\<^sub>t S" by (induct S rule: LI_preproc.induct) auto lemma LI_preproc_wf\<^sub>s\<^sub>t: assumes "wf\<^sub>s\<^sub>t X S" shows "wf\<^sub>s\<^sub>t X (LI_preproc S)" using assms proof (induction S arbitrary: X rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv X ts S) hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" "wf\<^sub>s\<^sub>t X (LI_preproc S)" by auto thus ?case using wf_Receive1_prefix by simp next case (ConsSnd X ts S) hence "wf\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) (LI_preproc S)" by simp thus ?case using wf_Send1_prefix by simp qed simp_all lemma LI_preproc_preserves_wellformedness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (LI_preproc S) \" using assms LI_preproc_vars_eq[of S] LI_preproc_wf\<^sub>s\<^sub>t[of "{}" S] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by argo lemma LI_preproc_prop_SendE: assumes "LI_preproc_prop S" and "Send ts \ set S" shows "(\x. ts = [Var x]) \ (\f T. ts = [Fun f T])" proof - obtain t where "ts = [t]" using assms unfolding LI_preproc_prop_def by auto thus ?thesis by (cases t) auto qed lemma LI_preproc_prop_split: "LI_preproc_prop (S@S') \ LI_preproc_prop S \ LI_preproc_prop S'" (is "?A \ ?B") proof show "?A \ ?B" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) show "?B \ ?A" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) qed subsection \Lemma: The Lazy Intruder is Well-founded\ context begin private lemma LI_compose_measure_lt: "((S@(map Send1 T)@S',\\<^sub>1), (S@Send [Fun f T]#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" using strand_fv_card_map_fun_eq[of S f T S'] strand_size_map_fun_lt[of T f] by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) private lemma LI_unify_measure_lt: assumes "Some \ = mgu (Fun f T) t" "fv t \ fv\<^sub>s\<^sub>t S" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Send [Fun f T]#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_fun_le[of S S' f T] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv (Fun f T) \ fv t" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using assms(2) by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) have "range_vars \ \ fv (Fun f T) \ fv\<^sub>s\<^sub>t S" using assms(2) mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Send [Fun f T]#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_equality_measure_lt: assumes "Some \ = mgu t t'" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Equality a t t'#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_eq_le[of S S' a t t'] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv t \ fv t'" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using assms by auto have "range_vars \ \ fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using assms mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_in_measure: "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) thus ?case using LI_compose_measure_lt[of S T S'] by metis next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand_subset(2)[of S] by force thus ?case using LI_unify_measure_lt[OF Unify.hyps(3), of S S'] by metis qed (metis LI_equality_measure_lt) private lemma LI_in_measure_trans: "(S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" by (induction rule: trancl.induct, metis surjective_pairing LI_in_measure) (metis (no_types, lifting) surjective_pairing LI_in_measure measure\<^sub>s\<^sub>t_trans trans_def) private lemma LI_converse_wellfounded_trans: "wf ((LI_rel\<^sup>+)\)" proof - have "(LI_rel\<^sup>+)\ \ measure\<^sub>s\<^sub>t" using LI_in_measure_trans by auto thus ?thesis using measure\<^sub>s\<^sub>t_wellfounded wf_subset by metis qed private lemma LI_acyclic_trans: "acyclic (LI_rel\<^sup>+)" using wf_acyclic[OF LI_converse_wellfounded_trans] acyclic_converse by metis private lemma LI_acyclic: "acyclic LI_rel" using LI_acyclic_trans acyclic_subset by (simp add: acyclic_def) lemma LI_no_infinite_chain: "\(\f. \i. f i \\<^sup>+ f (Suc i))" proof - have "\(\f. \i. (f (Suc i), f i) \ (LI_rel\<^sup>+)\)" using wf_iff_no_infinite_down_chain LI_converse_wellfounded_trans by metis thus ?thesis by simp qed private lemma LI_unify_finite: assumes "finite M" shows "finite {((S@Send [Fun f T]#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)) | \ T'. simple S \ Fun f T' \ M \ Some \ = mgu (Fun f T) (Fun f T')}" using assms proof (induction M rule: finite_induct) case (insert m M) thus ?case proof (cases m) case (Fun g U) let ?a = "\\. ((S@Send [Fun f T]#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \))" let ?A = "\B. {?a \ | \ T'. simple S \ Fun f T' \ B \ Some \ = mgu (Fun f T) (Fun f T')}" have "?A (insert m M) = (?A M) \ (?A {m})" by auto moreover have "finite (?A {m})" proof (cases "\\. Some \ = mgu (Fun f T) (Fun g U)") case True then obtain \ where \: "Some \ = mgu (Fun f T) (Fun g U)" by blast have A_m_eq: "\\'. ?a \' \ ?A {m} \ ?a \ = ?a \'" proof - fix \' assume "?a \' \ ?A {m}" hence "\\. Some \ = mgu (Fun f T) (Fun g U) \ ?a \ = ?a \'" using \m = Fun g U\ by auto thus "?a \ = ?a \'" by (metis \ option.inject) qed have "?A {m} = {} \ ?A {m} = {?a \}" proof (cases "simple S \ ?A {m} \ {}") case True hence "simple S" "?A {m} \ {}" by meson+ hence "?A {m} = {?a \ | \. Some \ = mgu (Fun f T) (Fun g U)}" using \m = Fun g U\ by auto hence "?a \ \ ?A {m}" using \ by auto show ?thesis proof (rule ccontr) assume "\(?A {m} = {} \ ?A {m} = {?a \})" then obtain B where B: "?A {m} = insert (?a \) B" "?a \ \ B" "B \ {}" using \?A {m} \ {}\ \?a \ \ ?A {m}\ by (metis (no_types, lifting) Set.set_insert) then obtain b where b: "?a \ \ b" "b \ B" by (metis (no_types, lifting) ex_in_conv) then obtain \' where \': "b = ?a \'" using B(1) by blast moreover have "?a \' \ ?A {m}" using B(1) b(2) \' by auto hence "?a \ = ?a \'" by (blast dest!: A_m_eq) ultimately show False using b(1) by simp qed qed auto thus ?thesis by (metis (no_types, lifting) finite.emptyI finite_insert) next case False hence "?A {m} = {}" using \m = Fun g U\ by blast thus ?thesis by (metis finite.emptyI) qed ultimately show ?thesis using insert.IH by auto qed simp qed fastforce end subsection \Lemma: The Lazy Intruder Preserves Well-formedness\ context begin private lemma LI_preserves_subst_wf_single: assumes "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) { case 1 thus ?case using vars_st_snd_map by auto } { case 2 thus ?case using vars_st_snd_map by auto } { case 3 thus ?case using vars_st_snd_map by force } { case 4 thus ?case using vars_st_snd_map by auto } next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik' by blast hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Unify.hyps(3)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Unify.hyps(3)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Send [Fun f T]#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } next case (Equality S \ t t' a S' \) hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] unfolding range_vars_alt_def by fastforce have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Equality.hyps(2)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Equality.hyps(2)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Equality a t t'#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } qed private lemma LI_preserves_subst_wf: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) { case 1 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 2 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 3 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 4 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } qed metis lemma LI_preserves_wellformedness: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>2 \\<^sub>2" proof - have *: "wf\<^sub>s\<^sub>t {} S\<^sub>j" when "(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>i \\<^sub>i" for S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j using that proof (induction rule: LI_rel.induct) case (Compose S T f S' \) thus ?case by (metis wf_send_compose wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) next case (Unify S f U \ T S' \) have "fv (Fun f T) \ fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using Unify.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S') = {}" using Unify.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_unify[OF _ Unify.hyps(2) MGU_is_Unifier[OF mgu_gives_MGU], of "{}", OF _ Unify.hyps(3)[symmetric], of S'] Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) next case (Equality S \ t t' a S' \) have "fv t \ fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using Equality.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_equality[OF _ Equality.hyps(2)[symmetric], of "{}" S a S'] Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed show ?thesis using assms proof (induction rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] *[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] by (metis wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed simp qed lemma LI_preserves_trm_wf: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof - { fix S \ S' \' assume "(S,\) \ (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and *: "t \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" "t \ set S' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" for t by auto hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" when "t \ set (map Send1 T)" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?case using * by force next case (Unify S f U \ T S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f U)" using Unify.prems(1) Unify.hyps(2) wf_trm_subterm[of _ "Fun f U"] by (simp, force) hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Unify.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto next case (Equality S \ t t' a S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" by simp_all hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Equality.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto qed } with assms show ?thesis by (induction rule: rtrancl_induct2) metis+ qed lemma LI_preproc_prop_subst: "LI_preproc_prop S \ LI_preproc_prop (S \\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) lemma LI_preserves_LI_preproc_prop: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "LI_preproc_prop S\<^sub>1" shows "LI_preproc_prop S\<^sub>2" using assms proof (induction rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) hence "LI_preproc_prop S\<^sub>i" by metis with step.hyps(2) show ?case proof (induction rule: LI_rel.induct) case (Unify S f T' \ T S' \) thus ?case using LI_preproc_prop_subst LI_preproc_prop_split by (metis append.left_neutral append_Cons) next case (Equality S \ t t' uu S' \) thus ?case using LI_preproc_prop_subst LI_preproc_prop_split by (metis append.left_neutral append_Cons) qed (auto simp add: LI_preproc_prop_def) qed simp end subsection \Theorem: Soundness of the Lazy Intruder\ context begin private lemma LI_soundness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2,\\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1,\\<^sub>1\" using assms(2,1,3) proof (induction rule: LI_rel.induct) case (Compose S T f S' \) have "ik\<^sub>s\<^sub>t (map Send1 T) \\<^sub>s\<^sub>e\<^sub>t \ = {}" by fastforce hence *: "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send1 T\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" using Compose unfolding constr_sem_c_def by (force, force, fastforce) have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using *(2) Compose.hyps(2) ComposeC[OF _ Compose.hyps(3), of "map (\x. x \ \) T"] unfolding subst_compose_def by force thus "\ \\<^sub>c \S@Send [Fun f T]#S',\\" using *(1,3) \\ \\<^sub>c \S@map Send1 T@S',\\\ by (auto simp add: constr_sem_c_def) next case (Unify S f U \ T S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Unify.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have \fun_id: "Fun f U \ \ = Fun f U" "Fun f T \ \ = Fun f T" using Unify.prems(1) subst_apply_term_ident[of "Fun f U" \] fv_subset_if_in_strand_ik[of "Fun f U" S] Unify.hyps(2) fv_snd_rcv_strand_subset(2)[of S] strand_vars_split(1)[of S "Send [Fun f T]#S'"] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def apply blast using Unify.prems(1) subst_apply_term_ident[of "Fun f T" \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] apply (blast,blast) using Unify.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv (Fun f T) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using Unify.hyps(2) by force+ hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send [Fun f T]]\\<^sub>c \" proof - from Unify.hyps(2) have "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast hence "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast moreover have "Unifier \ (Fun f T) (Fun f U)" by (fact MGU_is_Unifier[OF mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]]) ultimately have "Fun f T \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" using \ by (metis \fun_id subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) strand_subst_vars_union_bound[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "subst_idem \" by (fact mgu_gives_subst_idem[OF Unify.hyps(3)[symmetric]]) moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Unify.prems(1)] wf_constr_bvars_disj'[OF Unify.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Send [Fun f T]#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send [Fun f T]]\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) next case (Equality S \ t t' a S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Equality.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have "fv t \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto moreover have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by auto ultimately have \fun_id: "t \ \ = t" "t' \ \ = t'" by auto hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] apply (blast,blast) using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv t \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \" proof - have "t \ \ = t' \ \" using MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis hence "t \ (\ \\<^sub>s \) = t' \ (\ \\<^sub>s \)" by (metis \fun_id subst_subst_compose) hence "t \ \ = t' \ \" by (metis \ subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Equality.prems(1) by (fastforce simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def simp del: subst_range.simps) hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S') = {}" by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) subst_sends_strand_fv_to_img[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Equality.prems(1)] wf_constr_bvars_disj'[OF Equality.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] mgu_gives_subst_idem[OF Equality.hyps(2)[symmetric]] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Equality a t t'#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) qed theorem LI_soundness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2, \\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" using assms(2,1,3) proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_preserves_wellformedness[OF \(LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>* (S\<^sub>i, \\<^sub>i)\] LI_soundness_single[OF _ \(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)\ \\ \\<^sub>c \S\<^sub>j, \\<^sub>j\\] by metis qed (metis LI_preproc_sem_eq') end subsection \Theorem: Completeness of the Lazy Intruder\ context begin private lemma LI_completeness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" "\simple S\<^sub>1" "LI_preproc_prop S\<^sub>1" shows "\S\<^sub>2 \\<^sub>2. (S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using not_simple_elim[OF \\simple S\<^sub>1\] proof - { \ \In this case \S\<^sub>1\ isn't simple because it contains an equality constraint, so we can simply proceed with the reduction by computing the MGU for the equation\ assume "\S' S'' a t t'. S\<^sub>1 = S'@Equality a t t'#S'' \ simple S'" - then obtain S a t t' S' where S\<^sub>1: "S\<^sub>1 = S@Equality a t t'#S'" "simple S" by moura + then obtain S a t t' S' where S\<^sub>1: "S\<^sub>1 = S@Equality a t t'#S'" "simple S" by atomize_elim force hence *: "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" "t \ \ = t' \ \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ wf_eq_fv[of "{}" S t t' S'] fv_snd_rcv_strand_subset(5)[of S] by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) from * have "Unifier \ t t'" by simp then obtain \ where \: "Some \ = mgu t t'" "subst_idem \" "subst_domain \ \ range_vars \ \ fv t \ fv t'" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ t t'\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) by auto hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Equality a t t'#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Equality[OF \simple S\ \(1)] S\<^sub>1 by metis ultimately have ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) } moreover { \ \In this case \S\<^sub>1\ isn't simple because it contains a deduction constraint for a composed term, so we must look at how this composed term is derived under the interpretation \\\\ assume "\S' S'' ts. S\<^sub>1 = S'@Send ts#S'' \ (\x. ts = [Var x]) \ simple S'" hence "\S' S'' f T. S\<^sub>1 = S'@Send [Fun f T]#S'' \ simple S'" using LI_preproc_prop_SendE[OF \LI_preproc_prop S\<^sub>1\] by fastforce - with assms obtain S f T S' where S\<^sub>1: "S\<^sub>1 = S@Send [Fun f T]#S'" "simple S" by moura + with assms obtain S f T S' where S\<^sub>1: "S\<^sub>1 = S@Send [Fun f T]#S'" "simple S" by atomize_elim auto hence "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) \ \Lemma for a common subcase\ have fun_sat: "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" when T: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have "\t. t \ set T \ \ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send1 t]\\<^sub>c \" using T by simp hence "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send1 T\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_Send_map by blast moreover have "ik\<^sub>s\<^sub>t (S@[Send1 (Fun f T)]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>t (S@(map Send1 T)) \\<^sub>s\<^sub>e\<^sub>t \" by auto hence "\ik\<^sub>s\<^sub>t (S@(map Send1 T)) \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ unfolding S\<^sub>1(1) constr_sem_c_def by force ultimately show ?thesis using \\ \\<^sub>c \S, \\<^sub>1\\ strand_sem_append(1)[of "{}" S \ "map Send1 T"] strand_sem_append(1)[of "{}" "S@map Send1 T" \ S'] unfolding constr_sem_c_def by simp qed from S\<^sub>1 \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" by (auto simp add: constr_sem_c_def) hence ?thesis proof cases \ \Case 1: \\(f(T))\ has been derived using the \AxiomC\ rule.\ case AxiomC hence ex_t: "\t. t \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = t \ \" by auto show ?thesis proof (cases "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \") \ \Case 1.1: \f(T)\ is equal to a variable in the intruder knowledge under \\\. Hence there must exists a deduction constraint in the simple prefix of the constraint in which this variable occurs/"is sent" for the first time. Since this variable itself cannot have been derived from the \AxiomC\ rule (because it must be equal under the interpretation to \f(T)\, which is by assumption not in the intruder knowledge under \\\) it must be the case that we can derive it using the \ComposeC\ rule. Hence we can apply the \Compose\ rule of the lazy intruder to \f(T)\.\ case True have "\v. Var v \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = \ v" proof - - obtain t where "t \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = t \ \" using ex_t by moura + obtain t where "t \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = t \ \" using ex_t by atomize_elim auto thus ?thesis using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ by (cases t) auto qed hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. Fun f T \ \ = \ v" using vars_subset_if_in_strand_ik2[of _ S] by fastforce then obtain v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where S: "S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ \ = \ v" "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ = \ w)" using \wf\<^sub>s\<^sub>t {} S\ wf_simple_strand_first_Send_var_split[OF _ \simple S\, of "Fun f T" \] by auto hence "\w. Var w \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Var w \ \" by force moreover have "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ Fun f T \ \ \ Fun f T' \ \" using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ S(1) by (meson contra_subsetD ik_append_subset(1)) hence "\g T'. Fun g T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Fun g T' \ \" using S(2) by simp ultimately have "\t \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. \ v \ t \ \" by (metis term.exhaust) hence "\ v \ (ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e) \\<^sub>s\<^sub>e\<^sub>t \" by auto have "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using S\<^sub>1(1) S(1) \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by (auto simp add: constr_sem_c_def) hence "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using \Fun f T \ \ = \ v\ by metis hence "length T = arity f" "public f" "\t. t \ set T \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using \Fun f T \ \ = \ v\ \\ v \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \\ intruder_synth.simps[of "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \" "\ v"] by auto hence *: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using S(1) by (auto intro: ideduct_synth_mono) hence "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" by (metis fun_sat) moreover have "(S@Send [Fun f T]#S', \\<^sub>1) \ (S@map Send1 T@S', \\<^sub>1)" by (metis LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\]) ultimately show ?thesis using S\<^sub>1 by auto next \ \Case 1.2: \\(f(T))\ can be derived from an interpreted composed term in the intruder knowledge. Use the \Unify\ rule on this composed term to further reduce the constraint.\ case False then obtain T' where t: "Fun f T' \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = Fun f T' \ \" by auto hence "fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1" using S\<^sub>1(1) fv_subset_if_in_strand_ik'[OF t(1)] fv_snd_rcv_strand_subset(2)[of S] by auto from t have "Unifier \ (Fun f T) (Fun f T')" by simp then obtain \ where \: "Some \ = mgu (Fun f T) (Fun f T')" "subst_idem \" "subst_domain \ \ range_vars \ \ fv (Fun f T) \ fv (Fun f T')" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ (Fun f T) (Fun f T')\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) \fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1\ unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Send [Fun f T]#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Unify[OF \simple S\ t(1) \(1)] S\<^sub>1 by metis ultimately show ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) qed next \ \Case 2: \\(f(T))\ has been derived using the \ComposeC\ rule. Simply use the \Compose\ rule of the lazy intruder to proceed with the reduction.\ case (ComposeC T' g) hence "f = g" "length T = arity f" "public f" and "\x. x \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c x \ \" by auto hence "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" using fun_sat by metis moreover have "(S\<^sub>1, \\<^sub>1) \ (S@(map Send1 T)@S', \\<^sub>1)" using S\<^sub>1 LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\] by metis ultimately show ?thesis by metis qed } moreover have "\A B X F. S\<^sub>1 = A@Inequality X F#B \ ineq_model \ X F" using assms(2) by (auto simp add: constr_sem_c_def) ultimately show ?thesis using not_simple_elim[OF \\simple S\<^sub>1\] by metis qed theorem LI_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" shows "\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" proof (cases "simple (LI_preproc S\<^sub>1)") case False let ?Stuck = "\S\<^sub>2 \\<^sub>2. \(\S\<^sub>3 \\<^sub>3. (S\<^sub>2,\\<^sub>2) \ (S\<^sub>3,\\<^sub>3) \ (\ \\<^sub>c \S\<^sub>3, \\<^sub>3\))" let ?Sats = "{((S,\),(S',\')). (S,\) \ (S',\') \ (\ \\<^sub>c \S, \\) \ (\ \\<^sub>c \S', \'\)}" have simple_if_stuck: "\S\<^sub>2 \\<^sub>2. \(LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2); \ \\<^sub>c \S\<^sub>2, \\<^sub>2\; ?Stuck S\<^sub>2 \\<^sub>2\ \ simple S\<^sub>2" using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_preserves_LI_preproc_prop[OF _ LI_preproc_preproc_prop] LI_completeness_single[OF LI_preserves_wellformedness] trancl_into_rtrancl by metis have base: "\b. ((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats" using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_completeness_single[OF _ _ False LI_preproc_preproc_prop] LI_preproc_sem_eq' \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by auto have *: "\S \ S' \'. ((S,\),(S',\')) \ ?Sats\<^sup>+ \ (S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" proof - fix S \ S' \' assume "((S,\),(S',\')) \ ?Sats\<^sup>+" thus "(S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" by (induct rule: trancl_induct2) auto qed have "\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2" proof (rule ccontr) assume "\(\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2)" hence sat_not_stuck: "\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ \?Stuck S\<^sub>2 \\<^sub>2" by blast have "\S \. ((LI_preproc S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+ \ (\b. ((S,\),b) \ ?Sats)" proof (intro allI impI) fix S \ assume a: "((LI_preproc S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+" have "\b. ((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+ \ \c. b \ c \ ((LI_preproc S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" proof - fix b assume in_sat: "((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+" hence "\c. (b,c) \ ?Sats" using * sat_not_stuck by (cases b) blast thus "\c. b \ c \ ((LI_preproc S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" using trancl_into_trancl[OF in_sat] by blast qed hence "\S' \'. (S,\) \ (S',\') \ ((LI_preproc S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" using a by auto then obtain S' \' where S'\': "(S,\) \ (S',\')" "((LI_preproc S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" by auto hence "\ \\<^sub>c \S', \'\" using * by blast moreover have "(LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>+ (S,\)" using a trancl_mono by blast ultimately have "((S,\),(S',\')) \ ?Sats" using S'\'(1) * a by blast thus "\b. ((S,\),b) \ ?Sats" using S'\'(2) by blast qed hence "\f. \i::nat. (f i, f (Suc i)) \ ?Sats" using infinite_chain_intro'[OF base] by blast moreover have "?Sats \ LI_rel\<^sup>+" by auto hence "\(\f. \i::nat. (f i, f (Suc i)) \ ?Sats)" using LI_no_infinite_chain infinite_chain_mono by blast ultimately show False by auto qed hence "\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>+ (S\<^sub>2, \\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using simple_if_stuck * by blast thus ?thesis by (meson trancl_into_rtrancl) qed (use LI_preproc_sem_eq' \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ in blast) end subsection \Corollary: Soundness and Completeness as a Single Theorem\ corollary LI_soundness_and_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\ \ (\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\))" by (metis LI_soundness[OF assms] LI_completeness[OF assms]) end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Miscellaneous.thy b/thys/Stateful_Protocol_Composition_and_Typing/Miscellaneous.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Miscellaneous.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Miscellaneous.thy @@ -1,747 +1,747 @@ (* Title: Miscellaneous.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Miscellaneous Lemmata\ theory Miscellaneous imports Main "HOL-Library.Sublist" "HOL-Library.Infinite_Set" "HOL-Library.While_Combinator" begin subsection \List: zip, filter, map\ lemma zip_arg_subterm_split: assumes "(x,y) \ set (zip xs ys)" obtains xs' xs'' ys' ys'' where "xs = xs'@x#xs''" "ys = ys'@y#ys''" "length xs' = length ys'" proof - from assms have "\zs zs' vs vs'. xs = zs@x#zs' \ ys = vs@y#vs' \ length zs = length vs" proof (induction ys arbitrary: xs) case (Cons y' ys' xs) then obtain x' xs' where x': "xs = x'#xs'" by (metis empty_iff list.exhaust list.set(1) set_zip_leftD) show ?case by (cases "(x, y) \ set (zip xs' ys')", metis \xs = x'#xs'\ Cons.IH[of xs'] Cons_eq_appendI list.size(4), use Cons.prems x' in fastforce) qed simp thus ?thesis using that by blast qed lemma zip_arg_index: assumes "(x,y) \ set (zip xs ys)" obtains i where "xs ! i = x" "ys ! i = y" "i < length xs" "i < length ys" proof - obtain xs1 xs2 ys1 ys2 where "xs = xs1@x#xs2" "ys = ys1@y#ys2" "length xs1 = length ys1" - using zip_arg_subterm_split[OF assms] by moura + using zip_arg_subterm_split[OF assms] by blast thus ?thesis using nth_append_length[of xs1 x xs2] nth_append_length[of ys1 y ys2] that by simp qed lemma in_set_zip_swap: "(x,y) \ set (zip xs ys) \ (y,x) \ set (zip ys xs)" unfolding in_set_zip by auto lemma filter_nth: "i < length (filter P xs) \ P (filter P xs ! i)" using nth_mem by force lemma list_all_filter_eq: "list_all P xs \ filter P xs = xs" by (metis list_all_iff filter_True) lemma list_all_filter_nil: assumes "list_all P xs" and "\x. P x \ \Q x" shows "filter Q xs = []" using assms by (induct xs) simp_all lemma list_all_concat: "list_all (list_all f) P \ list_all f (concat P)" by (induct P) auto lemma list_all2_in_set_ex: assumes P: "list_all2 P xs ys" and x: "x \ set xs" shows "\y \ set ys. P x y" proof - obtain i where i: "i < length xs" "xs ! i = x" by (meson x in_set_conv_nth) have "i < length ys" "P (xs ! i) (ys ! i)" using P i(1) by (simp_all add: list_all2_iff list_all2_nthD) thus ?thesis using i(2) by auto qed lemma list_all2_in_set_ex': assumes P: "list_all2 P xs ys" and y: "y \ set ys" shows "\x \ set xs. P x y" proof - obtain i where i: "i < length ys" "ys ! i = y" by (meson y in_set_conv_nth) have "i < length xs" "P (xs ! i) (ys ! i)" using P i(1) by (simp_all add: list_all2_iff list_all2_nthD) thus ?thesis using i(2) by auto qed lemma list_all2_sym: assumes "\x y. P x y \ P y x" and "list_all2 P xs ys" shows "list_all2 P ys xs" using assms(2) by (induct rule: list_all2_induct) (simp_all add: assms(1)) lemma map_upt_index_eq: assumes "j < length xs" shows "(map (\i. xs ! is i) [0..(i,p) \ insert x (set xs). \(i',p') \ insert x (set xs). p = p' \ i = i'" shows "map snd (List.insert x xs) = List.insert (snd x) (map snd xs)" using assms proof (induction xs rule: List.rev_induct) case (snoc y xs) hence IH: "map snd (List.insert x xs) = List.insert (snd x) (map snd xs)" by fastforce obtain iy py where y: "y = (iy,py)" by (metis surj_pair) obtain ix px where x: "x = (ix,px)" by (metis surj_pair) have "(ix,px) \ insert x (set (y#xs))" "(iy,py) \ insert x (set (y#xs))" using y x by auto hence *: "iy = ix" when "py = px" using that snoc.prems by auto show ?case proof (cases "px = py") case True hence "y = x" using * y x by auto thus ?thesis using IH by simp next case False hence "y \ x" using y x by simp have "List.insert x (xs@[y]) = (List.insert x xs)@[y]" proof - have 1: "insert y (set xs) = set (xs@[y])" by simp have 2: "x \ insert y (set xs) \ x \ set xs" using \y \ x\ by blast show ?thesis using 1 2 by (metis (no_types) List.insert_def append_Cons insertCI) qed thus ?thesis using IH y x False by (auto simp add: List.insert_def) qed qed simp lemma map_append_inv: "map f xs = ys@zs \ \vs ws. xs = vs@ws \ map f vs = ys \ map f ws = zs" proof (induction xs arbitrary: ys zs) case (Cons x xs') note prems = Cons.prems note IH = Cons.IH show ?case proof (cases ys) case (Cons y ys') then obtain vs' ws where *: "xs' = vs'@ws" "map f vs' = ys'" "map f ws = zs" using prems IH[of ys' zs] by auto hence "x#xs' = (x#vs')@ws" "map f (x#vs') = y#ys'" using Cons prems by force+ thus ?thesis by (metis Cons *(3)) qed (use prems in simp) qed simp lemma map2_those_Some_case: assumes "those (map2 f xs ys) = Some zs" and "(x,y) \ set (zip xs ys)" shows "\z. f x y = Some z" using assms proof (induction xs arbitrary: ys zs) case (Cons x' xs') obtain y' ys' where ys: "ys = y'#ys'" using Cons.prems(2) by (cases ys) simp_all obtain z where z: "f x' y' = Some z" using Cons.prems(1) ys by fastforce obtain zs' where zs: "those (map2 f xs' ys') = Some zs'" using z Cons.prems(1) ys by auto show ?case proof (cases "(x,y) = (x',y')") case False hence "(x,y) \ set (zip xs' ys')" using Cons.prems(2) unfolding ys by force thus ?thesis using Cons.IH[OF zs] by blast qed (use ys z in fast) qed simp lemma those_Some_Cons_ex: assumes "those (x#xs) = Some ys" shows "\y ys'. ys = y#ys' \ those xs = Some ys' \ x = Some y" using assms by (cases x) auto lemma those_Some_iff: "those xs = Some ys \ ((\x' \ set xs. \x. x' = Some x) \ ys = map the xs)" (is "?A xs ys \ ?B xs ys") proof show "?A xs ys \ ?B xs ys" proof (induction xs arbitrary: ys) case (Cons x' xs') obtain y' ys' where ys: "ys = y'#ys'" "those xs' = Some ys'" and x: "x' = Some y'" using Cons.prems those_Some_Cons_ex by blast show ?case using Cons.IH[OF ys(2)] unfolding x ys by simp qed simp show "?B xs ys \ ?A xs ys" by (induct xs arbitrary: ys) (simp, fastforce) qed lemma those_map2_SomeD: assumes "those (map2 f ts ss) = Some \" and "\ \ set \" shows "\(t,s) \ set (zip ts ss). f t s = Some \" using those_Some_iff[of "map2 f ts ss" \] assms by fastforce lemma those_map2_SomeI: assumes "\i. i < length xs \ f (xs ! i) (ys ! i) = Some (g i)" and "length xs = length ys" shows "those (map2 f xs ys) = Some (map g [0..z \ set (map2 f xs ys). \z'. z = Some z'" proof fix z assume z: "z \ set (map2 f xs ys)" then obtain i where i: "i < length xs" "i < length ys" "z = f (xs ! i) (ys ! i)" using in_set_conv_nth[of z "map2 f xs ys"] by auto thus "\z'. z = Some z'" using assms(1) by blast qed moreover have "map Some (map g [0..i. f (xs ! i) (ys ! i)) [0.. Some) (map g [0..List: subsequences\ lemma subseqs_set_subset: assumes "ys \ set (subseqs xs)" shows "set ys \ set xs" using assms subseqs_powset[of xs] by auto lemma subset_sublist_exists: "ys \ set xs \ \zs. set zs = ys \ zs \ set (subseqs xs)" proof (induction xs arbitrary: ys) case Cons thus ?case by (metis (no_types, lifting) Pow_iff imageE subseqs_powset) qed simp lemma map_subseqs: "map (map f) (subseqs xs) = subseqs (map f xs)" proof (induct xs) case (Cons x xs) have "map (Cons (f x)) (map (map f) (subseqs xs)) = map (map f) (map (Cons x) (subseqs xs))" by (induct "subseqs xs") auto thus ?case by (simp add: Let_def Cons) qed simp lemma subseqs_Cons: assumes "ys \ set (subseqs xs)" shows "ys \ set (subseqs (x#xs))" by (metis assms Un_iff set_append subseqs.simps(2)) lemma subseqs_subset: assumes "ys \ set (subseqs xs)" shows "set ys \ set xs" using assms by (metis Pow_iff image_eqI subseqs_powset) subsection \List: prefixes, suffixes\ lemma suffix_Cons': "suffix [x] (y#ys) \ suffix [x] ys \ (y = x \ ys = [])" using suffix_Cons[of "[x]"] by auto lemma prefix_Cons': "prefix (x#xs) (x#ys) \ prefix xs ys" by simp lemma prefix_map: "prefix xs (map f ys) \ \zs. prefix zs ys \ map f zs = xs" using map_append_inv unfolding prefix_def by fast lemma concat_mono_prefix: "prefix xs ys \ prefix (concat xs) (concat ys)" unfolding prefix_def by fastforce lemma concat_map_mono_prefix: "prefix xs ys \ prefix (concat (map f xs)) (concat (map f ys))" by (rule map_mono_prefix[THEN concat_mono_prefix]) lemma length_prefix_ex: assumes "n \ length xs" shows "\ys zs. xs = ys@zs \ length ys = n" using assms proof (induction n) case (Suc n) - then obtain ys zs where IH: "xs = ys@zs" "length ys = n" by moura + then obtain ys zs where IH: "xs = ys@zs" "length ys = n" by atomize_elim auto hence "length zs > 0" using Suc.prems(1) by auto then obtain v vs where v: "zs = v#vs" by (metis Suc_length_conv gr0_conv_Suc) hence "length (ys@[v]) = Suc n" using IH(2) by simp thus ?case using IH(1) v by (metis append.assoc append_Cons append_Nil) qed simp lemma length_prefix_ex': assumes "n < length xs" shows "\ys zs. xs = ys@xs ! n#zs \ length ys = n" proof - - obtain ys zs where xs: "xs = ys@zs" "length ys = n" using assms length_prefix_ex[of n xs] by moura + obtain ys zs where xs: "xs = ys@zs" "length ys = n" using assms length_prefix_ex[of n xs] by atomize_elim auto hence "length zs > 0" using assms by auto then obtain v vs where v: "zs = v#vs" by (metis Suc_length_conv gr0_conv_Suc) hence "(ys@zs) ! n = v" using xs by auto thus ?thesis using v xs by auto qed lemma length_prefix_ex2: assumes "i < length xs" "j < length xs" "i < j" shows "\ys zs vs. xs = ys@xs ! i#zs@xs ! j#vs \ length ys = i \ length zs = j - i - 1" proof - obtain xs0 vs where xs0: "xs = xs0@xs ! j#vs" "length xs0 = j" using length_prefix_ex'[OF assms(2)] by blast then obtain ys zs where "xs0 = ys@xs ! i#zs" "length ys = i" by (metis assms(3) length_prefix_ex' nth_append[of _ _ i]) thus ?thesis using xs0 by force qed lemma prefix_prefix_inv: assumes xs: "prefix xs (ys@zs)" and x: "suffix [x] xs" shows "prefix xs ys \ x \ set zs" proof - have "prefix xs ys" when "x \ set zs" using that xs proof (induction zs rule: rev_induct) case (snoc z zs) show ?case proof (rule snoc.IH) have "x \ z" using snoc.prems(1) by simp thus "prefix xs (ys@zs)" using snoc.prems(2) x by (metis append1_eq_conv append_assoc prefix_snoc suffixE) qed (use snoc.prems(1) in simp) qed simp thus ?thesis by blast qed lemma prefix_snoc_obtain: assumes xs: "prefix (xs@[x]) (ys@zs)" and ys: "\prefix (xs@[x]) ys" obtains vs where "xs@[x] = ys@vs@[x]" "prefix (vs@[x]) zs" proof - have "\vs. xs@[x] = ys@vs@[x] \ prefix (vs@[x]) zs" using xs proof (induction zs rule: List.rev_induct) case (snoc z zs) show ?case proof (cases "xs@[x] = ys@zs@[z]") case False hence "prefix (xs@[x]) (ys@zs)" using snoc.prems by (metis append_assoc prefix_snoc) thus ?thesis using snoc.IH by auto qed simp qed (simp add: ys) thus ?thesis using that by blast qed lemma prefix_snoc_in_iff: "x \ set xs \ (\B. prefix (B@[x]) xs)" by (induct xs rule: List.rev_induct) auto subsection \List: products\ lemma product_lists_Cons: "x#xs \ set (product_lists (y#ys)) \ (xs \ set (product_lists ys) \ x \ set y)" by auto lemma product_lists_in_set_nth: assumes "xs \ set (product_lists ys)" shows "\i set (ys ! i)" proof - have 0: "length ys = length xs" using assms(1) by (simp add: in_set_product_lists_length) thus ?thesis using assms proof (induction ys arbitrary: xs) case (Cons y ys) obtain x xs' where xs: "xs = x#xs'" using Cons.prems(1) by (metis length_Suc_conv) hence "xs' \ set (product_lists ys) \ \i set (ys ! i)" "length ys = length xs'" "x#xs' \ set (product_lists (y#ys))" using Cons by simp_all thus ?case using xs product_lists_Cons[of x xs' y ys] by (simp add: nth_Cons') qed simp qed lemma product_lists_in_set_nth': assumes "\i set (xs ! i)" and "length xs = length ys" shows "ys \ set (product_lists xs)" using assms proof (induction xs arbitrary: ys) case (Cons x xs) obtain y ys' where ys: "ys = y#ys'" using Cons.prems(2) by (metis length_Suc_conv) hence "ys' \ set (product_lists xs)" "y \ set x" "length xs = length ys'" using Cons by fastforce+ thus ?case using ys product_lists_Cons[of y ys' x xs] by (simp add: nth_Cons') qed simp subsection \Other Lemmata\ lemma finite_ballI: "\l \ {}. P l" "P x \ \l \ xs. P l \ \l \ insert x xs. P l" by (blast, blast) lemma list_set_ballI: "\l \ set []. P l" "P x \ \l \ set xs. P l \ \l \ set (x#xs). P l" by (simp, simp) lemma inv_set_fset: "finite M \ set (inv set M) = M" unfolding inv_def by (metis (mono_tags) finite_list someI_ex) lemma lfp_eqI': assumes "mono f" and "f C = C" and "\X \ Pow C. f X = X \ X = C" shows "lfp f = C" by (metis PowI assms lfp_lowerbound lfp_unfold subset_refl) lemma lfp_while': fixes f::"'a set \ 'a set" and M::"'a set" defines "N \ while (\A. f A \ A) f {}" assumes f_mono: "mono f" and N_finite: "finite N" and N_supset: "f N \ N" shows "lfp f = N" proof - have *: "f X \ N" when "X \ N" for X using N_supset monoD[OF f_mono that] by blast show ?thesis using lfp_while[OF f_mono * N_finite] by (simp add: N_def) qed lemma lfp_while'': fixes f::"'a set \ 'a set" and M::"'a set" defines "N \ while (\A. f A \ A) f {}" assumes f_mono: "mono f" and lfp_finite: "finite (lfp f)" shows "lfp f = N" proof - have *: "f X \ lfp f" when "X \ lfp f" for X using lfp_fixpoint[OF f_mono] monoD[OF f_mono that] by blast show ?thesis using lfp_while[OF f_mono * lfp_finite] by (simp add: N_def) qed lemma preordered_finite_set_has_maxima: assumes "finite A" "A \ {}" shows "\a::'a::{preorder} \ A. \b \ A. \(a < b)" using assms proof (induction A rule: finite_induct) case (insert a A) thus ?case by (cases "A = {}", simp, metis insert_iff order_trans less_le_not_le) qed simp lemma partition_index_bij: fixes n::nat obtains I k where "bij_betw I {0.. n" "\i. i < k \ P (I i)" "\i. k \ i \ i < n \ \(P (I i))" proof - define A where "A = filter P [0..i. \P i) [0..n. (A@B) ! n)" note defs = A_def B_def k_def I_def have k1: "k \ n" by (metis defs(1,3) diff_le_self dual_order.trans length_filter_le length_upt) have "i < k \ P (A ! i)" for i by (metis defs(1,3) filter_nth) hence k2: "i < k \ P ((A@B) ! i)" for i by (simp add: defs nth_append) have "i < length B \ \(P (B ! i))" for i by (metis defs(2) filter_nth) hence "i < length B \ \(P ((A@B) ! (k + i)))" for i using k_def by simp hence "k \ i \ i < k + length B \ \(P ((A@B) ! i))" for i by (metis add.commute add_less_imp_less_right le_add_diff_inverse2) hence k3: "k \ i \ i < n \ \(P ((A@B) ! i))" for i by (simp add: defs sum_length_filter_compl) have *: "length (A@B) = n" "set (A@B) = {0.. {0.. y \ {0.. (I x = I y) = (x = y)" by (metis *(1,3) defs(4) nth_eq_iff_index_eq atLeastLessThan_iff) next fix x show "x \ {0.. I x \ {0.. {0.. \x \ {0..x. x \ set xs \ finite {y. P x y}" shows "finite {ys. length xs = length ys \ (\y \ set ys. \x \ set xs. P x y)}" proof - define Q where "Q \ \ys. \y \ set ys. \x \ set xs. P x y" define M where "M \ {y. \x \ set xs. P x y}" have 0: "finite M" using assms unfolding M_def by fastforce have "Q ys \ set ys \ M" "(Q ys \ length ys = length xs) \ (length xs = length ys \ Q ys)" for ys unfolding Q_def M_def by auto thus ?thesis using finite_lists_length_eq[OF 0, of "length xs"] unfolding Q_def by presburger qed lemma trancl_eqI: assumes "\(a,b) \ A. \(c,d) \ A. b = c \ (a,d) \ A" shows "A = A\<^sup>+" proof show "A\<^sup>+ \ A" proof fix x assume x: "x \ A\<^sup>+" then obtain a b where ab: "x = (a,b)" by (metis surj_pair) hence "(a,b) \ A\<^sup>+" using x by metis hence "(a,b) \ A" using assms by (induct rule: trancl_induct) auto thus "x \ A" using ab by metis qed qed auto lemma trancl_eqI': assumes "\(a,b) \ A. \(c,d) \ A. b = c \ a \ d \ (a,d) \ A" and "\(a,b) \ A. a \ b" shows "A = {(a,b) \ A\<^sup>+. a \ b}" proof show "{(a,b) \ A\<^sup>+. a \ b} \ A" proof fix x assume x: "x \ {(a,b) \ A\<^sup>+. a \ b}" then obtain a b where ab: "x = (a,b)" by (metis surj_pair) hence "(a,b) \ A\<^sup>+" "a \ b" using x by blast+ hence "(a,b) \ A" proof (induction rule: trancl_induct) case base thus ?case by blast next case step thus ?case using assms(1) by force qed thus "x \ A" using ab by metis qed qed (use assms(2) in auto) lemma distinct_concat_idx_disjoint: assumes xs: "distinct (concat xs)" and ij: "i < length xs" "j < length xs" "i < j" shows "set (xs ! i) \ set (xs ! j) = {}" proof - obtain ys zs vs where ys: "xs = ys@xs ! i#zs@xs ! j#vs" "length ys = i" "length zs = j - i - 1" - using length_prefix_ex2[OF ij] by moura + using length_prefix_ex2[OF ij] by atomize_elim auto thus ?thesis using xs concat_append[of "ys@xs ! i#zs" "xs ! j#vs"] distinct_append[of "concat (ys@xs ! i#zs)" "concat (xs ! j#vs)"] by auto qed lemma remdups_ex2: "length (remdups xs) > 1 \ \a \ set xs. \b \ set xs. a \ b" by (metis distinct_Ex1 distinct_remdups less_trans nth_mem set_remdups zero_less_one zero_neq_one) lemma trancl_minus_refl_idem: defines "cl \ \ts. {(a,b) \ ts\<^sup>+. a \ b}" shows "cl (cl ts) = cl ts" proof - have 0: "(ts\<^sup>+)\<^sup>+ = ts\<^sup>+" "cl ts \ ts\<^sup>+" "(cl ts)\<^sup>+ \ (ts\<^sup>+)\<^sup>+" proof - show "(ts\<^sup>+)\<^sup>+ = ts\<^sup>+" "cl ts \ ts\<^sup>+" unfolding cl_def by auto thus "(cl ts)\<^sup>+ \ (ts\<^sup>+)\<^sup>+" using trancl_mono[of _ "cl ts" "ts\<^sup>+"] by blast qed have 1: "t \ cl (cl ts)" when t: "t \ cl ts" for t using t 0 unfolding cl_def by fast have 2: "t \ cl ts" when t: "t \ cl (cl ts)" for t proof - obtain a b where ab: "t = (a,b)" by (metis surj_pair) have "t \ (cl ts)\<^sup>+" and a_neq_b: "a \ b" using t unfolding cl_def ab by force+ hence "t \ ts\<^sup>+" using 0 by blast thus ?thesis using a_neq_b unfolding cl_def ab by blast qed show ?thesis using 1 2 by blast qed lemma ex_list_obtain: assumes ts: "\t. t \ set ts \ \s. P t s" obtains ss where "length ss = length ts" "\i < length ss. P (ts ! i) (ss ! i)" proof - have "\ss. length ss = length ts \ (\i < length ss. P (ts ! i) (ss ! i))" using ts proof (induction ts rule: List.rev_induct) case (snoc t ts) obtain s ss where s: "length ss = length ts" "\i < length ss. P (ts ! i) (ss ! i)" "P t s" using snoc.IH snoc.prems by force have *: "length (ss@[s]) = length (ts@[t])" using s(1) by simp hence "P ((ts@[t]) ! i) ((ss@[s]) ! i)" when i: "i < length (ss@[s])" for i using s(2,3) i nth_append[of ts "[t]"] nth_append[of ss "[s]"] by force thus ?case using * by blast qed simp thus thesis using that by blast qed lemma length_1_conv[iff]: "(length ts = 1) = (\a. ts = [a])" by (cases ts) simp_all lemma length_2_conv[iff]: "(length ts = 2) = (\a b. ts = [a,b])" proof (cases ts) case (Cons a ts') thus ?thesis by (cases ts') simp_all qed simp lemma length_3_conv[iff]: "(length ts = 3) \ (\a b c. ts = [a,b,c])" proof (cases ts) case (Cons a ts') note * = this thus ?thesis proof (cases ts') case (Cons b ts'') thus ?thesis using * by (cases ts'') simp_all qed simp qed simp lemma Max_nat_finite_le: assumes "finite M" and "\m. m \ M \ f m \ (n::nat)" shows "Max (insert 0 (f ` M)) \ n" proof - have 0: "finite (insert 0 (f ` M))" using assms(1) by blast have 1: "insert 0 (f ` M) \ {}" by force have 2: "m \ n" when "m \ insert 0 (f ` M)" for m using assms(2) that by fastforce show ?thesis using Max.boundedI[OF 0 1 2] by blast qed lemma Max_nat_finite_lt: assumes "finite M" and "M \ {}" and "\m. m \ M \ f m < (n::nat)" shows "Max (f ` M) < n" proof - define g where "g \ \m. Suc (f m)" have 0: "finite (f ` M)" "finite (g ` M)" using assms(1) by (blast, blast) have 1: "f ` M \ {}" "g ` M \ {}" using assms(2) by (force, force) have 2: "m \ n" when "m \ g ` M" for m using assms(3) that unfolding g_def by fastforce have 3: "Max (g ` M) \ n" using Max.boundedI[OF 0(2) 1(2) 2] by blast have 4: "Max (f ` M) < Max (g ` M)" using Max_in[OF 0(1) 1(1)] Max_gr_iff[OF 0(2) 1(2), of "Max (f ` M)"] unfolding g_def by blast show ?thesis using 3 4 by linarith qed lemma ex_finite_disj_nat_inj: fixes N N'::"nat set" assumes N: "finite N" and N': "finite N'" shows "\M::nat set. \\::nat \ nat. inj \ \ \ ` N = M \ M \ N' = {}" using N proof (induction N rule: finite_induct) case empty thus ?case using injI[of "\x::nat. x"] by blast next case (insert n N) then obtain M \ where M: "inj \" "\ ` N = M" "M \ N' = {}" by blast obtain m where m: "m \ M" "m \ insert n N" "m \ N'" using M(2) finite_imageI[OF insert.hyps(1), of \] insert.hyps(1) N' by (metis finite_UnI finite_insert UnCI infinite_nat_iff_unbounded_le finite_nat_set_iff_bounded_le) define \ where "\ \ \k. if k \ insert n N then (\(n := m)) k else Suc (Max (insert m M)) + k" have "insert m M \ N' = {}" using m M(3) unfolding \_def by auto moreover have "\ ` insert n N = insert m M" using insert.hyps(2) M(2) unfolding \_def by auto moreover have "inj \" proof (intro injI) fix i j assume ij: "\ i = \ j" have 0: "finite (insert m (\ ` N))" using insert.hyps(1) by simp have 1: "Suc (Max (insert m (\ ` N))) > k" when k: "k \ insert m (\ ` N)" for k using Max_ge[OF 0 k] by linarith have 2: "(\(n := m)) k \ insert m (\ ` N)" when k: "k \ insert n N" for k using k by auto have 3: "(\(n := m)) k \ Suc (Max (insert m (\ ` N))) + k'" when k: "k \ insert n N" for k k' using 1[OF 2[OF k]] by linarith have 4: "i \ insert n N \ j \ insert n N" using ij 3 M(2) unfolding \_def by metis show "i = j" proof (cases "i \ insert n N") case True hence *: "\ i = (\(n := m)) i" "\ j = (\(n := m)) j" "i \ insert n N" "j \ insert n N" using ij iffD1[OF 4] unfolding \_def by (argo, argo, argo, argo) show ?thesis proof (cases "i = n \ j = n") case True moreover have ?thesis when "i = n" "j = n" using that by simp moreover have False when "(i = n \ j \ n) \ (i \ n \ j = n)" by (metis M(2) that ij * m(1) fun_upd_other fun_upd_same image_eqI insertE) ultimately show ?thesis by argo next case False thus ?thesis using ij injD[OF M(1), of i j] unfolding *(1,2) by simp qed next case False thus ?thesis using ij 4 unfolding \_def by force qed qed ultimately show ?case by blast qed subsection \Infinite Paths in Relations as Mappings from Naturals to States\ context begin private fun rel_chain_fun::"nat \ 'a \ 'a \ ('a \ 'a) set \ 'a" where "rel_chain_fun 0 x _ _ = x" | "rel_chain_fun (Suc i) x y r = (if i = 0 then y else SOME z. (rel_chain_fun i x y r, z) \ r)" lemma infinite_chain_intro: fixes r::"('a \ 'a) set" assumes "\(a,b) \ r. \c. (b,c) \ r" "r \ {}" shows "\f. \i::nat. (f i, f (Suc i)) \ r" proof - from assms(2) obtain a b where "(a,b) \ r" by auto let ?P = "\i. (rel_chain_fun i a b r, rel_chain_fun (Suc i) a b r) \ r" let ?Q = "\i. \z. (rel_chain_fun i a b r, z) \ r" have base: "?P 0" using \(a,b) \ r\ by auto have step: "?P (Suc i)" when i: "?P i" for i proof - have "?Q (Suc i)" using assms(1) i by auto thus ?thesis using someI_ex[OF \?Q (Suc i)\] by auto qed have "\i::nat. (rel_chain_fun i a b r, rel_chain_fun (Suc i) a b r) \ r" using base step nat_induct[of ?P] by simp thus ?thesis by fastforce qed end lemma infinite_chain_intro': fixes r::"('a \ 'a) set" assumes base: "\b. (x,b) \ r" and step: "\b. (x,b) \ r\<^sup>+ \ (\c. (b,c) \ r)" shows "\f. \i::nat. (f i, f (Suc i)) \ r" proof - let ?s = "{(a,b) \ r. a = x \ (x,a) \ r\<^sup>+}" have "?s \ {}" using base by auto have "\c. (b,c) \ ?s" when ab: "(a,b) \ ?s" for a b proof (cases "a = x") case False hence "(x,a) \ r\<^sup>+" using ab by auto hence "(x,b) \ r\<^sup>+" using \(a,b) \ ?s\ by auto thus ?thesis using step by auto qed (use ab step in auto) hence "\f. \i. (f i, f (Suc i)) \ ?s" using infinite_chain_intro[of ?s] \?s \ {}\ by blast thus ?thesis by auto qed lemma infinite_chain_mono: assumes "S \ T" "\f. \i::nat. (f i, f (Suc i)) \ S" shows "\f. \i::nat. (f i, f (Suc i)) \ T" using assms by auto end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy @@ -1,3460 +1,3460 @@ (* Based on src/HOL/ex/Unification.thy packaged with Isabelle/HOL 2015 having the following license: ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER. Copyright (c) 1986-2015, University of Cambridge, Technische Universitaet Muenchen, and contributors. All rights reserved. Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: * Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. * Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. * Neither the name of the University of Cambridge or the Technische Universitaet Muenchen nor the names of their contributors may be used to endorse or promote products derived from this software without specific prior written permission. THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (* Title: More_Unification.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause Originally based on src/HOL/ex/Unification.thy (Isabelle/HOL 2015) by: Author: Martin Coen, Cambridge University Computer Laboratory Author: Konrad Slind, TUM & Cambridge University Computer Laboratory Author: Alexander Krauss, TUM *) section \Definitions and Properties Related to Substitutions and Unification\ theory More_Unification imports Messages "First_Order_Terms.Unification" begin subsection \Substitutions\ abbreviation subst_apply_list (infix "\\<^sub>l\<^sub>i\<^sub>s\<^sub>t" 51) where "T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ map (\t. t \ \) T" abbreviation subst_apply_pair (infixl "\\<^sub>p" 60) where "d \\<^sub>p \ \ (case d of (t,t') \ (t \ \, t' \ \))" abbreviation subst_apply_pair_set (infixl "\\<^sub>p\<^sub>s\<^sub>e\<^sub>t" 60) where "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ (\d. d \\<^sub>p \) ` M" definition subst_apply_pairs (infix "\\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s" 51) where "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \ map (\f. f \\<^sub>p \) F" abbreviation subst_more_general_than (infixl "\\<^sub>\" 50) where "\ \\<^sub>\ \ \ \\. \ = \ \\<^sub>s \" abbreviation subst_support (infix "supports" 50) where "\ supports \ \ (\x. \ x \ \ = \ x)" abbreviation rm_var where "rm_var v s \ s(v := Var v)" abbreviation rm_vars where "rm_vars vs \ \ (\v. if v \ vs then Var v else \ v)" definition subst_elim where "subst_elim \ v \ \t. v \ fv (t \ \)" definition subst_idem where "subst_idem s \ s \\<^sub>s s = s" lemma subst_support_def: "\ supports \ \ \ = \ \\<^sub>s \" unfolding subst_compose_def by metis lemma subst_supportD: "\ supports \ \ \ \\<^sub>\ \" using subst_support_def by auto lemma rm_vars_empty[simp]: "rm_vars {} s = s" "rm_vars (set []) s = s" by simp_all lemma rm_vars_singleton: "rm_vars {v} s = rm_var v s" by auto lemma subst_apply_terms_empty: "M \\<^sub>s\<^sub>e\<^sub>t Var = M" by simp lemma subst_agreement: "(t \ r = t \ s) \ (\v \ fv t. Var v \ r = Var v \ s)" by (induct t) auto lemma repl_invariance[dest?]: "v \ fv t \ t \ s(v := u) = t \ s" by (simp add: subst_agreement) lemma subst_idx_map: assumes "\i \ set I. i < length T" shows "(map ((!) T) I) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) I" using assms by auto lemma subst_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t ((!) (map (\t. t \ \) T))" (is "?A = ?B") proof - have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for i using that assms by auto hence "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for k using that by (induction k) force+ thus ?thesis by auto qed lemma subst_remove_var: "v \ fv s \ v \ fv (t \ Var(v := s))" by (induct t) simp_all lemma subst_set_map: "x \ set X \ x \ s \ set (map (\x. x \ s) X)" by simp lemma subst_set_idx_map: assumes "\i \ I. i < length T" shows "(!) T ` I \\<^sub>s\<^sub>e\<^sub>t \ = (!) (map (\t. t \ \) T) ` I" (is "?A = ?B") proof have *: "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto show "?A \ ?B" using * assms by blast show "?B \ ?A" using * assms by auto qed lemma subst_set_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t K. i < length T" shows "K \\<^sub>s\<^sub>e\<^sub>t (!) T \\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (!) (map (\t. t \ \) T)" (is "?A = ?B") proof have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t K" for i using that assms by auto hence *: "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t K" for k using that by (induction k) force+ show "?A \ ?B" using * by auto show "?B \ ?A" using * by force qed lemma subst_term_list_obtain: assumes "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" and "length T = length S" shows "\U. length T = length U \ (\i < length T. P (T ! i) (U ! i)) \ S = map (\u. u \ \) U" using assms proof (induction T arbitrary: S) case (Cons t T S') then obtain s S where S': "S' = s#S" by (cases S') auto have "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" "length T = length S" using Cons.prems S' by force+ then obtain U where U: "length T = length U" "\i < length T. P (T ! i) (U ! i)" "S = map (\u. u \ \) U" - using Cons.IH by moura + using Cons.IH by atomize_elim auto obtain u where u: "P t u" "s = u \ \" using Cons.prems(1) S' by auto have 1: "length (t#T) = length (u#U)" using Cons.prems(2) U(1) by fastforce have 2: "\i < length (t#T). P ((t#T) ! i) ((u#U) ! i)" using u(1) U(2) by (simp add: nth_Cons') have 3: "S' = map (\u. u \ \) (u#U)" using U u S' by simp show ?case using 1 2 3 by blast qed simp lemma subst_mono: "t \ u \ t \ s \ u \ s" by (induct u) auto lemma subst_mono_fv: "x \ fv t \ s x \ t \ s" by (induct t) auto lemma subst_mono_neq: assumes "t \ u" shows "t \ s \ u \ s" proof (cases u) case (Var v) hence False using \t \ u\ by simp thus ?thesis .. next case (Fun f X) then obtain x where "x \ set X" "t \ x" using \t \ u\ by auto hence "t \ s \ x \ s" using subst_mono by metis obtain Y where "Fun f X \ s = Fun f Y" by auto hence "x \ s \ set Y" using \x \ set X\ by auto hence "x \ s \ Fun f X \ s" using \Fun f X \ s = Fun f Y\ Fun_param_is_subterm by simp hence "t \ s \ Fun f X \ s" using \t \ s \ x \ s\ by (metis term.dual_order.trans term.order.eq_iff) thus ?thesis using \u = Fun f X\ \t \ u\ by metis qed lemma subst_no_occs[dest]: "\Var v \ t \ t \ Var(v := s) = t" by (induct t) (simp_all add: map_idI) lemma var_comp[simp]: "\ \\<^sub>s Var = \" "Var \\<^sub>s \ = \" unfolding subst_compose_def by simp_all lemma subst_comp_all: "M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = (M \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" unfolding subst_subst_compose by auto lemma subst_all_mono: "M \ M' \ M \\<^sub>s\<^sub>e\<^sub>t s \ M' \\<^sub>s\<^sub>e\<^sub>t s" by auto lemma subst_comp_set_image: "(\ \\<^sub>s \) ` X = \ ` X \\<^sub>s\<^sub>e\<^sub>t \" using subst_compose by fastforce lemma subst_ground_ident[dest?]: "fv t = {} \ t \ s = t" by (induct t, simp, metis subst_agreement empty_iff subst_apply_term_empty) lemma subst_ground_ident_compose: "fv (\ x) = {} \ (\ \\<^sub>s \) x = \ x" "fv (t \ \) = {} \ t \ (\ \\<^sub>s \) = t \ \" unfolding subst_subst_compose by (simp_all add: subst_compose_def subst_ground_ident) lemma subst_all_ground_ident[dest?]: "ground M \ M \\<^sub>s\<^sub>e\<^sub>t s = M" proof - assume "ground M" hence "\t. t \ M \ fv t = {}" by auto hence "\t. t \ M \ t \ s = t" by (metis subst_ground_ident) moreover have "\t. t \ M \ t \ s \ M \\<^sub>s\<^sub>e\<^sub>t s" by (metis imageI) ultimately show "M \\<^sub>s\<^sub>e\<^sub>t s = M" by (simp add: image_cong) qed lemma subst_cong: "\\ = \'; \ = \'\ \ (\ \\<^sub>s \) = (\' \\<^sub>s \')" by auto lemma subst_mgt_bot[simp]: "Var \\<^sub>\ \" by simp lemma subst_mgt_refl[simp]: "\ \\<^sub>\ \" by (metis var_comp(1)) lemma subst_mgt_trans: "\\ \\<^sub>\ \; \ \\<^sub>\ \\ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma subst_mgt_comp: "\ \\<^sub>\ \ \\<^sub>s \" by auto lemma subst_mgt_comp': "\ \\<^sub>s \ \\<^sub>\ \ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma var_self: "(\w. if w = v then Var v else Var w) = Var" using subst_agreement by auto lemma var_same[simp]: "Var(v := t) = Var \ t = Var v" by (intro iffI, metis fun_upd_same, simp add: var_self) lemma subst_eq_if_eq_vars: "(\v. (Var v) \ \ = (Var v) \ \) \ \ = \" by (auto simp add: subst_agreement) lemma subst_all_empty[simp]: "{} \\<^sub>s\<^sub>e\<^sub>t \ = {}" by simp lemma subst_all_insert:"(insert t M) \\<^sub>s\<^sub>e\<^sub>t \ = insert (t \ \) (M \\<^sub>s\<^sub>e\<^sub>t \)" by auto lemma subst_apply_fv_subset: "fv t \ V \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by (induct t) auto lemma subst_apply_fv_empty: assumes "fv t = {}" shows "fv (t \ \) = {}" using assms subst_apply_fv_subset[of t "{}" \] by auto lemma subst_compose_fv: assumes "fv (\ x) = {}" shows "fv ((\ \\<^sub>s \) x) = {}" using assms subst_apply_fv_empty unfolding subst_compose_def by fast lemma subst_compose_fv': fixes \ \::"('a,'b) subst" assumes "y \ fv ((\ \\<^sub>s \) x)" shows "\z. z \ fv (\ x)" using assms subst_compose_fv by fast lemma subst_apply_fv_unfold: "fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by (induct t) auto lemma subst_apply_fv_unfold_set: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (simp add: subst_apply_fv_unfold) lemma subst_apply_fv_unfold': "fv (t \ \) = (\v \ fv t. fv (\ v))" using subst_apply_fv_unfold by simp lemma subst_apply_fv_union: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))" proof - have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by auto thus ?thesis using subst_apply_fv_unfold by metis qed lemma fv\<^sub>s\<^sub>e\<^sub>t_subst: "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t M)" by (simp add: subst_apply_fv_unfold) lemma subst_list_set_fv: "fv\<^sub>s\<^sub>e\<^sub>t (set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using subst_apply_fv_unfold_set[of \ ts] by simp lemma subst_elimI[intro]: "(\t. v \ fv (t \ \)) \ subst_elim \ v" by (auto simp add: subst_elim_def) lemma subst_elimI'[intro]: "(\w. v \ fv (Var w \ \)) \ subst_elim \ v" by (simp add: subst_elim_def subst_apply_fv_unfold') lemma subst_elimD[dest]: "subst_elim \ v \ v \ fv (t \ \)" by (auto simp add: subst_elim_def) lemma subst_elimD'[dest]: "subst_elim \ v \ \ v \ Var v" by (metis subst_elim_def eval_term.simps(1) term.set_intros(3)) lemma subst_elimD''[dest]: "subst_elim \ v \ v \ fv (\ w)" by (metis subst_elim_def eval_term.simps(1)) lemma subst_elim_rm_vars_dest[dest]: "subst_elim (\::('a,'b) subst) v \ v \ vs \ subst_elim (rm_vars vs \) v" proof - assume assms: "subst_elim \ v" "v \ vs" obtain f::"('a, 'b) subst \ 'b \ 'b" where "\\ v. (\w. v \ fv (Var w \ \)) = (v \ fv (Var (f \ v) \ \))" by moura hence *: "\a \. a \ fv (Var (f \ a) \ \) \ subst_elim \ a" by blast have "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \ \ v \ fv (Var (f (rm_vars vs \) v) \ rm_vars vs \)" using assms(1) by fastforce moreover { assume "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \" hence "rm_vars vs \ (f (rm_vars vs \) v) \ \ (f (rm_vars vs \) v)" by auto hence "f (rm_vars vs \) v \ vs" by meson hence ?thesis using * assms(2) by force } ultimately show ?thesis using * by blast qed lemma occs_subst_elim: "\Var v \ t \ subst_elim (Var(v := t)) v \ (Var(v := t)) = Var" proof (cases "Var v = t") assume "Var v \ t" "\Var v \ t" hence "v \ fv t" by (simp add: vars_iff_subterm_or_eq) thus ?thesis by (auto simp add: subst_remove_var) qed auto lemma occs_subst_elim': "\Var v \ t \ subst_elim (Var(v := t)) v" proof - assume "\Var v \ t" hence "v \ fv t" by (auto simp add: vars_iff_subterm_or_eq) thus "subst_elim (Var(v := t)) v" by (simp add: subst_elim_def subst_remove_var) qed lemma subst_elim_comp: "subst_elim \ v \ subst_elim (\ \\<^sub>s \) v" by (auto simp add: subst_elim_def) lemma var_subst_idem: "subst_idem Var" by (simp add: subst_idem_def) lemma var_upd_subst_idem: assumes "\Var v \ t" shows "subst_idem (Var(v := t))" using subst_no_occs[OF assms] by (simp add: subst_idem_def subst_def[symmetric]) lemma zip_map_subst: "zip xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. (t, t \ \)) xs" by (induction xs) auto lemma map2_map_subst: "map2 f xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. f t (t \ \)) xs" by (induction xs) auto subsection \Lemmata: Domain and Range of Substitutions\ lemma range_vars_alt_def: "range_vars s \ fv\<^sub>s\<^sub>e\<^sub>t (subst_range s)" unfolding range_vars_def by simp lemma subst_dom_var_finite[simp]: "finite (subst_domain Var)" by simp lemma subst_range_Var[simp]: "subst_range Var = {}" by simp lemma range_vars_Var[simp]: "range_vars Var = {}" by fastforce lemma finite_subst_img_if_finite_dom: "finite (subst_domain \) \ finite (range_vars \)" unfolding range_vars_alt_def by auto lemma finite_subst_img_if_finite_dom': "finite (subst_domain \) \ finite (subst_range \)" by auto lemma subst_img_alt_def: "subst_range s = {t. \v. s v = t \ t \ Var v}" by (auto simp add: subst_domain_def) lemma subst_fv_img_alt_def: "range_vars s = (\t \ {t. \v. s v = t \ t \ Var v}. fv t)" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_domI[intro]: "\ v \ Var v \ v \ subst_domain \" by (simp add: subst_domain_def) lemma subst_imgI[intro]: "\ v \ Var v \ \ v \ subst_range \" by (simp add: subst_domain_def) lemma subst_fv_imgI[intro]: "\ v \ Var v \ fv (\ v) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_eqI': assumes "t \ \ = t \ \" "subst_domain \ = subst_domain \" "subst_domain \ \ fv t" shows "\ = \" by (metis assms(2,3) term_subst_eq_rev[OF assms(1)] in_mono ext subst_domI) lemma subst_domain_subst_Fun_single[simp]: "subst_domain (Var(x := Fun f T)) = {x}" (is "?A = ?B") unfolding subst_domain_def by simp lemma subst_range_subst_Fun_single[simp]: "subst_range (Var(x := Fun f T)) = {Fun f T}" (is "?A = ?B") by simp lemma range_vars_subst_Fun_single[simp]: "range_vars (Var(x := Fun f T)) = fv (Fun f T)" unfolding range_vars_alt_def by force lemma var_renaming_is_Fun_iff: assumes "subst_range \ \ range Var" shows "is_Fun t = is_Fun (t \ \)" proof (cases t) case (Var x) hence "\y. \ x = Var y" using assms by auto thus ?thesis using Var by auto qed simp lemma subst_fv_dom_img_subset: "fv t \ subst_domain \ \ fv (t \ \) \ range_vars \" unfolding range_vars_alt_def by (induct t) auto lemma subst_fv_dom_img_subset_set: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ range_vars \" proof - assume assms: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \" obtain f::"'a set \ (('b, 'a) term \ 'a set) \ ('b, 'a) terms \ ('b, 'a) term" where "\x y z. (\v. v \ z \ \ y v \ x) \ (f x y z \ z \ \ y (f x y z) \ x)" by moura hence *: "\T g A. (\ \ (g ` T) \ A \ (\t. t \ T \ g t \ A)) \ (\ (g ` T) \ A \ f A g T \ T \ \ g (f A g T) \ A)" by (metis (no_types) SUP_le_iff) hence **: "\t. t \ M \ fv t \ subst_domain \" by (metis (no_types) assms fv\<^sub>s\<^sub>e\<^sub>t.simps) have "\t::('b, 'a) term. \f T. t \ f ` T \ (\t'::('b, 'a) term. t = f t' \ t' \ T)" by blast hence "f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ fv (f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \)) \ range_vars \" by (metis (full_types) ** subst_fv_dom_img_subset) thus ?thesis by (metis (no_types) * fv\<^sub>s\<^sub>e\<^sub>t.simps) qed lemma subst_fv_dom_ground_if_ground_img: assumes "fv t \ subst_domain s" "ground (subst_range s)" shows "fv (t \ s) = {}" using subst_fv_dom_img_subset[OF assms(1)] assms(2) by force lemma subst_fv_dom_ground_if_ground_img': assumes "fv t \ subst_domain s" "\x. x \ subst_domain s \ fv (s x) = {}" shows "fv (t \ s) = {}" using subst_fv_dom_ground_if_ground_img[OF assms(1)] assms(2) by auto lemma subst_fv_unfold: "fv (t \ s) = (fv t - subst_domain s) \ fv\<^sub>s\<^sub>e\<^sub>t (s ` (fv t \ subst_domain s))" proof (induction t) case (Var v) thus ?case proof (cases "v \ subst_domain s") case True thus ?thesis by auto next case False hence "fv (Var v \ s) = {v}" "fv (Var v) \ subst_domain s = {}" by auto thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma subst_fv_unfold_ground_img: "range_vars s = {} \ fv (t \ s) = fv t - subst_domain s" by (auto simp: range_vars_alt_def subst_fv_unfold) lemma subst_img_update: "\\ v = Var v; t \ Var v\ \ range_vars (\(v := t)) = range_vars \ \ fv t" proof - assume "\ v = Var v" "t \ Var v" hence "(\s \ {s. \w. (\(v := t)) w = s \ s \ Var w}. fv s) = fv t \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "range_vars (\(v := t)) = range_vars \ \ fv t" by (metis Un_commute subst_fv_img_alt_def) qed lemma subst_dom_update1: "v \ subst_domain \ \ subst_domain (\(v := Var v)) = subst_domain \" by (auto simp add: subst_domain_def) lemma subst_dom_update2: "t \ Var v \ subst_domain (\(v := t)) = insert v (subst_domain \)" by (auto simp add: subst_domain_def) lemma subst_dom_update3: "t = Var v \ subst_domain (\(v := t)) = subst_domain \ - {v}" by (auto simp add: subst_domain_def) lemma var_not_in_subst_dom[elim]: "v \ subst_domain s \ s v = Var v" by (simp add: subst_domain_def) lemma subst_dom_vars_in_subst[elim]: "v \ subst_domain s \ s v \ Var v" by (simp add: subst_domain_def) lemma subst_not_dom_fixed: "\v \ fv t; v \ subst_domain s\ \ v \ fv (t \ s)" by (induct t) auto lemma subst_not_img_fixed: "\v \ fv (t \ s); v \ range_vars s\ \ v \ fv t" unfolding range_vars_alt_def by (induct t) force+ lemma ground_range_vars[intro]: "ground (subst_range s) \ range_vars s = {}" unfolding range_vars_alt_def by metis lemma ground_subst_no_var[intro]: "ground (subst_range s) \ x \ range_vars s" using ground_range_vars[of s] by blast lemma ground_img_obtain_fun: assumes "ground (subst_range s)" "x \ subst_domain s" obtains f T where "s x = Fun f T" "Fun f T \ subst_range s" "fv (Fun f T) = {}" proof - - from assms(2) obtain t where t: "s x = t" "t \ subst_range s" by moura + from assms(2) obtain t where t: "s x = t" "t \ subst_range s" by atomize_elim auto hence "fv t = {}" using assms(1) by auto thus ?thesis using t that by (cases t) simp_all qed lemma ground_term_subst_domain_fv_subset: "fv (t \ \) = {} \ fv t \ subst_domain \" by (induct t) auto lemma ground_subst_range_empty_fv: "ground (subst_range \) \ x \ subst_domain \ \ fv (\ x) = {}" by simp lemma subst_Var_notin_img: "x \ range_vars s \ t \ s = Var x \ t = Var x" using subst_not_img_fixed[of x t s] by (induct t) auto lemma fv_in_subst_img: "\s v = t; t \ Var v\ \ fv t \ range_vars s" unfolding range_vars_alt_def by auto lemma empty_dom_iff_empty_subst: "subst_domain \ = {} \ \ = Var" by auto lemma subst_dom_cong: "(\v t. \ v = t \ \ v = t) \ subst_domain \ \ subst_domain \" by (auto simp add: subst_domain_def) lemma subst_img_cong: "(\v t. \ v = t \ \ v = t) \ range_vars \ \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_dom_elim: "subst_domain s \ range_vars s = {} \ fv (t \ s) \ subst_domain s = {}" proof (induction t) case (Var v) thus ?case using fv_in_subst_img[of s] by (cases "s v = Var v") (auto simp add: subst_domain_def) next case Fun thus ?case by auto qed lemma subst_dom_insert_finite: "finite (subst_domain s) = finite (subst_domain (s(v := t)))" proof assume "finite (subst_domain s)" have "subst_domain (s(v := t)) \ insert v (subst_domain s)" by (auto simp add: subst_domain_def) thus "finite (subst_domain (s(v := t)))" by (meson \finite (subst_domain s)\ finite_insert rev_finite_subset) next assume *: "finite (subst_domain (s(v := t)))" hence "finite (insert v (subst_domain s))" proof (cases "t = Var v") case True hence "finite (subst_domain s - {v})" by (metis * subst_dom_update3) thus ?thesis by simp qed (metis * subst_dom_update2[of t v s]) thus "finite (subst_domain s)" by simp qed lemma trm_subst_disj: "t \ \ = t \ fv t \ subst_domain \ = {}" proof (induction t) case (Fun f X) hence "map (\x. x \ \) X = X" by simp hence "\x. x \ set X \ x \ \ = x" using map_eq_conv by fastforce thus ?case using Fun.IH by auto qed (simp add: subst_domain_def) declare subst_apply_term_ident[intro] lemma trm_subst_ident'[intro]: "v \ subst_domain \ \ (Var v) \ \ = Var v" using subst_apply_term_ident by (simp add: subst_domain_def) lemma trm_subst_ident''[intro]: "(\x. x \ fv t \ \ x = Var x) \ t \ \ = t" proof - assume "\x. x \ fv t \ \ x = Var x" hence "fv t \ subst_domain \ = {}" by (auto simp add: subst_domain_def) thus ?thesis using subst_apply_term_ident by auto qed lemma set_subst_ident: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t \ = M" proof - assume "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {}" hence "\t \ M. t \ \ = t" by auto thus ?thesis by force qed lemma trm_subst_ident_subterms[intro]: "fv t \ subst_domain \ = {} \ subterms t \\<^sub>s\<^sub>e\<^sub>t \ = subterms t" using set_subst_ident[of "subterms t" \] fv_subterms[of t] by simp lemma trm_subst_ident_subterms'[intro]: "v \ fv t \ subterms t \\<^sub>s\<^sub>e\<^sub>t Var(v := s) = subterms t" using trm_subst_ident_subterms[of t "Var(v := s)"] by (meson subst_no_occs trm_subst_disj vars_iff_subtermeq) lemma const_mem_subst_cases: assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ \ ` fv\<^sub>s\<^sub>e\<^sub>t M" proof - obtain m where m: "m \ M" "m \ \ = Fun c []" using assms by auto thus ?thesis by (cases m) force+ qed lemma const_mem_subst_cases': assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ subst_range \" using const_mem_subst_cases[OF assms] by force lemma fv_subterms_substI[intro]: "y \ fv t \ \ y \ subterms t \\<^sub>s\<^sub>e\<^sub>t \" using image_iff vars_iff_subtermeq by fastforce lemma fv_subterms_subst_eq[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms (t \ \)) = fv\<^sub>s\<^sub>e\<^sub>t (subterms t \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms by (induct t) force+ lemma fv_subterms_set_subst: "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \))" using fv_subterms_subst_eq[of _ \] by auto lemma fv_subterms_set_subst': "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms_set[of "M \\<^sub>s\<^sub>e\<^sub>t \"] fv_subterms_set_subst[of \ M] by simp lemma fv_subst_subset: "x \ fv t \ fv (\ x) \ fv (t \ \)" by (metis fv_subset image_eqI subst_apply_fv_unfold) lemma fv_subst_subset': "fv s \ fv t \ fv (s \ \) \ fv (t \ \)" using fv_subst_subset by (induct s) force+ lemma fv_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv (t \ \)" shows "\y \ fv t. x \ fv (\ y)" using assms by (induct t) force+ lemma set_subst_all_ident: "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = M \\<^sub>s\<^sub>e\<^sub>t \" by (metis set_subst_ident subst_comp_all) lemma subterms_subst: "subterms (t \ d) = (subterms t \\<^sub>s\<^sub>e\<^sub>t d) \ subterms\<^sub>s\<^sub>e\<^sub>t (d ` (fv t \ subst_domain d))" by (induct t) (auto simp add: subst_domain_def) lemma subterms_subst': fixes \::"('a,'b) subst" assumes "\x \ fv t. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ subst_domain \") case True hence "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using Var by simp hence "subterms (\ x) = {\ x}" by auto thus ?thesis by simp qed auto qed auto lemma subterms_subst'': fixes \::"('a,'b) subst" assumes "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subterms_subst'[of _ \] assms by auto lemma subterms_subst_subterm: fixes \::"('a,'b) subst" assumes "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" and "b \ subterms (a \ \)" shows "\c \ subterms a. c \ \ = b" using subterms_subst'[OF assms(1)] assms(2) by auto lemma subterms_subst_subset: "subterms t \\<^sub>s\<^sub>e\<^sub>t \ \ subterms (t \ \)" by (induct t) auto lemma subterms_subst_subset': "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using subterms_subst_subset by fast lemma subterms\<^sub>s\<^sub>e\<^sub>t_subst: fixes \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. t \ subterms (\ x))" using assms subterms_subst[of _ \] by auto lemma rm_vars_dom: "subst_domain (rm_vars V s) = subst_domain s - V" by (auto simp add: subst_domain_def) lemma rm_vars_dom_subset: "subst_domain (rm_vars V s) \ subst_domain s" by (auto simp add: subst_domain_def) lemma rm_vars_dom_eq': "subst_domain (rm_vars (UNIV - V) s) = subst_domain s \ V" using rm_vars_dom[of "UNIV - V" s] by blast lemma rm_vars_dom_eqI: assumes "t \ \ = t \ \" shows "subst_domain (rm_vars (UNIV - fv t) \) = subst_domain (rm_vars (UNIV - fv t) \)" by (meson assms Diff_iff UNIV_I term_subst_eq_rev) lemma rm_vars_img: "subst_range (rm_vars V s) = s ` subst_domain (rm_vars V s)" by (auto simp add: subst_domain_def) lemma rm_vars_img_subset: "subst_range (rm_vars V s) \ subst_range s" by (auto simp add: subst_domain_def) lemma rm_vars_img_fv_subset: "range_vars (rm_vars V s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma rm_vars_fv_obtain: assumes "x \ fv (t \ rm_vars X \) - X" shows "\y \ fv t - X. x \ fv (rm_vars X \ y)" using assms by (induct t) (fastforce, force) lemma rm_vars_apply: "v \ subst_domain (rm_vars V s) \ (rm_vars V s) v = s v" by (auto simp add: subst_domain_def) lemma rm_vars_apply': "subst_domain \ \ vs = {} \ rm_vars vs \ = \" by force lemma rm_vars_ident: "fv t \ vs = {} \ t \ (rm_vars vs \) = t \ \" by (induct t) auto lemma rm_vars_fv_subset: "fv (t \ rm_vars X \) \ fv t \ fv (t \ \)" by (induct t) auto lemma rm_vars_fv_disj: assumes "fv t \ X = {}" "fv (t \ \) \ X = {}" shows "fv (t \ rm_vars X \) \ X = {}" using rm_vars_ident[OF assms(1)] assms(2) by auto lemma rm_vars_ground_supports: assumes "ground (subst_range \)" shows "rm_vars X \ supports \" proof fix x have *: "ground (subst_range (rm_vars X \))" using rm_vars_img_subset[of X \] assms by (auto simp add: subst_domain_def) show "rm_vars X \ x \ \ = \ x " proof (cases "x \ subst_domain (rm_vars X \)") case True hence "fv (rm_vars X \ x) = {}" using * by auto thus ?thesis using True by auto qed (simp add: subst_domain_def) qed lemma rm_vars_split: assumes "ground (subst_range \)" shows "\ = rm_vars X \ \\<^sub>s rm_vars (subst_domain \ - X) \" proof - let ?s1 = "rm_vars X \" let ?s2 = "rm_vars (subst_domain \ - X) \" have doms: "subst_domain ?s1 \ subst_domain \" "subst_domain ?s2 \ subst_domain \" by (auto simp add: subst_domain_def) { fix x assume "x \ subst_domain \" hence "\ x = Var x" "?s1 x = Var x" "?s2 x = Var x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = Var x" "?s2 x = \ x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = \ x" "fv (\ x) = {}" using assms doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose subst_ground_ident) } ultimately show ?thesis by blast qed lemma rm_vars_fv_img_disj: assumes "fv t \ X = {}" "X \ range_vars \ = {}" shows "fv (t \ rm_vars X \) \ X = {}" using assms proof (induction t) case (Var x) hence *: "(rm_vars X \) x = \ x" by auto show ?case proof (cases "x \ subst_domain \") case True hence "\ x \ subst_range \" by auto hence "fv (\ x) \ X = {}" using Var.prems(2) unfolding range_vars_alt_def by fastforce thus ?thesis using * by auto next case False thus ?thesis using Var.prems(1) by auto qed next case Fun thus ?case by auto qed lemma subst_apply_dom_ident: "t \ \ = t \ subst_domain \ \ subst_domain \ \ t \ \ = t" proof (induction t) case (Fun f T) thus ?case by (induct T) auto qed (auto simp add: subst_domain_def) lemma rm_vars_subst_apply_ident: assumes "t \ \ = t" shows "t \ (rm_vars vs \) = t" using rm_vars_dom[of vs \] subst_apply_dom_ident[OF assms, of "rm_vars vs \"] by auto lemma rm_vars_subst_eq: "t \ \ = t \ rm_vars (subst_domain \ - subst_domain \ \ fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_subst_eq': "t \ \ = t \ rm_vars (UNIV - fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_comp: assumes "range_vars \ \ vs = {}" shows "t \ rm_vars vs (\ \\<^sub>s \) = t \ (rm_vars vs \ \\<^sub>s rm_vars vs \)" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ vs") case True thus ?thesis using Var by (simp add: subst_compose_def) next case False have "subst_domain (rm_vars vs \) \ vs = {}" by (auto simp add: subst_domain_def) moreover have "fv (\ x) \ vs = {}" using Var False unfolding range_vars_alt_def by force ultimately have "\ x \ (rm_vars vs \) = \ x \ \" using rm_vars_ident by (simp add: subst_domain_def) moreover have "(rm_vars vs (\ \\<^sub>s \)) x = (\ \\<^sub>s \) x" by (metis False) ultimately show ?thesis by (auto simp: subst_compose) qed next case Fun thus ?case by auto qed lemma rm_vars_fv\<^sub>s\<^sub>e\<^sub>t_subst: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (rm_vars X \ ` Y)" shows "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) \ x \ X" using assms by auto lemma disj_dom_img_var_notin: assumes "subst_domain \ \ range_vars \ = {}" "\ v = t" "t \ Var v" shows "v \ fv t" "\v \ fv (t \ \). v \ subst_domain \" proof - have "v \ subst_domain \" "fv t \ range_vars \" using fv_in_subst_img[of \ v t, OF assms(2)] assms(2,3) by (auto simp add: subst_domain_def) thus "v \ fv t" using assms(1) by auto have *: "fv t \ subst_domain \ = {}" using assms(1) \fv t \ range_vars \\ by auto hence "t \ \ = t" by blast thus "\v \ fv (t \ \). v \ subst_domain \" using * by auto qed lemma subst_sends_dom_to_img: "v \ subst_domain \ \ fv (Var v \ \) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_sends_fv_to_img: "fv (t \ s) \ fv t \ range_vars s" proof (induction t) case (Var v) thus ?case proof (cases "Var v \ s = Var v") case True thus ?thesis by simp next case False hence "v \ subst_domain s" by (meson trm_subst_ident') hence "fv (Var v \ s) \ range_vars s" using subst_sends_dom_to_img by simp thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma ident_comp_subst_trm_if_disj: assumes "subst_domain \ \ range_vars \ = {}" "v \ subst_domain \" shows "(\ \\<^sub>s \) v = \ v" proof - from assms have " subst_domain \ \ fv (\ v) = {}" using fv_in_subst_img unfolding range_vars_alt_def by auto thus "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast qed lemma ident_comp_subst_trm_if_disj': "fv (\ v) \ subst_domain \ = {} \ (\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast lemma subst_idemI[intro]: "subst_domain \ \ range_vars \ = {} \ subst_idem \" using ident_comp_subst_trm_if_disj[of \ \] var_not_in_subst_dom[of _ \] subst_eq_if_eq_vars[of \] by (metis subst_idem_def subst_compose_def var_comp(2)) lemma subst_idemI'[intro]: "ground (subst_range \) \ subst_idem \" proof (intro subst_idemI) assume "ground (subst_range \)" hence "range_vars \ = {}" by (metis ground_range_vars) thus "subst_domain \ \ range_vars \ = {}" by blast qed lemma subst_idemE: "subst_idem \ \ subst_domain \ \ range_vars \ = {}" proof - assume "subst_idem \" hence "\v. fv (\ v) \ subst_domain \ = {}" unfolding subst_idem_def subst_compose_def by (metis trm_subst_disj) thus ?thesis unfolding range_vars_alt_def by auto qed lemma subst_idem_rm_vars: "subst_idem \ \ subst_idem (rm_vars X \)" proof - assume "subst_idem \" hence "subst_domain \ \ range_vars \ = {}" by (metis subst_idemE) moreover have "subst_domain (rm_vars X \) \ subst_domain \" "range_vars (rm_vars X \) \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately show ?thesis by blast qed lemma subst_fv_bounded_if_img_bounded: "range_vars \ \ fv t \ V \ fv (t \ \) \ fv t \ V" proof (induction t) case (Var v) thus ?case unfolding range_vars_alt_def by (cases "\ v = Var v") auto qed (metis (no_types, lifting) Un_assoc Un_commute subst_sends_fv_to_img sup.absorb_iff2) lemma subst_fv_bound_singleton: "fv (t \ Var(v := t')) \ fv t \ fv t'" using subst_fv_bounded_if_img_bounded[of "Var(v := t')" t "fv t'"] unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_fv_bounded_if_img_bounded': assumes "range_vars \ \ fv\<^sub>s\<^sub>e\<^sub>t M" shows "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" proof fix v assume *: "v \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" obtain t where t: "t \ M" "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" "v \ fv (t \ \)" proof - assume **: "\t. \t \ M; t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \; v \ fv (t \ \)\ \ thesis" have "v \ \ (fv ` ((\t. t \ \) ` M))" using * by (metis fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "\t. t \ M \ v \ fv (t \ \)" by blast thus ?thesis using ** imageI by blast qed from \t \ M\ obtain M' where "t \ M'" "M = insert t M'" by (meson Set.set_insert) hence "fv\<^sub>s\<^sub>e\<^sub>t M = fv t \ fv\<^sub>s\<^sub>e\<^sub>t M'" by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded assms by simp thus "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms \v \ fv (t \ \)\ by auto qed lemma ground_img_if_ground_subst: "(\v t. s v = t \ fv t = {}) \ range_vars s = {}" unfolding range_vars_alt_def by auto lemma ground_subst_fv_subset: "ground (subst_range \) \ fv (t \ \) \ fv t" using subst_fv_bounded_if_img_bounded[of \] unfolding range_vars_alt_def by force lemma ground_subst_fv_subset': "ground (subst_range \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded'[of \ M] unfolding range_vars_alt_def by auto lemma subst_to_var_is_var[elim]: "t \ s = Var v \ \w. t = Var w" by (auto elim!: eval_term.elims) lemma subst_dom_comp_inI: assumes "y \ subst_domain \" and "y \ subst_domain \" shows "y \ subst_domain (\ \\<^sub>s \)" using assms subst_domain_subst_compose[of \ \] by blast lemma subst_comp_notin_dom_eq: "x \ subst_domain \1 \ (\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by fastforce lemma subst_dom_comp_eq: assumes "subst_domain \ \ range_vars \ = {}" shows "subst_domain (\ \\<^sub>s \) = subst_domain \ \ subst_domain \" proof (rule ccontr) assume "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" hence "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" using subst_domain_compose[of \ \] by (simp add: subst_domain_def) then obtain v where "v \ subst_domain (\ \\<^sub>s \)" "v \ subst_domain \ \ subst_domain \" by auto hence v_in_some_subst: "\ v \ Var v \ \ v \ Var v" and "\ v \ \ = Var v" unfolding subst_compose_def by (auto simp add: subst_domain_def) then obtain w where "\ v = Var w" using subst_to_var_is_var by fastforce show False proof (cases "v = w") case True hence "\ v = Var v" using \\ v = Var w\ by simp hence "\ v \ Var v" using v_in_some_subst by simp thus False using \\ v = Var v\ \\ v \ \ = Var v\ by simp next case False hence "v \ subst_domain \" using v_in_some_subst \\ v \ \ = Var v\ by auto hence "v \ range_vars \" using assms by auto moreover have "\ w = Var v" using \\ v \ \ = Var v\ \\ v = Var w\ by simp hence "v \ range_vars \" using \v \ w\ subst_fv_imgI[of \ w] by simp ultimately show False .. qed qed lemma subst_img_comp_subset[simp]: "range_vars (\1 \\<^sub>s \2) \ range_vars \1 \ range_vars \2" proof let ?img = "range_vars" fix x assume "x \ ?img (\1 \\<^sub>s \2)" then obtain v t where vt: "x \ fv t" "t = (\1 \\<^sub>s \2) v" "t \ Var v" unfolding range_vars_alt_def subst_compose_def by (auto simp add: subst_domain_def) { assume "x \ ?img \1" hence "x \ ?img \2" by (metis (no_types, opaque_lifting) fv_in_subst_img Un_iff subst_compose_def vt subsetCE eval_term.simps(1) subst_sends_fv_to_img) } thus "x \ ?img \1 \ ?img \2" by auto qed lemma subst_img_comp_subset': assumes "t \ subst_range (\1 \\<^sub>s \2)" shows "t \ subst_range \2 \ (\t' \ subst_range \1. t = t' \ \2)" proof - obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "(\1 \\<^sub>s \2) x = t" "t \ Var x" using assms by (auto simp add: subst_domain_def) { assume "x \ subst_domain \1" hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto hence ?thesis using x by auto } moreover { assume "x \ subst_domain \1" hence ?thesis using subst_compose x(2) by fastforce } ultimately show ?thesis by metis qed lemma subst_img_comp_subset'': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1)) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2))" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t \ subterms ((\1 \\<^sub>s \2) x)" by auto show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True thus ?thesis using subst_compose[of \1 \2] x(2) subterms_subst by fastforce next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed qed lemma subst_img_comp_subset''': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var" then obtain f T where fT: "t = Fun f T" by (cases t) simp_all then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "Fun f T \ subterms ((\1 \\<^sub>s \2) x)" using t by auto have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True hence "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2)) \ (subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2)" using x(2) by (auto simp: subst_compose subterms_subst) moreover have ?thesis when *: "Fun f T \ subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2" proof - - obtain s where s: "s \ subterms (\1 x)" "Fun f T = s \ \2" using * by moura + obtain s where s: "s \ subterms (\1 x)" "Fun f T = s \ \2" using * by atomize_elim auto show ?thesis proof (cases s) case (Var y) hence "Fun f T \ subst_range \2" using s by force thus ?thesis by blast next case (Fun g S) hence "Fun f T \ (subterms (\1 x) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2" using s by blast thus ?thesis using True by auto qed qed ultimately show ?thesis by blast next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" using fT by auto qed lemma subst_img_comp_subset_const: assumes "Fun c [] \ subst_range (\1 \\<^sub>s \2)" shows "Fun c [] \ subst_range \2 \ Fun c [] \ subst_range \1 \ (\x. Var x \ subst_range \1 \ \2 x = Fun c [])" proof (cases "Fun c [] \ subst_range \2") case False then obtain t where t: "t \ subst_range \1" "Fun c [] = t \ \2" using subst_img_comp_subset'[OF assms] by auto thus ?thesis by (cases t) auto qed (simp add: subst_img_comp_subset'[OF assms]) lemma subst_img_comp_subset_const': fixes \ \::"('f,'v) subst" assumes "(\ \\<^sub>s \) x = Fun c []" shows "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" proof (cases "\ x = Fun c []") case False then obtain t where "\ x = t" "t \ \ = Fun c []" using assms unfolding subst_compose_def by auto thus ?thesis by (cases t) auto qed simp lemma subst_img_comp_subset_ground: assumes "ground (subst_range \1)" shows "subst_range (\1 \\<^sub>s \2) \ subst_range \1 \ subst_range \2" proof fix t assume t: "t \ subst_range (\1 \\<^sub>s \2)" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t = (\1 \\<^sub>s \2) x" by auto show "t \ subst_range \1 \ subst_range \2" proof (cases "x \ subst_domain \1") case True hence "fv (\1 x) = {}" using assms ground_subst_range_empty_fv by fast hence "t = \1 x" using x(2) unfolding subst_compose_def by blast thus ?thesis using True by simp next case False hence "t = \2 x" "x \ subst_domain \2" using x subst_domain_compose[of \1 \2] by (metis subst_comp_notin_dom_eq, blast) thus ?thesis using x by simp qed qed lemma subst_fv_dom_img_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "subst_domain \ = {v}" "range_vars \ = fv t" proof - show "subst_domain \ = {v}" using assms by (fastforce simp add: subst_domain_def) have "fv t \ range_vars \" by (metis fv_in_subst_img assms(1,2) vars_iff_subterm_or_eq) moreover have "\v. \ v \ Var v \ \ v = t" using assms by fastforce ultimately show "range_vars \ = fv t" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) qed lemma subst_comp_upd1: "\(v := t) \\<^sub>s \ = (\ \\<^sub>s \)(v := t \ \)" unfolding subst_compose_def by auto lemma subst_comp_upd2: assumes "v \ subst_domain s" "v \ range_vars s" shows "s(v := t) = s \\<^sub>s (Var(v := t))" unfolding subst_compose_def proof - { fix w have "(s(v := t)) w = s w \ Var(v := t)" proof (cases "w = v") case True hence "s w = Var w" using \v \ subst_domain s\ by (simp add: subst_domain_def) thus ?thesis using \w = v\ by simp next case False hence "(s(v := t)) w = s w" by simp moreover have "s w \ Var(v := t) = s w" using \w \ v\ \v \ range_vars s\ by (metis fv_in_subst_img fun_upd_apply insert_absorb insert_subset repl_invariance eval_term.simps(1) subst_apply_term_empty) ultimately show ?thesis .. qed } thus "s(v := t) = (\w. s w \ Var(v := t))" by auto qed lemma ground_subst_dom_iff_img: "ground (subst_range \) \ x \ subst_domain \ \ \ x \ subst_range \" by (auto simp add: subst_domain_def) lemma finite_dom_subst_exists: "finite S \ \\::('f,'v) subst. subst_domain \ = S" proof (induction S rule: finite.induct) case (insertI A a) then obtain \::"('f,'v) subst" where "subst_domain \ = A" by blast fix f::'f have "subst_domain (\(a := Fun f [])) = insert a A" using \subst_domain \ = A\ by (auto simp add: subst_domain_def) thus ?case by metis qed (auto simp add: subst_domain_def) lemma subst_inj_is_bij_betw_dom_img_if_ground_img: assumes "ground (subst_range \)" shows "inj \ \ bij_betw \ (subst_domain \) (subst_range \)" (is "?A \ ?B") proof show "?A \ ?B" by (metis bij_betw_def injD inj_onI subst_range.simps) next assume ?B hence "inj_on \ (subst_domain \)" unfolding bij_betw_def by auto moreover have "\x. x \ UNIV - subst_domain \ \ \ x = Var x" by auto hence "inj_on \ (UNIV - subst_domain \)" using inj_onI[of "UNIV - subst_domain \"] by (metis term.inject(1)) moreover have "\x y. x \ subst_domain \ \ y \ subst_domain \ \ \ x \ \ y" using assms by (auto simp add: subst_domain_def) ultimately show ?A by (metis injI inj_onD subst_domI term.inject(1)) qed lemma bij_finite_ground_subst_exists: assumes "finite (S::'v set)" "infinite (U::('f,'v) term set)" "ground U" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ U" proof - obtain T' where "T' \ U" "card T' = card S" "finite T'" by (meson assms(2) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ ('f,'v) term" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis hence *: "\v. v \ S \ f v \ Var v" using \ground U\ \T' \ U\ bij_betwE by fastforce let ?\ = "\v. if v \ S then f v else Var v" have "subst_domain ?\ = S" proof show "subst_domain ?\ \ S" by (auto simp add: subst_domain_def) { fix v assume "v \ S" "v \ subst_domain ?\" hence "f v = Var v" by (simp add: subst_domain_def) hence False using *[OF \v \ S\] by metis } thus "S \ subst_domain ?\" by blast qed hence "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" using \ground U\ bij_betwE[OF f_bij] set_rev_mp[OF _ \T' \ U\] by (metis (no_types, lifting) UN_iff empty_iff vars_iff_subterm_or_eq fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "inj_on ?\ (subst_domain ?\)" using f_bij \subst_domain ?\ = S\ unfolding bij_betw_def inj_on_def by metis hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = T'" using \bij_betw f S T'\ \subst_domain ?\ = S\ unfolding bij_betw_def by auto hence "subst_range ?\ \ U" using \T' \ U\ by auto ultimately show ?thesis using \subst_domain ?\ = S\ by (metis (lifting)) qed lemma bij_finite_const_subst_exists: assumes "finite (S::'v set)" "finite (T::'f set)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ (\c. Fun c []) ` (U - T)" proof - obtain T' where "T' \ U - T" "card T' = card S" "finite T'" by (meson assms(2,3) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ 'f" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis let ?\ = "\v. if v \ S then Fun (f v) [] else Var v" have "subst_domain ?\ = S" by (simp add: subst_domain_def) moreover have "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" by auto hence "inj_on ?\ (subst_domain ?\)" using f_bij unfolding bij_betw_def inj_on_def by (metis \subst_domain ?\ = S\ term.inject(2)) hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = ((\c. Fun c []) ` T')" using \bij_betw f S T'\ unfolding bij_betw_def inj_on_def by (auto simp add: subst_domain_def) hence "subst_range ?\ \ ((\c. Fun c []) ` (U - T))" using \T' \ U - T\ by auto ultimately show ?thesis by (metis (lifting)) qed lemma bij_finite_const_subst_exists': assumes "finite (S::'v set)" "finite (T::('f,'v) terms)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` U) - T" proof - have "finite (\(funs_term ` T))" using assms(2) by auto then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ (\c. Fun c []) ` (U - (\(funs_term ` T)))" using bij_finite_const_subst_exists[OF assms(1) _ assms(3)] by blast moreover have "(\c. Fun c []) ` (U - (\(funs_term ` T))) \ ((\c. Fun c []) ` U) - T" by auto ultimately show ?thesis by blast qed lemma bij_betw_iteI: assumes "bij_betw f A B" "bij_betw g C D" "A \ C = {}" "B \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ C) (B \ D)" proof - have "bij_betw (\x. if x \ A then f x else g x) A B" by (metis bij_betw_cong[of A f "\x. if x \ A then f x else g x" B] assms(1)) moreover have "bij_betw (\x. if x \ A then f x else g x) C D" using bij_betw_cong[of C g "\x. if x \ A then f x else g x" D] assms(2,3) by force ultimately show ?thesis using bij_betw_combine[OF _ _ assms(4)] by metis qed lemma subst_comp_split: assumes "subst_domain \ \ range_vars \ = {}" shows "\ = (rm_vars (subst_domain \ - V) \) \\<^sub>s (rm_vars V \)" (is ?P) and "\ = (rm_vars V \) \\<^sub>s (rm_vars (subst_domain \ - V) \)" (is ?Q) proof - let ?rm1 = "rm_vars (subst_domain \ - V) \" and ?rm2 = "rm_vars V \" have "subst_domain ?rm2 \ range_vars ?rm1 = {}" "subst_domain ?rm1 \ range_vars ?rm2 = {}" using assms unfolding range_vars_alt_def by (force simp add: subst_domain_def)+ hence *: "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" using ident_comp_subst_trm_if_disj[of ?rm2 ?rm1] ident_comp_subst_trm_if_disj[of ?rm1 ?rm2] by (auto simp add: subst_domain_def) hence "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" unfolding subst_compose_def by (auto simp add: subst_domain_def) hence "\v. (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. (?rm2 \\<^sub>s ?rm1) v = \ v" using * by blast+ thus ?P ?Q by auto qed lemma subst_comp_eq_if_disjoint_vars: assumes "(subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" shows "\ \\<^sub>s \ = \ \\<^sub>s \" proof - { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (force simp add: subst_compose)+ hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (auto simp add: subst_compose subst_domain_def) hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" "x \ subst_domain \" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by (simp add: subst_compose subst_domain_def) } ultimately show ?thesis by auto qed lemma subst_eq_if_disjoint_vars_ground: fixes \ \::"('f,'v) subst" assumes "subst_domain \ \ subst_domain \ = {}" "ground (subst_range \)" "ground (subst_range \)" shows "t \ \ \ \ = t \ \ \ \" by (metis assms subst_comp_eq_if_disjoint_vars range_vars_alt_def subst_subst_compose sup_bot.right_neutral) lemma subst_img_bound: "subst_domain \ \ range_vars \ \ fv t \ range_vars \ \ fv (t \ \)" proof - assume "subst_domain \ \ range_vars \ \ fv t" hence "subst_domain \ \ fv t" by blast thus ?thesis by (metis (no_types) range_vars_alt_def le_iff_sup subst_apply_fv_unfold subst_apply_fv_union subst_range.simps) qed lemma subst_all_fv_subset: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof - assume *: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" { fix v assume "v \ fv t" hence "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using * by auto then obtain t' where "t' \ M" "v \ fv t'" by auto hence "fv (\ v) \ fv (t' \ \)" by (metis eval_term.simps(1) subst_apply_fv_subset subst_apply_fv_unfold subtermeq_vars_subset vars_iff_subtermeq) hence "fv (\ v) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using \t' \ M\ by auto } thus ?thesis by (auto simp: subst_apply_fv_unfold) qed lemma subst_support_if_mgt_subst_idem: assumes "\ \\<^sub>\ \" "subst_idem \" shows "\ supports \" proof - from \\ \\<^sub>\ \\ obtain \ where \: "\ = \ \\<^sub>s \" by blast hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \ \\<^sub>s \)" by (simp add: subst_compose) hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \)" using \subst_idem \ \ unfolding subst_idem_def by simp hence "\v. \ v \ \ = Var v \ \" using \ by simp thus "\ supports \" by simp qed lemma subst_support_iff_mgt_if_subst_idem: assumes "subst_idem \" shows "\ \\<^sub>\ \ \ \ supports \" proof show "\ \\<^sub>\ \ \ \ supports \" by (fact subst_support_if_mgt_subst_idem[OF _ \subst_idem \\]) show "\ supports \ \ \ \\<^sub>\ \" by (fact subst_supportD) qed lemma subst_support_comp: fixes \ \ \::"('a,'b) subst" assumes "\ supports \" "\ supports \" shows "(\ \\<^sub>s \) supports \" by (metis (no_types) assms subst_agreement eval_term.simps(1) subst_subst_compose) lemma subst_support_comp': fixes \ \ \::"('a,'b) subst" assumes "\ supports \" shows "\ supports (\ \\<^sub>s \)" "\ supports \ \ \ supports (\ \\<^sub>s \)" using assms unfolding subst_support_def by (metis subst_compose_assoc, metis) lemma subst_support_comp_split: fixes \ \ \::"('a,'b) subst" assumes "(\ \\<^sub>s \) supports \" shows "subst_domain \ \ range_vars \ = {} \ \ supports \" and "subst_domain \ \ subst_domain \ = {} \ \ supports \" proof - assume "subst_domain \ \ range_vars \ = {}" hence "subst_idem \" by (metis subst_idemI) have "\ \\<^sub>\ \" using assms subst_compose_assoc[of \ \ \] unfolding subst_compose_def by metis show "\ supports \" using subst_support_if_mgt_subst_idem[OF \\ \\<^sub>\ \\ \subst_idem \\] by auto next assume "subst_domain \ \ subst_domain \ = {}" moreover have "\v \ subst_domain (\ \\<^sub>s \). (\ \\<^sub>s \) v \ \ = \ v" using assms by metis ultimately have "\v \ subst_domain \. \ v \ \ = \ v" using var_not_in_subst_dom unfolding subst_compose_def by (metis IntI empty_iff eval_term.simps(1)) thus "\ supports \" by force qed lemma subst_idem_support: "subst_idem \ \ \ supports \ \\<^sub>s \" unfolding subst_idem_def by (metis subst_support_def subst_compose_assoc) lemma subst_idem_iff_self_support: "subst_idem \ \ \ supports \" using subst_support_def[of \ \] unfolding subst_idem_def by auto lemma subterm_subst_neq: "t \ t' \ t \ s \ t' \ s" by (metis subst_mono_neq) lemma fv_Fun_subst_neq: "x \ fv (Fun f T) \ \ x \ Fun f T \ \" using subterm_subst_neq[of "Var x" "Fun f T"] vars_iff_subterm_or_eq[of x "Fun f T"] by auto lemma subterm_subst_unfold: assumes "t \ s \ \" shows "(\s'. s' \ s \ t = s' \ \) \ (\x \ fv s. t \ \ x)" using assms proof (induction s) case (Fun f T) thus ?case proof (cases "t = Fun f T \ \") case True thus ?thesis using Fun by auto next case False then obtain s' where s': "s' \ set T" "t \ s' \ \" using Fun by auto hence "(\s''. s'' \ s' \ t = s'' \ \) \ (\x \ fv s'. t \ \ x)" by (metis Fun.IH) thus ?thesis using s'(1) by auto qed qed simp lemma subterm_subst_img_subterm: assumes "t \ s \ \" "\s'. s' \ s \ t \ s' \ \" shows "\w \ fv s. t \ \ w" using subterm_subst_unfold[OF assms(1)] assms(2) by force lemma subterm_subst_not_img_subterm: assumes "t \ s \ \" "\(\w \ fv s. t \ \ w)" shows "\f T. Fun f T \ s \ t = Fun f T \ \" proof (rule ccontr) assume "\(\f T. Fun f T \ s \ t = Fun f T \ \)" hence "\f T. Fun f T \ s \ t \ Fun f T \ \" by simp moreover have "\x. Var x \ s \ t \ Var x \ \" using assms(2) vars_iff_subtermeq by force ultimately have "\s'. s' \ s \ t \ s' \ \" by (metis "term.exhaust") thus False using assms subterm_subst_img_subterm by blast qed lemma subst_apply_img_var: assumes "v \ fv (t \ \)" "v \ fv t" obtains w where "w \ fv t" "v \ fv (\ w)" using assms by (induct t) auto lemma subst_apply_img_var': assumes "x \ fv (t \ \)" "x \ fv t" shows "\y \ fv t. x \ fv (\ y)" by (metis assms subst_apply_img_var) lemma nth_map_subst: fixes \::"('f,'v) subst" and T::"('f,'v) term list" and i::nat shows "i < length T \ (map (\t. t \ \) T) ! i = (T ! i) \ \" by (fact nth_map) lemma subst_subterm: assumes "Fun f T \ t \ \" shows "(\S. Fun f S \ t \ Fun f S \ \ = Fun f T) \ (\s \ subst_range \. Fun f T \ s)" using assms subterm_subst_not_img_subterm by (cases "\s \ subst_range \. Fun f T \ s") fastforce+ lemma subst_subterm': assumes "Fun f T \ t \ \" shows "\S. length S = length T \ (Fun f S \ t \ (\s \ subst_range \. Fun f S \ s))" using subst_subterm[OF assms] by auto lemma subst_subterm'': assumes "s \ subterms (t \ \)" shows "(\u \ subterms t. s = u \ \) \ s \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" proof (cases s) case (Var x) thus ?thesis using assms subterm_subst_not_img_subterm vars_iff_subtermeq by (cases "s = t \ \") fastforce+ next case (Fun f T) thus ?thesis using subst_subterm assms by fastforce qed lemma fv_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv (t \ \ \\<^sub>s \) = fv (t \ \ \\<^sub>s \)" proof (induction t) case (Var x) show ?case proof (cases "x \ subst_domain \") case True thus ?thesis using assms unfolding range_vars_alt_def by (auto simp: subst_compose subst_apply_fv_empty) next case False hence "\ x = Var x" "\ x = Var x" using assms(1) by (blast,blast) thus ?thesis by (simp add: subst_compose) qed qed simp subsection \More Small Lemmata\ lemma funs_term_subst: "funs_term (t \ \) = funs_term t \ (\x \ fv t. funs_term (\ x))" by (induct t) auto lemma fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq: assumes "X \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (Y - X)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) - X" using assms unfolding range_vars_alt_def by force lemma subst_Fun_index_eq: assumes "i < length T" "Fun f T \ \ = Fun g T' \ \" shows "T ! i \ \ = T' ! i \ \" proof - have "map (\x. x \ \) T = map (\x. x \ \) T'" using assms by simp thus ?thesis by (metis assms(1) length_map nth_map) qed lemma fv_exists_if_unifiable_and_neq: fixes t t'::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ t'" "t \ \ = t' \ \" shows "fv t \ fv t' \ {}" proof assume "fv t \ fv t' = {}" hence "fv t = {}" "fv t' = {}" by auto hence "t \ \ = t" "t' \ \ = t'" by auto hence "t = t'" using assms(2) by metis thus False using assms(1) by auto qed lemma const_subterm_subst: "Fun c [] \ t \ Fun c [] \ t \ \" by (induct t) auto lemma const_subterm_subst_var_obtain: assumes "Fun c [] \ t \ \" "\Fun c [] \ t" obtains x where "x \ fv t" "Fun c [] \ \ x" using assms by (induct t) auto lemma const_subterm_subst_cases: assumes "Fun c [] \ t \ \" shows "Fun c [] \ t \ (\x \ fv t. x \ subst_domain \ \ Fun c [] \ \ x)" proof (cases "Fun c [] \ t") case False then obtain x where "x \ fv t" "Fun c [] \ \ x" - using const_subterm_subst_var_obtain[OF assms] by moura + using const_subterm_subst_var_obtain[OF assms] by atomize_elim auto thus ?thesis by (cases "x \ subst_domain \") auto qed simp lemma const_subterms_subst_cases: assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. x \ subst_domain \ \ Fun c [] \ \ x)" using assms const_subterm_subst_cases[of c _ \] by auto lemma const_subterms_subst_cases': assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ Fun c [] \\<^sub>s\<^sub>e\<^sub>t subst_range \" using const_subterms_subst_cases[OF assms] by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset: assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" shows "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using assms proof (induction F) case (Cons f F) then obtain t t' where f: "f = (t,t')" by (metis surj_pair) show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F") case True thus ?thesis using Cons.IH unfolding subst_apply_pairs_def by auto next case False hence "x \ fv t \ fv t'" using Cons.prems f by simp hence "fv (\ x) \ fv (t \ \) \ fv (t' \ \)" using fv_subst_subset[of x] by force thus ?thesis using f unfolding subst_apply_pairs_def by auto qed qed simp lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof (induction F) case (Cons f F) - obtain t t' where "f = (t,t')" by moura + obtain t t' where "f = (t,t')" by atomize_elim auto thus ?case using Cons by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp_all add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "\y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. x \ fv (\ y)" using assms proof (induction F) case (Cons f F) then obtain t s where f: "f = (t,s)" by (metis surj_pair) from Cons.IH show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False hence "x \ fv (t \ \) \ x \ fv (s \ \)" using f Cons.prems by (simp add: subst_apply_pairs_def) hence "(\y \ fv t. x \ fv (\ y)) \ (\y \ fv s. x \ fv (\ y))" by (metis fv_subst_obtain_var) thus ?thesis using f by (auto simp add: subst_apply_pairs_def) qed (auto simp add: Cons.IH) qed (simp add: subst_apply_pairs_def) lemma pair_subst_ident[intro]: "(fv t \ fv t') \ subst_domain \ = {} \ (t,t') \\<^sub>p \ = (t,t')" by auto lemma pairs_substI[intro]: assumes "subst_domain \ \ (\(s,t) \ M. fv s \ fv t) = {}" shows "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = M" proof - { fix m assume M: "m \ M" then obtain s t where m: "m = (s,t)" by (metis surj_pair) hence "(fv s \ fv t) \ subst_domain \ = {}" using assms M by auto hence "m \\<^sub>p \ = m" using m by auto } thus ?thesis by (simp add: image_cong) qed lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F))" proof (induction F) case (Cons g G) obtain t t' where "g = (t,t')" by (metis surj_pair) thus ?case using Cons.IH by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset: assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" shows "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ subst_domain \ \ subst_domain \" using assms proof (induction F) case (Cons g G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ subst_domain \ \ subst_domain \" by (simp add: subst_apply_pairs_def) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv (t \ \) \ subst_domain \" "fv (t' \ \) \ subst_domain \" using Cons.prems by (simp_all add: subst_apply_pairs_def) hence "fv t \ subst_domain \ \ subst_domain \" "fv t' \ subst_domain \ \ subst_domain \" unfolding subst_apply_fv_unfold by force+ thus ?case using IH g by (simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) lemma pairs_subst_comp: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \\<^sub>s \ = ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (induct F) (auto simp add: subst_apply_pairs_def) lemma pairs_substI'[intro]: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {} \ F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by (induct F) (force simp add: subst_apply_pairs_def)+ lemma subst_pair_compose[simp]: "d \\<^sub>p (\ \\<^sub>s \) = d \\<^sub>p \ \\<^sub>p \" proof - - obtain t s where "d = (t,s)" by moura + obtain t s where "d = (t,s)" by atomize_elim auto thus ?thesis by auto qed lemma subst_pairs_compose[simp]: "D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by auto lemma subst_apply_pair_pair: "(t, s) \\<^sub>p \ = (t \ \, s \ \)" by (rule prod.case) lemma subst_apply_pairs_nil[simp]: "[] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = []" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_singleton[simp]: "[(t,s)] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = [(t \ \,s \ \)]" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_Var[iff]: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s Var = F" by (simp add: subst_apply_pairs_def) lemma subst_apply_pairs_pset_subst: "set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" unfolding subst_apply_pairs_def by force lemma subst_subterms: "t \\<^sub>s\<^sub>e\<^sub>t M \ t \ \ \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_mono_neq by fastforce lemma subst_subterms_fv: "x \ fv\<^sub>s\<^sub>e\<^sub>t M \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using fv_subterms_substI by fastforce lemma subst_subterms_Var: "Var x \\<^sub>s\<^sub>e\<^sub>t M \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_subterms_fv[of x M \] by force lemma fv_subset_subterms_subset: "\ ` fv\<^sub>s\<^sub>e\<^sub>t M \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_subterms_fv by fast lemma subst_const_swap_eq: fixes \ \::"('a,'b) subst" assumes t: "t \ \ = s \ \" and \: "\x \ fv t \ fv s. \k. \ x = Fun k []" "\x \ fv t. \(\ x \ s)" "\x \ fv s. \(\ x \ t)" and \_def: "\ \ \x. p (\ x)" shows "t \ \ = s \ \" using t \ proof (induction t arbitrary: s) case (Var x) thus ?case unfolding \_def by (cases s) auto next case (Fun f ts) note prems = Fun.prems obtain ss where s: "s = Fun f ss" and ss: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using prems by (cases s) auto have "ts ! i \ \ = ss ! i \ \" when i: "i < length ts" for i proof - have *: "ts ! i \ set ts" using i by simp have **: "ts ! i \ \ = ss ! i \ \" using i prems(1) unfolding s by (metis subst_Fun_index_eq) have ***: "ss ! i \ set ss" using i ss by (metis length_map nth_mem) show ?thesis using Fun.IH[OF * **] prems(2,3,4) * *** unfolding s by auto qed hence IH: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using ss by (metis (mono_tags, lifting) length_map nth_equalityI nth_map) show ?case using IH unfolding s by auto qed lemma term_subst_set_eq: assumes "\x. x \ fv\<^sub>s\<^sub>e\<^sub>t M \ \ x = \ x" shows "M \\<^sub>s\<^sub>e\<^sub>t \ = M \\<^sub>s\<^sub>e\<^sub>t \" proof - have "t \ \ = t \ \" when "t \ M" for t using that assms term_subst_eq[of _ \ \] by fastforce thus ?thesis by simp qed lemma subst_const_swap_eq': assumes "t \ \ = s \ \" and "\x \ fv t \ fv s. \ x = \ x \ \(\ x \ t) \ \(\ x \ s)" (is "?A t s") and "\x \ fv t \ fv s. \c. \ x = Fun c []" (is "?B t s") and "\x \ fv t \ fv s. \c. \ x = Fun c []" (is "?C t s") and "\x \ fv t \ fv s. \y \ fv t \ fv s. \ x = \ y \ \ x = \ y" (is "?D t s") shows "t \ \ = s \ \" using assms proof (induction t arbitrary: s) case (Var x) note prems = Var.prems have "(\y. s = Var y) \ (\c. s = Fun c [])" using prems(1,3) by (cases s) auto thus ?case proof assume "\y. s = Var y" then obtain y where y: "s = Var y" by blast hence "\ x = \ y" using prems(1) by simp hence "\ x = \ y" using prems(5) y by fastforce thus ?thesis using y by force next assume "\c. s = Fun c []" then obtain c where c: "s = Fun c []" by blast have "\ x = \ x \ \(\ x \ Fun c [])" using prems(2) c by auto thus ?thesis using prems(1) c by simp qed next case (Fun f ts) note prems = Fun.prems note IH = Fun.IH show ?case proof (cases s) case (Var x) note s = this hence ts: "ts = []" using prems(1,3) by auto show ?thesis using prems unfolding s ts by auto next case (Fun g ss) note s = this hence g: "f = g" using prems(1) by fastforce have ss: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using prems(1) unfolding s by (cases s) auto have len: "length ts = length ss" using ss by (metis length_map) have "ts ! i \ \ = ss ! i \ \" when i: "i < length ts" for i proof - have 0: "ts ! i \ set ts" using i by simp have 1: "ts ! i \ \ = ss ! i \ \" using i prems(1) unfolding s by (metis subst_Fun_index_eq) have 2: "ss ! i \ set ss" using i by (metis len nth_mem) have 3: "fv (ts ! i) \ fv (Fun f ts)" "fv (ss ! i) \ fv (Fun g ss)" "subterms (ts ! i) \ subterms (Fun f ts)" "subterms (ss ! i) \ subterms (Fun g ss)" subgoal by (meson index_Fun_fv_subset i) subgoal by (metis index_Fun_fv_subset i len) subgoal using ss i by fastforce subgoal using ss i len by fastforce done have 4: "?A (ts ! i) (ss ! i)" "?B (ts ! i) (ss ! i)" "?C (ts ! i) (ss ! i)" "?D (ts ! i) (ss ! i)" subgoal using 3 prems(2) unfolding s by blast subgoal using 3(1,2) prems(3) unfolding s by blast subgoal using 3(1,2) prems(4) unfolding s by blast subgoal using 3(1,2) prems(5) unfolding s by blast done thus ?thesis using IH[OF 0 1 4] prems(2-) 0 2 unfolding s by blast qed hence "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis (mono_tags, lifting) ss length_map nth_equalityI nth_map) thus ?thesis unfolding s g by auto qed qed lemma subst_const_swap_eq_mem: assumes "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ x \ \(\ x \\<^sub>s\<^sub>e\<^sub>t insert t M)" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" (is "?B (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" (is "?C (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ y \ \ x = \ y" (is "?D (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") shows "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" proof - let ?A = "\t s. \x \ fv t \ fv s. \ x = \ x \ \(\ x \ t) \ \(\ x \ s)" obtain s where s: "s \ M" "s \ \ = t \ \" using assms(1) by fastforce have 0: "fv s \ fv\<^sub>s\<^sub>e\<^sub>t M" "subterms s \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" "subterms t \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using s(1) by auto have 1: "?A s t" "?B (fv s \ fv t)" "?C (fv s \ fv t)" "?D (fv s \ fv t)" subgoal using assms(2) 0 by fast subgoal using assms(3) 0 by blast subgoal using assms(4) 0 by blast subgoal using assms(5) 0 by blast done have "s \ \ = t \ \" by (rule subst_const_swap_eq'[OF s(2) 1]) thus ?thesis by (metis s(1) imageI) qed subsection \Finite Substitutions\ inductive_set fsubst::"('a,'b) subst set" where fvar: "Var \ fsubst" | FUpdate: "\\ \ fsubst; v \ subst_domain \; t \ Var v\ \ \(v := t) \ fsubst" lemma finite_dom_iff_fsubst: "finite (subst_domain \) \ \ \ fsubst" proof assume "finite (subst_domain \)" thus "\ \ fsubst" proof (induction "subst_domain \" arbitrary: \ rule: finite.induct) case emptyI hence "\ = Var" using empty_dom_iff_empty_subst by metis thus ?case using fvar by simp next case (insertI \'\<^sub>d\<^sub>o\<^sub>m v) thus ?case proof (cases "v \ \'\<^sub>d\<^sub>o\<^sub>m") case True hence "\'\<^sub>d\<^sub>o\<^sub>m = subst_domain \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by auto thus ?thesis using insertI.hyps(2) by metis next case False let ?\' = "\w. if w \ \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w" have "subst_domain ?\' = \'\<^sub>d\<^sub>o\<^sub>m" using \v \ \'\<^sub>d\<^sub>o\<^sub>m\ \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) hence "?\' \ fsubst" using insertI.hyps(2) by simp moreover have "?\'(v := \ v) = (\w. if w \ insert v \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w)" by auto hence "?\'(v := \ v) = \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) ultimately show ?thesis using FUpdate[of ?\' v "\ v"] False insertI.hyps(3) by (auto simp add: subst_domain_def) qed qed next assume "\ \ fsubst" thus "finite (subst_domain \)" by (induct \, simp, metis subst_dom_insert_finite) qed lemma fsubst_induct[case_names fvar FUpdate, induct set: finite]: assumes "finite (subst_domain \)" "P Var" and "\\ v t. \finite (subst_domain \); v \ subst_domain \; t \ Var v; P \\ \ P (\(v := t))" shows "P \" using assms finite_dom_iff_fsubst fsubst.induct by metis lemma fun_upd_fsubst: "s(v := t) \ fsubst \ s \ fsubst" using subst_dom_insert_finite[of s] finite_dom_iff_fsubst by blast lemma finite_img_if_fsubst: "s \ fsubst \ finite (subst_range s)" using finite_dom_iff_fsubst finite_subst_img_if_finite_dom' by blast subsection \Unifiers and Most General Unifiers (MGUs)\ abbreviation Unifier::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "Unifier \ t u \ (t \ \ = u \ \)" abbreviation MGU::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "MGU \ t u \ Unifier \ t u \ (\\. Unifier \ t u \ \ \\<^sub>\ \)" lemma MGUI[intro]: shows "\t \ \ = u \ \; \\::('f,'v) subst. t \ \ = u \ \ \ \ \\<^sub>\ \\ \ MGU \ t u" by auto lemma UnifierD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "Unifier \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - from assms show "f = g" by auto from assms have "Fun f X \ \ = Fun g Y \ \" by auto hence "length (map (\x. x \ \) X) = length (map (\x. x \ \) Y)" by auto thus "length X = length Y" by auto qed lemma MGUD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" using assms by (auto dest: map_eq_imp_length_eq) lemma MGU_sym[sym]: "MGU \ s t \ MGU \ t s" by auto lemma Unifier_sym[sym]: "Unifier \ s t \ Unifier \ t s" by auto lemma MGU_nil: "MGU Var s t \ s = t" by fastforce lemma Unifier_comp: "Unifier (\ \\<^sub>s \) t u \ Unifier \ (t \ \) (u \ \)" by simp lemma Unifier_comp': "Unifier \ (t \ \) (u \ \) \ Unifier (\ \\<^sub>s \) t u" by simp lemma Unifier_excludes_subterm: assumes \: "Unifier \ t u" shows "\t \ u" proof assume "t \ u" hence "t \ \ \ u \ \" using subst_mono_neq by metis hence "t \ \ \ u \ \" by simp moreover from \ have "t \ \ = u \ \" by auto ultimately show False .. qed lemma MGU_is_Unifier: "MGU \ t u \ Unifier \ t u" by (rule conjunct1) lemma MGU_Var1: assumes "\Var v \ t" shows "MGU (Var(v := t)) (Var v) t" proof (intro MGUI exI) show "Var v \ (Var(v := t)) = t \ (Var(v := t))" using assms subst_no_occs by fastforce next fix \::"('a,'b) subst" assume th: "Var v \ \ = t \ \" show "\ = (Var(v := t)) \\<^sub>s \" using th by (auto simp: subst_compose_def) qed lemma MGU_Var2: "v \ fv t \ MGU (Var(v := t)) (Var v) t" by (metis (no_types) MGU_Var1 vars_iff_subterm_or_eq) lemma MGU_Var3: "MGU Var (Var v) (Var w) \ v = w" by fastforce lemma MGU_Const1: "MGU Var (Fun c []) (Fun d []) \ c = d" by fastforce lemma MGU_Const2: "MGU \ (Fun c []) (Fun d []) \ c = d" by auto lemma MGU_Fun: assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - let ?F = "\\ X. map (\x. x \ \) X" from assms have "\f = g; ?F \ X = ?F \ Y; \\'. f = g \ ?F \' X = ?F \' Y \ \ \\<^sub>\ \'\ \ length X = length Y" using map_eq_imp_length_eq by auto thus "f = g" "length X = length Y" using assms by auto qed lemma Unifier_Fun: assumes "Unifier \ (Fun f (x#X)) (Fun g (y#Y))" shows "Unifier \ x y" "Unifier \ (Fun f X) (Fun g Y)" using assms by simp_all lemma Unifier_subst_idem_subst: "subst_idem r \ Unifier s (t \ r) (u \ r) \ Unifier (r \\<^sub>s s) (t \ r) (u \ r)" by (metis (no_types, lifting) subst_idem_def subst_subst_compose) lemma subst_idem_comp: "subst_idem r \ Unifier s (t \ r) (u \ r) \ (\q. Unifier q (t \ r) (u \ r) \ s \\<^sub>s q = q) \ subst_idem (r \\<^sub>s s)" by (frule Unifier_subst_idem_subst, blast, metis subst_idem_def subst_compose_assoc) lemma Unifier_mgt: "\Unifier \ t u; \ \\<^sub>\ \\ \ Unifier \ t u" by auto lemma Unifier_support: "\Unifier \ t u; \ supports \\ \ Unifier \ t u" using subst_supportD Unifier_mgt by metis lemma MGU_mgt: "\MGU \ t u; MGU \ t u\ \ \ \\<^sub>\ \" by auto lemma Unifier_trm_fv_bound: "\Unifier s t u; v \ fv t\ \ v \ subst_domain s \ range_vars s \ fv u" proof (induction t arbitrary: s u) case (Fun f X) hence "v \ fv (u \ s) \ v \ subst_domain s" by (metis subst_not_dom_fixed) thus ?case by (metis (no_types) Un_iff contra_subsetD subst_sends_fv_to_img) qed (metis (no_types) UnI1 UnI2 subsetCE no_var_subterm subst_sends_dom_to_img subst_to_var_is_var trm_subst_ident' vars_iff_subterm_or_eq) lemma Unifier_rm_var: "\Unifier \ s t; v \ fv s \ fv t\ \ Unifier (rm_var v \) s t" by (auto simp add: repl_invariance) lemma Unifier_ground_rm_vars: assumes "ground (subst_range s)" "Unifier (rm_vars X s) t t'" shows "Unifier s t t'" by (rule Unifier_support[OF assms(2) rm_vars_ground_supports[OF assms(1)]]) lemma Unifier_dom_restrict: assumes "Unifier s t t'" "fv t \ fv t' \ S" shows "Unifier (rm_vars (UNIV - S) s) t t'" proof - let ?s = "rm_vars (UNIV - S) s" show ?thesis using term_subst_eq_conv[of t s ?s] term_subst_eq_conv[of t' s ?s] assms by auto qed subsection \Well-formedness of Substitutions and Unifiers\ inductive_set wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set::"('a,'b) subst set" where Empty[simp]: "Var \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" | Insert[simp]: "\\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; v \ subst_domain \; v \ range_vars \; fv t \ (insert v (subst_domain \)) = {}\ \ \(v := t) \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" definition wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" definition wf\<^sub>M\<^sub>G\<^sub>U::"('a,'b) subst \ ('a,'b) term \ ('a,'b) term \ bool" where "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ MGU \ s t \ subst_domain \ \ range_vars \ \ fv s \ fv t" lemma wf_subst_subst_idem: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" using subst_idemI[of \] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by fast lemma wf_subst_properties: "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set = wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof - assume "subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" hence "finite (subst_domain \)" "subst_domain \ \ range_vars \ = {}" by auto thus "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" proof (induction \ rule: fsubst_induct) case fvar thus ?case by simp next case (FUpdate \ v t) have "subst_domain \ \ subst_domain (\(v := t))" "range_vars \ \ range_vars (\(v := t))" using FUpdate.hyps(2,3) subst_img_update unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def)+ hence "subst_domain \ \ range_vars \ = {}" using FUpdate.prems(1) by blast hence "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" using FUpdate.IH by metis have *: "range_vars (\(v := t)) = range_vars \ \ fv t" using FUpdate.hyps(2) subst_img_update[OF _ FUpdate.hyps(3)] by fastforce hence "fv t \ insert v (subst_domain \) = {}" using FUpdate.prems subst_dom_update2[OF FUpdate.hyps(3)] by blast moreover have "subst_domain (\(v := t)) = insert v (subst_domain \)" by (meson FUpdate.hyps(3) subst_dom_update2) hence "v \ range_vars \" using FUpdate.prems * by blast ultimately show ?case using Insert[OF \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set\ \v \ subst_domain \\] by metis qed qed show "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct) case Empty thus ?case by simp next case (Insert \ v t) hence 1: "subst_domain \ \ range_vars \ = {}" by simp hence 2: "subst_domain (\(v := t)) \ range_vars \ = {}" using Insert.hyps(3) by (auto simp add: subst_domain_def) have 3: "fv t \ subst_domain (\(v := t)) = {}" using Insert.hyps(4) by (auto simp add: subst_domain_def) have 4: "\ v = Var v" using \v \ subst_domain \\ by (simp add: subst_domain_def) from Insert.IH have "finite (subst_domain \)" by simp hence 5: "finite (subst_domain (\(v := t)))" using subst_dom_insert_finite[of \] by simp have "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" proof (cases "t = Var v") case True hence "range_vars (\(v := t)) = range_vars \" using 4 fun_upd_triv term.inject(1) unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by auto next case False hence "range_vars (\(v := t)) = fv t \ (range_vars \)" using 4 subst_img_update[of \ v] by auto thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by blast qed thus ?case using 5 by blast qed qed lemma wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct[consumes 1, case_names Empty Insert]: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "P Var" and "\\ v t. \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" shows "P \" proof - from assms(1,3) wf_subst_properties have "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" "\\ v t. \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" by blast+ thus "P \" using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct assms(2) by blast qed lemma wf_subst_fsubst: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ fsubst" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using finite_dom_iff_fsubst by blast lemma wf_subst_nil: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp lemma wf_MGU_nil: "MGU Var s t \ wf\<^sub>M\<^sub>G\<^sub>U Var s t" using wf_subst_nil subst_domain_Var range_vars_Var unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by fast lemma wf_MGU_dom_bound: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ subst_domain \ \ fv s \ fv t" unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma wf_subst_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - have *: "subst_domain \ = {v}" by (metis subst_fv_dom_img_single(1)[OF assms]) have "subst_domain \ \ range_vars \ = {}" using * assms subst_fv_dom_img_single(2) by (metis inf_bot_left insert_disjoint(1)) moreover have "finite (subst_domain \)" using * by simp ultimately show ?thesis by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_reduction: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" proof - assume "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s" moreover have "subst_domain (rm_var v s) \ subst_domain s" by (auto simp add: subst_domain_def) moreover have "range_vars (rm_var v s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately have "subst_domain (rm_var v s) \ range_vars (rm_var v s) = {}" by (meson compl_le_compl_iff disjoint_eq_subset_Compl subset_trans wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "finite (subst_domain (rm_var v s))" using \subst_domain (rm_var v s) \ subst_domain s\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s\ rev_finite_subset unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast ultimately show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_compose: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) have "t \ Var v" using Insert.hyps(4) by auto hence dom1v_unfold: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using subst_dom_update2 by metis hence doms_disj: "subst_domain \1 \ subst_domain \2 = {}" using Insert.prems(2) disjoint_insert(1) by blast moreover have dom_img_disj: "subst_domain \1 \ range_vars \2 = {}" using Insert.hyps(2) Insert.prems(3) by (fastforce simp add: subst_domain_def) ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using Insert.IH[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\] by metis have dom_comp_is_union: "subst_domain (\1 \\<^sub>s \2) = subst_domain \1 \ subst_domain \2" using subst_dom_comp_eq[OF dom_img_disj] . have "v \ subst_domain \2" using Insert.prems(2) \t \ Var v\ by (fastforce simp add: subst_domain_def) hence "\2 v = Var v" "\1 v = Var v" using Insert.hyps(2) by (simp_all add: subst_domain_def) hence "(\1 \\<^sub>s \2) v = Var v" "(\1(v := t) \\<^sub>s \2) v = t \ \2" "((\1 \\<^sub>s \2)(v := t)) v = t" unfolding subst_compose_def by simp_all have fv_t2_bound: "fv (t \ \2) \ fv t \ range_vars \2" by (meson subst_sends_fv_to_img) have 1: "v \ subst_domain (\1 \\<^sub>s \2)" using \(\1 \\<^sub>s \2) v = Var v\ by (auto simp add: subst_domain_def) have "insert v (subst_domain \1) \ range_vars \2 = {}" using Insert.prems(3) dom1v_unfold by blast hence "v \ range_vars \1 \ range_vars \2" using Insert.hyps(3) by blast hence 2: "v \ range_vars (\1 \\<^sub>s \2)" by (meson set_rev_mp subst_img_comp_subset) have "subst_domain \2 \ range_vars \2 = {}" using \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\ unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp hence "fv (t \ \2) \ subst_domain \2 = {}" using subst_dom_elim unfolding range_vars_alt_def by simp moreover have "v \ range_vars \2" using Insert.prems(3) dom1v_unfold by blast hence "v \ fv t \ range_vars \2" using Insert.hyps(4) by blast hence "v \ fv (t \ \2)" using \fv (t \ \2) \ fv t \ range_vars \2\ by blast moreover have "fv (t \ \2) \ subst_domain \1 = {}" using dom_img_disj fv_t2_bound \fv t \ insert v (subst_domain \1) = {}\ by blast ultimately have 3: "fv (t \ \2) \ insert v (subst_domain (\1 \\<^sub>s \2)) = {}" using dom_comp_is_union by blast have "\1(v := t) \\<^sub>s \2 = (\1 \\<^sub>s \2)(v := t \ \2)" using subst_comp_upd1[of \1 v t \2] . moreover have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ((\1 \\<^sub>s \2)(v := t \ \2))" using "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert"[OF _ 1 2 3] \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)\ wf_subst_properties by metis ultimately show ?case by presburger qed lemma wf_subst_append: fixes \1 \2::"('f,'v) subst" assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" and "range_vars \1 \ subst_domain \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\v. if \1 v = Var v then \2 v else \1 v)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) let ?if = "\w. if \1 w = Var w then \2 w else \1 w" let ?if_upd = "\w. if (\1(v := t)) w = Var w then \2 w else (\1(v := t)) w" from Insert.hyps(4) have "?if_upd = ?if(v := t)" by fastforce have dom_insert: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using Insert.hyps(4) by (auto simp add: subst_domain_def) have "\1 v = Var v" "t \ Var v" using Insert.hyps(2,4) by auto hence img_insert: "range_vars (\1(v := t)) = range_vars \1 \ fv t" using subst_img_update by metis from Insert.prems(2) dom_insert have "subst_domain \1 \ subst_domain \2 = {}" by (auto simp add: subst_domain_def) moreover have "subst_domain \1 \ range_vars \2 = {}" using Insert.prems(3) dom_insert by (simp add: subst_domain_def) moreover have "range_vars \1 \ subst_domain \2 = {}" using Insert.prems(4) img_insert by blast ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if" using Insert.IH[OF Insert.prems(1)] by metis have dom_union: "subst_domain ?if = subst_domain \1 \ subst_domain \2" by (auto simp add: subst_domain_def) hence "v \ subst_domain ?if" using Insert.hyps(2) Insert.prems(2) dom_insert by (auto simp add: subst_domain_def) moreover have "v \ range_vars ?if" using Insert.prems(3) Insert.hyps(3) dom_insert unfolding range_vars_alt_def by (auto simp add: subst_domain_def) moreover have "fv t \ insert v (subst_domain ?if) = {}" using Insert.hyps(4) Insert.prems(4) img_insert unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def) ultimately show ?case using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if\ \?if_upd = ?if(v := t)\ wf_subst_properties by (metis (no_types, lifting)) qed lemma wf_subst_elim_append: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_elim \ v" "v \ fv t" shows "subst_elim (\(w := t)) v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ v' t') hence "\q. v \ fv (Var q \ \(v' := t'))" using subst_elimD by blast hence "\q. v \ fv (Var q \ \(v' := t', w := t))" using \v \ fv t\ by simp thus ?case by (metis subst_elimI' eval_term.simps(1)) qed (simp add: subst_elim_def) lemma wf_subst_elim_dom: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\v \ subst_domain \. subst_elim \ v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ w t) have dom_insert: "subst_domain (\(w := t)) \ insert w (subst_domain \)" by (auto simp add: subst_domain_def) hence "\v \ subst_domain \. subst_elim (\(w := t)) v" using Insert.IH Insert.hyps(2,4) by (metis Insert.hyps(1) IntI disjoint_insert(2) empty_iff wf_subst_elim_append) moreover have "w \ fv t" using Insert.hyps(4) by simp hence "\q. w \ fv (Var q \ \(w := t))" by (metis fv_simps(1) fv_in_subst_img Insert.hyps(3) contra_subsetD fun_upd_def singletonD eval_term.simps(1)) hence "subst_elim (\(w := t)) w" by (metis subst_elimI') ultimately show ?case using dom_insert by blast qed simp lemma wf_subst_support_iff_mgt: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ supports \ \ \ \\<^sub>\ \" using subst_support_def subst_support_if_mgt_subst_idem wf_subst_subst_idem by blast subsection \Interpretations\ abbreviation interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ = UNIV \ ground (subst_range \)" lemma interpretation_substI: "(\v. fv (\ v) = {}) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - assume "\v. fv (\ v) = {}" moreover { fix v assume "fv (\ v) = {}" hence "v \ subst_domain \" by auto } ultimately show ?thesis by auto qed lemma interpretation_grounds[simp]: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ fv (t \ \) = {}" using subst_fv_dom_ground_if_ground_img[of t \] by blast lemma interpretation_grounds_all: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. fv (\ v) = {})" by (metis range_vars_alt_def UNIV_I fv_in_subst_img subset_empty subst_dom_vars_in_subst) lemma interpretation_grounds_all': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ ground (M \\<^sub>s\<^sub>e\<^sub>t \)" using subst_fv_dom_ground_if_ground_img[of _ \] by simp lemma interpretation_comp: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have \_fv: "fv (\ v) = {}" for v using interpretation_grounds_all[OF assms] by simp hence \_fv': "fv (t \ \) = {}" for t by (metis all_not_in_conv subst_elimD subst_elimI' eval_term.simps(1)) from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv' insert_not_empty) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v unfolding subst_compose_def using \_fv' by simp hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv insert_not_empty subst_to_var_is_var) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v unfolding subst_compose_def by (simp add: \_fv subst_apply_term_ident) hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. qed lemma interpretation_subst_exists: "\\::('f,'v) subst. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain c::"'f" where "c \ UNIV" by simp then obtain \::"('f,'v) subst" where "\v. \ v = Fun c []" by simp hence "subst_domain \ = UNIV" "ground (subst_range \)" by (simp_all add: subst_domain_def) thus ?thesis by auto qed lemma interpretation_subst_exists': "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \)" proof - obtain \::"('f,'v) subst" where \: "subst_domain \ = UNIV" "ground (subst_range \)" - using interpretation_subst_exists by moura + using interpretation_subst_exists by atomize_elim auto let ?\ = "rm_vars (UNIV - X) \" have 1: "subst_domain ?\ = X" using \ by (auto simp add: subst_domain_def) hence 2: "ground (subst_range ?\)" using \ by force show ?thesis using 1 2 by blast qed lemma interpretation_subst_idem: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" unfolding subst_idem_def using interpretation_grounds_all[of \] subst_apply_term_ident subst_eq_if_eq_vars by (fastforce simp: subst_compose) lemma subst_idem_comp_upd_eq: assumes "v \ subst_domain \" "subst_idem \" shows "\ \\<^sub>s \ = \(v := \ v) \\<^sub>s \" proof - from assms(1) have "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by auto moreover have "\w. w \ v \ (\ \\<^sub>s \) w = (\(v := \ v) \\<^sub>s \) w" unfolding subst_compose_def by auto moreover have "(\(v := \ v) \\<^sub>s \) v = \ v" using assms(2) unfolding subst_idem_def subst_compose_def by (metis fun_upd_same) ultimately show ?thesis by (metis fun_upd_same fun_upd_triv subst_comp_upd1) qed lemma interpretation_dom_img_disjoint: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {}" unfolding range_vars_alt_def by auto subsection \Basic Properties of MGUs\ lemma MGU_is_mgu_singleton: "MGU \ t u = is_mgu \ {(t,u)}" unfolding is_mgu_def unifiers_def by auto lemma Unifier_in_unifiers_singleton: "Unifier \ s t \ \ \ unifiers {(s,t)}" unfolding unifiers_def by auto lemma subst_list_singleton_fv_subset: "(\x \ set (subst_list (subst v t) E). fv (fst x) \ fv (snd x)) \ fv t \ (\x \ set E. fv (fst x) \ fv (snd x))" proof (induction E) case (Cons x E) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?fvx = "fv (fst x) \ fv (snd x)" let ?fvxsubst = "fv (fst x \ Var(v := t)) \ fv (snd x \ Var(v := t))" have "?fvs (subst_list (subst v t) (x#E)) = ?fvxsubst \ ?fvs (subst_list (subst v t) E)" unfolding subst_list_def subst_def by auto hence "?fvs (subst_list (subst v t) (x#E)) \ ?fvxsubst \ fv t \ ?fvs E" using Cons.IH by blast moreover have "?fvs (x#E) = ?fvx \ ?fvs E" by auto moreover have "?fvxsubst \ ?fvx \ fv t" using subst_fv_bound_singleton[of _ v t] by blast ultimately show ?case unfolding range_vars_alt_def by auto qed (simp add: subst_list_def) lemma subst_of_dom_subset: "subst_domain (subst_of L) \ set (map fst L)" proof (induction L rule: List.rev_induct) case (snoc x L) then obtain v t where x: "x = (v,t)" by (metis surj_pair) hence "subst_of (L@[x]) = Var(v := t) \\<^sub>s subst_of L" unfolding subst_of_def subst_def by (induct L) (auto simp: subst_compose) hence "subst_domain (subst_of (L@[x])) \ insert v (subst_domain (subst_of L))" using x subst_domain_compose[of "Var(v := t)" "subst_of L"] by (auto simp add: subst_domain_def) thus ?case using snoc.IH x by auto qed simp lemma wf_MGU_is_imgu_singleton: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ is_imgu \ {(s,t)}" proof - assume 1: "wf\<^sub>M\<^sub>G\<^sub>U \ s t" have 2: "subst_idem \" by (metis wf_subst_subst_idem 1 wf\<^sub>M\<^sub>G\<^sub>U_def) have 3: "\\' \ unifiers {(s,t)}. \ \\<^sub>\ \'" "\ \ unifiers {(s,t)}" by (metis 1 Unifier_in_unifiers_singleton wf\<^sub>M\<^sub>G\<^sub>U_def)+ have "\\ \ unifiers {(s,t)}. \ = \ \\<^sub>s \" by (metis 2 3 subst_idem_def subst_compose_assoc) thus "is_imgu \ {(s,t)}" by (metis is_imgu_def \\ \ unifiers {(s,t)}\) qed lemmas mgu_subst_range_vars = mgu_range_vars lemmas mgu_same_empty = mgu_same lemma mgu_var: assumes "x \ fv t" shows "mgu (Var x) t = Some (Var(x := t))" proof - have "unify [(Var x,t)] [] = Some [(x,t)]" using assms by (auto simp add: subst_list_def) moreover have "subst_of [(x,t)] = Var(x := t)" unfolding subst_of_def subst_def by simp ultimately show ?thesis by (simp add: mgu_def) qed lemma mgu_gives_wellformed_subst: assumes "mgu s t = Some \" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_finite_subst_domain[OF assms] mgu_subst_domain_range_vars_disjoint[OF assms] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma mgu_gives_wellformed_MGU: assumes "mgu s t = Some \" shows "wf\<^sub>M\<^sub>G\<^sub>U \ s t" using mgu_subst_domain[OF assms] mgu_sound[OF assms] mgu_subst_range_vars [OF assms] MGU_is_mgu_singleton[of s \ t] is_imgu_imp_is_mgu[of \ "{(s,t)}"] mgu_gives_wellformed_subst[OF assms] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_gives_subst_idem: "mgu s t = Some \ \ subst_idem \" using mgu_sound[of s t \] unfolding is_imgu_def subst_idem_def by auto lemma mgu_always_unifies: "Unifier \ M N \ \\. mgu M N = Some \" using mgu_complete Unifier_in_unifiers_singleton by blast lemma mgu_gives_MGU: "mgu s t = Some \ \ MGU \ s t" using mgu_sound[of s t \, THEN is_imgu_imp_is_mgu] MGU_is_mgu_singleton by metis lemma mgu_vars_bounded[dest?]: "mgu M N = Some \ \ subst_domain \ \ range_vars \ \ fv M \ fv N" using mgu_gives_wellformed_MGU unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_vars_bounded': assumes \: "mgu M N = Some \" and MN: "fv M = {} \ fv N = {}" shows "subst_domain \ = fv M \ fv N" (is ?A) and "range_vars \ = {}" (is ?B) proof - let ?C = "\t. subst_domain \ = fv t" have 0: "fv N = {} \ subst_domain \ \ fv M" "fv N = {} \ range_vars \ \ fv M" "fv M = {} \ subst_domain \ \ fv N" "fv M = {} \ range_vars \ \ fv N" using mgu_vars_bounded[OF \] by simp_all note 1 = mgu_gives_MGU[OF \] mgu_subst_domain_range_vars_disjoint[OF \] note 2 = subst_fv_imgI[of \] subst_dom_vars_in_subst[of _ \] note 3 = ground_term_subst_domain_fv_subset[of _ \] note 4 = subst_apply_fv_empty[of _ \] have "fv (\ x) = {}" when x: "x \ fv M" and N: "fv N = {}" for x using x N 0(1,2) 1 2[of x] 3[of M] 4[of N] by auto hence "?C M" ?B when N: "fv N = {}" using 0(1,2)[OF N] N by (fastforce, fastforce) moreover have "fv (\ x) = {}" when x: "x \ fv N" and M: "fv M = {}" for x using x M 0(3,4) 1 2[of x] 3[of N] 4[of M] by auto hence "?C N" ?B when M: "fv M = {}" using 0(3,4)[OF M] M by (fastforce, fastforce) ultimately show ?A ?B using MN by auto qed lemma mgu_eliminates[dest?]: assumes "mgu M N = Some \" shows "(\v \ fv M \ fv N. subst_elim \ v) \ \ = Var" (is "?P M N \") proof (cases "\ = Var") case False then obtain v where v: "v \ subst_domain \" by auto hence "v \ fv M \ fv N" using mgu_vars_bounded[OF assms] by blast thus ?thesis using wf_subst_elim_dom[OF mgu_gives_wellformed_subst[OF assms]] v by blast qed simp lemma mgu_eliminates_dom: assumes "mgu x y = Some \" "v \ subst_domain \" shows "subst_elim \ v" using mgu_gives_wellformed_subst[OF assms(1)] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_elim_def by (metis disjoint_iff_not_equal subst_dom_elim assms(2)) lemma unify_list_distinct: assumes "Unification.unify E B = Some U" "distinct (map fst B)" and "(\x \ set E. fv (fst x) \ fv (snd x)) \ set (map fst B) = {}" shows "distinct (map fst U)" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case 1 thus ?case by simp next case (2 f X g Y E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "\t t'. (t,t') \ set E' \ fv t \ fv (Fun f X) \ fv t' \ fv (Fun g Y)" by (metis zip_arg_subterm subtermeq_vars_subset) hence "?fvs E' \ fv (Fun f X) \ fv (Fun g Y)" by fastforce moreover have "fv (Fun f X) \ set (map fst B) = {}" "fv (Fun g Y) \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs E' \ set (map fst B) = {}" by blast moreover have "?fvs E \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs (E'@E) \ set (map fst B) = {}" by auto thus ?case using "2.IH"[OF * ** "2.prems"(2)] by metis next case (3 v t E B) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v t) E" from "3.prems"(3) have "v \ set (map fst B)" "fv t \ set (map fst B) = {}" by force+ hence *: "distinct (map fst ((v, t)#B))" using "3.prems"(2) by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "?fvs ?E' \ set (map fst ((v, t)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def by (simp add: \v \ fv t\ subst_remove_var) moreover have "?fvs ?E' \ fv t \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using "3.prems"(3) by auto ultimately show ?thesis by auto qed ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ *] by metis qed next case (4 f X v E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v (Fun f X)) E" have *: "?fvs E \ set (map fst B) = {}" using "4.prems"(3) by auto from "4.prems"(1) have "v \ fv (Fun f X)" by force from "4.prems"(3) have **: "v \ set (map fst B)" "fv (Fun f X) \ set (map fst B) = {}" by force+ hence ***: "distinct (map fst ((v, Fun f X)#B))" using "4.prems"(2) by auto from "4.prems"(3) have ****: "?fvs ?E' \ set (map fst ((v, Fun f X)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def using \v \ fv (Fun f X)\ subst_remove_var[of v "Fun f X"] by simp moreover have "?fvs ?E' \ fv (Fun f X) \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using * ** by blast ultimately show ?thesis by auto qed have "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, Fun f X) # B) = Some U" using \v \ fv (Fun f X)\ "4.prems"(1) by auto thus ?case using "4.IH"[OF \v \ fv (Fun f X)\ _ *** ****] by metis qed lemma mgu_None_is_subst_neq: fixes s t::"('a,'b) term" and \::"('a,'b) subst" assumes "mgu s t = None" shows "s \ \ \ t \ \" using assms mgu_always_unifies by force lemma mgu_None_if_neq_ground: assumes "t \ t'" "fv t = {}" "fv t' = {}" shows "mgu t t' = None" proof (rule ccontr) assume "mgu t t' \ None" then obtain \ where \: "mgu t t' = Some \" by auto hence "t \ \ = t" "t' \ \ = t'" using assms subst_ground_ident by auto thus False using assms(1) MGU_is_Unifier[OF mgu_gives_MGU[OF \]] by auto qed lemma mgu_None_commutes: "mgu s t = None \ mgu t s = None" thm mgu_complete[of s t] Unifier_in_unifiers_singleton[of ] using mgu_complete[of s t] Unifier_in_unifiers_singleton[of s _ t] Unifier_sym[of t _ s] Unifier_in_unifiers_singleton[of t _ s] mgu_sound[of t s] unfolding is_imgu_def by fastforce lemma mgu_img_subterm_subst: fixes \::"('f,'v) subst" and s t u::"('f,'v) term" assumes "mgu s t = Some \" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) - range Var" shows "u \ ((subterms s \ subterms t) - range Var) \\<^sub>s\<^sub>e\<^sub>t \" proof - define subterms_tuples::"('f,'v) equation list \ ('f,'v) terms" where subtt_def: "subterms_tuples \ \E. subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E) \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" define subterms_img::"('f,'v) subst \ ('f,'v) terms" where subti_def: "subterms_img \ \d. subterms\<^sub>s\<^sub>e\<^sub>t (subst_range d)" define d where "d \ \v t. subst v t::('f,'v) subst" define V where "V \ range Var::('f,'v) terms" define R where "R \ \d::('f,'v) subst. ((subterms s \ subterms t) - V) \\<^sub>s\<^sub>e\<^sub>t d" define M where "M \ \E d. subterms_tuples E \ subterms_img d" define Q where "Q \ (\E d. M E d - V \ R d - V)" define Q' where "Q' \ (\E d d'. (M E d - V) \\<^sub>s\<^sub>e\<^sub>t d' \ (R d - V) \\<^sub>s\<^sub>e\<^sub>t (d'::('f,'v) subst))" have Q_subst: "Q (subst_list (subst v t') E) (subst_of ((v, t')#B))" when v_fv: "v \ fv t'" and Q_assm: "Q ((Var v, t')#E) (subst_of B)" for v t' E B proof - define E' where "E' \ subst_list (subst v t') E" define B' where "B' \ subst_of ((v, t')#B)" have E': "E' = subst_list (d v t') E" and B': "B' = subst_of B \\<^sub>s d v t'" using subst_of_simps(3)[of "(v, t')"] unfolding subst_def E'_def B'_def d_def by simp_all have vt_img_subt: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (d v t')) = subterms t'" and vt_dom: "subst_domain (d v t') = {v}" using v_fv by (auto simp add: subst_domain_def d_def subst_def) have *: "subterms u1 \ subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E)" "subterms u2 \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" when "(u1,u2) \ set E" for u1 u2 using that by auto have **: "subterms\<^sub>s\<^sub>e\<^sub>t (d v t' ` (fv u \ subst_domain (d v t'))) \ subterms t'" for u::"('f,'v) term" using vt_dom unfolding d_def by force have 1: "subterms_tuples E' - V \ (subterms t' - V) \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t')" (is "?A \ ?B") proof fix u assume "u \ ?A" then obtain u1 u2 where u12: "(u1,u2) \ set E" "u \ (subterms (u1 \ (d v t')) - V) \ (subterms (u2 \ (d v t')) - V)" - unfolding subtt_def subst_list_def E'_def d_def by moura + unfolding subtt_def subst_list_def E'_def d_def by atomize_elim force hence "u \ (subterms t' - V) \ (((subterms_tuples E) \\<^sub>s\<^sub>e\<^sub>t d v t') - V)" using subterms_subst[of u1 "d v t'"] subterms_subst[of u2 "d v t'"] *[OF u12(1)] **[of u1] **[of u2] unfolding subtt_def subst_list_def by auto moreover have "(subterms_tuples E \\<^sub>s\<^sub>e\<^sub>t d v t') - V \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t') \ {t'}" unfolding subst_def subtt_def V_def d_def by force ultimately show "u \ ?B" using u12 v_fv by auto qed have 2: "subterms_img B' - V \ (subterms t' - V) \ (subterms_img (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t')" using B' vt_img_subt subst_img_comp_subset'''[of "subst_of B" "d v t'"] unfolding subti_def subst_def V_def by argo have 3: "subterms_tuples ((Var v, t')#E) - V = (subterms t' - V) \ (subterms_tuples E - V)" by (auto simp add: subst_def subtt_def V_def) have "fv\<^sub>s\<^sub>e\<^sub>t (subterms t' - V) \ subst_domain (d v t') = {}" using v_fv vt_dom fv_subterms[of t'] by fastforce hence 4: "subterms t' - V \\<^sub>s\<^sub>e\<^sub>t d v t' = subterms t' - V" using set_subst_ident[of "subterms t' - range Var" "d v t'"] by (simp add: V_def) have "M E' B' - V \ M ((Var v, t')#E) (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t'" using 1 2 3 4 unfolding M_def by blast moreover have "Q' ((Var v, t')#E) (subst_of B) (d v t')" using Q_assm unfolding Q_def Q'_def by auto moreover have "R (subst_of B) \\<^sub>s\<^sub>e\<^sub>t d v t' = R (subst_of ((v,t')#B))" unfolding R_def d_def by auto ultimately have "M (subst_list (d v t') E) (subst_of ((v, t')#B)) - V \ R (subst_of ((v, t')#B)) - V" unfolding Q'_def E'_def B'_def d_def by blast thus ?thesis unfolding Q_def M_def R_def d_def by blast qed have "u \ subterms s \ subterms t - V \\<^sub>s\<^sub>e\<^sub>t subst_of U" when assms': "unify E B = Some U" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (subst_of U)) - V" "Q E (subst_of B)" for E B U and T::"('f,'v) term list" using assms' proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B) thus ?case by (auto simp add: Q_def M_def R_def subti_def) next case (2 g X h Y E B U) from "2.prems"(1) obtain E' where E': "decompose (Fun g X) (Fun h Y) = Some E'" "g = h" "length X = length Y" "E' = zip X Y" "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) moreover have "subterms_tuples (E'@E) \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof fix u assume "u \ subterms_tuples (E'@E)" then obtain u1 u2 where u12: "(u1,u2) \ set (E'@E)" "u \ subterms u1 \ subterms u2" unfolding subtt_def by fastforce thus "u \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof (cases "(u1,u2) \ set E'") case True hence "subterms u1 \ subterms (Fun g X)" "subterms u2 \ subterms (Fun h Y)" using E'(4) subterms_subset params_subterms subsetCE by (metis set_zip_leftD, metis set_zip_rightD) thus ?thesis using u12 unfolding subtt_def by auto next case False thus ?thesis using u12 unfolding subtt_def by fastforce qed qed hence "Q (E'@E) (subst_of B)" using "2.prems"(3) unfolding Q_def M_def by blast ultimately show ?case using "2.IH"[of E' U] "2.prems" by meson next case (3 v t' E B) show ?case proof (cases "t' = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) unfolding Q_def M_def V_def subtt_def by auto next case False hence 1: "v \ fv t'" using "3.prems"(1) by auto hence "unify (subst_list (subst v t') E) ((v, t')#B) = Some U" using False "3.prems"(1) by auto thus ?thesis using Q_subst[OF 1 "3.prems"(3)] "3.IH"(2)[OF False 1 _ "3.prems"(2)] by metis qed next case (4 g X v E B U) have 1: "v \ fv (Fun g X)" using "4.prems"(1) not_None_eq by fastforce hence 2: "unify (subst_list (subst v (Fun g X)) E) ((v, Fun g X)#B) = Some U" using "4.prems"(1) by auto have 3: "Q ((Var v, Fun g X)#E) (subst_of B)" using "4.prems"(3) unfolding Q_def M_def subtt_def by auto show ?case using Q_subst[OF 1 3] "4.IH"[OF 1 2 "4.prems"(2)] by metis qed moreover obtain D where "unify [(s, t)] [] = Some D" "\ = subst_of D" using assms(1) by (auto simp: mgu_def split: option.splits) moreover have "Q [(s,t)] (subst_of [])" unfolding Q_def M_def R_def subtt_def subti_def by force ultimately show ?thesis using assms(2) unfolding V_def by auto qed lemma mgu_img_consts: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "Fun c [] \ subterms s \ subterms t" proof - obtain u where "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun c []" using mgu_img_subterm_subst[OF assms(1), of "Fun c []"] assms(2) by force thus ?thesis by (cases u) auto qed lemma mgu_img_consts': fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "\ z = Fun c []" shows "Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF assms(1)] assms(2) by (metis Un_iff in_subterms_Union subst_imgI term.distinct(1)) lemma mgu_img_composed_var_term: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and f::'f and Z::"'v list" assumes "mgu s t = Some \" "Fun f (map Var Z) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "\Z'. map \ Z' = map Var Z \ Fun f (map Var Z') \ subterms s \ subterms t" proof - obtain u where u: "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun f (map Var Z)" using mgu_img_subterm_subst[OF assms(1), of "Fun f (map Var Z)"] assms(2) by fastforce then obtain T where T: "u = Fun f T" "map (\t. t \ \) T = map Var Z" by (cases u) auto have "\t \ set T. \x. t = Var x" using T(2) by (induct T arbitrary: Z) auto then obtain Z' where Z': "map Var Z' = T" by (metis ex_map_conv) hence "map \ Z' = map Var Z" using T(2) by (induct Z' arbitrary: T Z) auto thus ?thesis using u(1) T(1) Z' by auto qed lemma mgu_ground_instance_case: assumes t: "fv (t \ \) = {}" shows "mgu t (t \ \) = Some (rm_vars (UNIV - fv t) \)" (is ?A) and mgu_ground_commutes: "mgu t (t \ \) = mgu (t \ \) t" (is ?B) proof - define \ where "\ \ rm_vars (UNIV - fv t) \" have \: "t \ \ = t \ \" using rm_vars_subst_eq'[of t \] unfolding \_def by metis have 0: "Unifier \ t (t \ \)" using subst_ground_ident[OF t, of \] term_subst_eq[of t \ \] unfolding \_def by (metis Diff_iff) obtain \ where \: "mgu t (t \ \) = Some \" "MGU \ t (t \ \)" using mgu_always_unifies[OF 0] mgu_gives_MGU[of t "t \ \"] unfolding \ by blast have 1: "subst_domain \ = fv t" using t MGU_is_Unifier[OF \(2)] subset_antisym[OF mgu_subst_domain[OF \(1)]] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \ by auto have 2: "subst_domain \ = fv t" using 0 rm_vars_dom[of "UNIV - fv t" \] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \_def by auto have "\ x = \ x" for x using 1 2 MGU_is_Unifier[OF \(2)] term_subst_eq_conv[of t \ \] subst_ground_ident[OF t[unfolded \], of \] subst_domI[of _ x] by metis hence "\ = \" by presburger thus A: ?A using \(1) unfolding \ \_def by blast have "Unifier \ (t \ \) t" using 0 by simp then obtain \' where \': "mgu (t \ \) t = Some \'" "MGU \' (t \ \) t" using mgu_always_unifies mgu_gives_MGU[of "t \ \" t] unfolding \ by fastforce have 3: "subst_domain \' = fv t" using t MGU_is_Unifier[OF \'(2)] subset_antisym[OF mgu_subst_domain[OF \'(1)]] ground_term_subst_domain_fv_subset[of t \'] subst_apply_fv_empty[OF t, of \'] unfolding \ by auto have "\' x = \ x" for x using 2 3 MGU_is_Unifier[OF \'(2)] term_subst_eq_conv[of t \' \] subst_ground_ident[OF t[unfolded \], of \'] subst_domI[of _ x] by metis hence "\' = \" by presburger thus ?B using A \'(1) unfolding \ \_def by argo qed subsection \Lemmata: The "Inequality Lemmata"\ text \Subterm injectivity (a stronger injectivity property)\ definition subterm_inj_on where "subterm_inj_on f A \ \x\A. \y\A. (\v. v \ f x \ v \ f y) \ x = y" lemma subterm_inj_on_imp_inj_on: "subterm_inj_on f A \ inj_on f A" unfolding subterm_inj_on_def inj_on_def by fastforce lemma subst_inj_on_is_bij_betw: "inj_on \ (subst_domain \) = bij_betw \ (subst_domain \) (subst_range \)" unfolding inj_on_def bij_betw_def by auto lemma subterm_inj_on_alt_def: "subterm_inj_on f A \ (inj_on f A \ (\s \ f`A. \u \ f`A. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") unfolding subterm_inj_on_def inj_on_def by fastforce lemma subterm_inj_on_alt_def': "subterm_inj_on \ (subst_domain \) \ (inj_on \ (subst_domain \) \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") by (metis subterm_inj_on_alt_def subst_range.simps) lemma subterm_inj_on_subset: assumes "subterm_inj_on f A" and "B \ A" shows "subterm_inj_on f B" proof - have "inj_on f A" "\s\f ` A. \u\f ` A. (\v. v \ s \ v \ u) \ s = u" using subterm_inj_on_alt_def[of f A] assms(1) by auto moreover have "f ` B \ f ` A" using assms(2) by auto ultimately have "inj_on f B" "\s\f ` B. \u\f ` B. (\v. v \ s \ v \ u) \ s = u" using inj_on_subset[of f A] assms(2) by blast+ thus ?thesis by (metis subterm_inj_on_alt_def) qed lemma inj_subst_unif_consts: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "\x \ (fv s \ fv t) - X. \c. \ x = Fun c []" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "ground (subst_range \)" "subst_domain \ \ X = {}" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" - using mgu_always_unifies by moura + using mgu_always_unifies by atomize_elim auto have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). \ z = Fun c [] \ Fun c [] \ s \ Fun c [] \ t" by (rule mgu_img_consts'[OF \]) have 4: "subst_domain \ \ range_vars \ = {}" by (metis mgu_gives_wellformed_subst[OF \] wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" by (metis mgu_gives_wellformed_MGU[OF \] wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \(4) ident_comp_subst_trm_if_disj[of \ \] unfolding range_vars_alt_def by fast } - then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura + then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by metis have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto - then obtain c where c: "\ x = Fun c []" using \(2,5) 5 by moura + then obtain c where c: "\ x = Fun c []" using \(2,5) 5 by atomize_elim auto hence *: "(\ \\<^sub>s \) x = Fun c []" using \ x by fastforce hence **: "x \ subst_domain (\ \\<^sub>s \)" "Fun c [] \ subst_range (\ \\<^sub>s \)" by (auto simp add: subst_domain_def) have "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" by (rule subst_img_comp_subset_const'[OF *]) moreover have "\ x \ Fun c []" proof (rule ccontr) assume "\\ x \ Fun c []" hence "Fun c [] \ s \ Fun c [] \ t" using 3 by metis moreover have "\u \ subst_range \. u \ subterms s \ subterms t" using \(3) by force hence "Fun c [] \ subterms s \ subterms t" by (metis c \ground (subst_range \)\x(2) ground_subst_dom_iff_img) ultimately show False by auto qed moreover have "\x' \ subst_domain \. \ x \ Var x'" proof (rule ccontr) assume "\(\x' \ subst_domain \. \ x \ Var x')" - then obtain x' where x': "x' \ subst_domain \" "\ x = Var x'" by moura + then obtain x' where x': "x' \ subst_domain \" "\ x = Var x'" by atomize_elim auto hence "\ x' = Fun c []" "(\ \\<^sub>s \) x = Fun c []" using * unfolding subst_compose_def by auto moreover have "x \ x'" using x(1) x'(2) 4 by (auto simp add: subst_domain_def) moreover have "x' \ subst_domain \" using x'(2) mgu_eliminates_dom[OF \] by (metis (no_types) subst_elim_def eval_term.simps(1) vars_iff_subterm_or_eq) moreover have "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x' = \ x'" using \ x(2) x'(1) by auto ultimately show False using subterm_inj_on_imp_inj_on[OF \(1)] * by (simp add: inj_on_def subst_compose_def x'(2) subst_domain_def) qed ultimately show "\y \ ?ys. \ x = Var y" by (metis 5 x(2) subtermeqI' vars_iff_subtermeq DiffI Un_iff subst_fv_imgI sup.orderE) qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp - then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura + then obtain y where y: "\ z = Var y" "y \ ?ys" using * by atomize_elim auto hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def by (auto simp add: subst_domain_def) hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by auto qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed lemma inj_subst_unif_comp_terms: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X)" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "ground (subst_range \)" using \(2) by auto have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" - using mgu_always_unifies by moura + using mgu_always_unifies by atomize_elim auto have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). Fun c [] \ \ z \ Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF \] by force have 4: "subst_domain \ \ range_vars \ = {}" using mgu_gives_wellformed_subst[OF \] by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" using mgu_gives_wellformed_MGU[OF \] by (metis wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \ground (subst_range \)\ ident_comp_subst_trm_if_disj[of \ \ x] unfolding range_vars_alt_def by blast } - then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura + then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by metis have ***: "\x. x \ subst_domain \ \ subst_domain \ \ fv (\ x) \ ?ys" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto moreover have "\(\x' \ ?xs. x' \ fv (\ x))" proof (rule ccontr) assume "\\(\x' \ ?xs. x' \ fv (\ x))" then obtain x' where x': "x' \ fv (\ x)" "x' \ ?xs" by metis have "x \ x'" "x' \ subst_domain \" "\ x' = Var x'" using 4 x(1) x'(1) unfolding range_vars_alt_def by auto hence "(\ \\<^sub>s \) x' \ (\ \\<^sub>s \) x" "\ x' = (\ \\<^sub>s \) x'" using \ x(2) x'(2) by (metis subst_compose subst_mono vars_iff_subtermeq x'(1), metis eval_term.simps(1) subst_compose_def) hence "\ x' \ \ x" using \ x(2) x'(2) by auto thus False using \(1) x'(2) x(2) \x \ x'\ unfolding subterm_inj_on_def by (meson subtermeqI') qed ultimately show "fv (\ x) \ ?ys" using 5 subst_dom_vars_in_subst[of x \] subst_fv_imgI[of \ x] by blast qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof (rule ccontr) fix xi assume xi_assms: "xi \ subst_domain \ \ subst_domain \" "\(\y \ ?ys. \ xi = Var y)" hence xi_\: "xi \ subst_domain \" and \_xi_comp: "\(\y. \ xi = Var y)" using ***[of xi] 5 by auto - then obtain f T where f: "\ xi = Fun f T" by (cases "\ xi") moura + then obtain f T where f: "\ xi = Fun f T" by (cases "\ xi") auto have "\g Y'. Y' \ [] \ Fun g (map Var Y') \ \ xi \ set Y' \ ?ys" proof - have "\c. Fun c [] \ \ xi \ Fun c [] \ \ xi" using \ xi_\ by (metis const_subterm_subst subst_compose) hence 1: "\c. \(Fun c [] \ \ xi)" using 3[of _ xi] xi_\ \(3) by auto have "\(\x. \ xi = Var x)" using f by auto hence "\g S. Fun g S \ \ xi \ (\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x))" using nonvar_term_has_composed_shallow_term[of "\ xi"] by auto then obtain g S where gS: "Fun g S \ \ xi" "\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x)" - by moura + by atomize_elim auto have "\s \ set S. \x. s = Var x" using 1 term.order_trans gS by (metis (no_types, lifting) UN_I term.order_refl subsetCE subterms.simps(2) sup_ge2) then obtain S' where 2: "map Var S' = S" by (metis ex_map_conv) have "S \ []" using 1 term.order_trans[OF _ gS(1)] by fastforce hence 3: "S' \ []" "Fun g (map Var S') \ \ xi" using gS(1) 2 by auto have "set S' \ fv (Fun g (map Var S'))" by simp hence 4: "set S' \ fv (\ xi)" using 3(2) fv_subterms by force show ?thesis using ***[OF xi_assms(1)] 2 3 4 by auto qed - then obtain g Y' where g: "Y' \ []" "Fun g (map Var Y') \ \ xi" "set Y' \ ?ys" by moura + then obtain g Y' where g: "Y' \ []" "Fun g (map Var Y') \ \ xi" "set Y' \ ?ys" by atomize_elim auto then obtain X where X: "map \ X = map Var Y'" "Fun g (map Var X) \ subterms s \ subterms t" using mgu_img_composed_var_term[OF \, of g Y'] by force hence "\(u::('f,'v) term) \ set (map Var X). u \ Var ` ?ys" using \(4) tfr g(1) by fastforce then obtain j where j: "j < length X" "X ! j \ ?ys" by (metis image_iff[of _ Var "fv s \ fv t - subst_domain \"] nth_map[of _ X Var] in_set_conv_nth[of _ "map Var X"] length_map[of Var X]) define yj' where yj': "yj' \ Y' ! j" define xj where xj: "xj \ X ! j" have "xj \ fv s \ fv t" using j X(1) g(3) 5 xj yj' by (metis length_map nth_map term.simps(1) in_set_conv_nth le_supE subsetCE subst_domI) hence xj_\: "xj \ subst_domain \" using j unfolding xj by simp have len: "length X = length Y'" by (rule map_eq_imp_length_eq[OF X(1)]) have "Var yj' \ \ xi" using term.order_trans[OF _ g(2)] j(1) len unfolding yj' by auto hence "\ yj' \ \ xi" using \ xi_\ by (metis eval_term.simps(1) subst_compose_def subst_mono) moreover have \_xj_var: "Var yj' = \ xj" using X(1) len j(1) nth_map unfolding xj yj' by metis hence "\ yj' = \ xj" using \ xj_\ by (metis eval_term.simps(1) subst_compose_def) moreover have "xi \ xj" using \_xi_comp \_xj_var by auto ultimately show False using \(1) xi_\ xj_\ unfolding subterm_inj_on_def by blast qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp - then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura + then obtain y where y: "\ z = Var y" "y \ ?ys" using * by atomize_elim auto hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def subst_domain_def by auto hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by blast qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed context begin private lemma sat_ineq_subterm_inj_subst_aux: fixes \::"('f,'v) subst" assumes "Unifier \ (s \ \) (t \ \)" "ground (subst_range \)" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" shows "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ = t \ \ \ \" proof - have "\\. Unifier \ (s \ \) (t \ \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain \'::"('f,'v) subst" where *: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" using interpretation_subst_exists by metis hence "Unifier (\ \\<^sub>s \') (s \ \) (t \ \)" using assms(1) by simp thus ?thesis using * interpretation_comp by blast qed - then obtain \' where \': "Unifier \' (s \ \) (t \ \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" by moura + then obtain \' where \': "Unifier \' (s \ \) (t \ \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" by atomize_elim auto define \'' where "\'' = rm_vars (UNIV - X) \'" have *: "fv (s \ \) \ X" "fv (t \ \) \ X" using assms(2,3) subst_fv_unfold_ground_img[of \] unfolding range_vars_alt_def by (simp_all add: Diff_subset_conv Un_commute) hence **: "subst_domain \'' = X" "ground (subst_range \'')" using rm_vars_img_subset[of "UNIV - X" \'] rm_vars_dom[of "UNIV - X" \'] \'(2) unfolding \''_def by auto hence "\t. t \ \ \ \'' = t \ \'' \ \" using subst_eq_if_disjoint_vars_ground[OF _ _ assms(2)] assms(4) by blast moreover have "Unifier \'' (s \ \) (t \ \)" using Unifier_dom_restrict[OF \'(1)] \''_def * by blast ultimately show ?thesis using ** by auto qed text \ The "inequality lemma": This lemma gives sufficient syntactic conditions for finding substitutions \\\ under which terms \s\ and \t\ are not unifiable. This is useful later when establishing the typing results since we there want to find well-typed solutions to inequality constraints / "negative checks" constraints, and this lemma gives conditions for protocols under which such constraints are well-typed satisfiable if satisfiable. \ lemma sat_ineq_subterm_inj_subst: fixes \ \ \::"('f,'v) subst" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subst_domain \ \ X = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "(\x \ (fv s \ fv t) - X. \c. \ x = Fun c []) \ (\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X))" and \: "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ \ t \ \ \ \" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" "ground (subst_range \)" "subst_domain \ = subst_domain \" and \: "subst_domain \ = X" "ground (subst_range \)" shows "s \ \ \ \ \ t \ \ \ \" proof - have "\\. \Unifier \ (s \ \) (t \ \)" by (metis \(1) sat_ineq_subterm_inj_subst_aux[OF _ \(4,2,3)]) hence "\Unifier \ (s \ \) (t \ \)" using inj_subst_unif_consts[OF \(1) _ \(4,2,3) \(4,5)] inj_subst_unif_comp_terms[OF \(1,2,4,5) _ \(4,5)] tfr by metis moreover have "subst_domain \ \ subst_domain \ = {}" using \(2,3) \(1) by auto ultimately show ?thesis using \ subst_eq_if_disjoint_vars_ground[OF _ \(2) \(2)] by metis qed end lemma ineq_subterm_inj_cond_subst: assumes "X \ range_vars \ = {}" and "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t S \ T = [] \ (\u \ set T. u \ Var`X)" shows "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \) \ T = [] \ (\u \ set T. u \ Var`X)" proof (intro allI impI) let ?M = "\S. subterms\<^sub>s\<^sub>e\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" let ?N = "\S. subterms\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>e\<^sub>t S \ subst_domain \))" fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \)" hence 1: "Fun f T \ ?M S \ Fun f T \ ?N S" using subterms_subst[of _ \] by auto have 2: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ \u \ set T. u \ Var`X" using fv_subset_subterms[of "Fun f T" "subst_range \"] assms(1) unfolding range_vars_alt_def by force have 3: "\x \ subst_domain \. \ x \ Var`X" proof fix x assume "x \ subst_domain \" hence "fv (\ x) \ range_vars \" using subst_dom_vars_in_subst subst_fv_imgI unfolding range_vars_alt_def by auto thus "\ x \ Var`X" using assms(1) by auto qed show "T = [] \ (\s \ set T. s \ Var`X)" using 1 proof assume "Fun f T \ ?M S" then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t S" "u \ \ = Fun f T" by fastforce show ?thesis proof (cases u) case (Var x) hence "Fun f T \ subst_range \" using u(2) by (simp add: subst_domain_def) hence "\u \ set T. u \ Var`X" using 2 by force thus ?thesis by auto next case (Fun g S) hence "S = [] \ (\u \ set S. u \ Var`X)" using assms(2) u(1) by metis thus ?thesis proof assume "S = []" thus ?thesis using u(2) Fun by simp next assume "\u \ set S. u \ Var`X" - then obtain u' where u': "u' \ set S" "u' \ Var`X" by moura + then obtain u' where u': "u' \ set S" "u' \ Var`X" by atomize_elim auto hence "u' \ \ \ set T" using u(2) Fun by auto thus ?thesis using u'(2) 3 by (cases u') force+ qed qed next assume "Fun f T \ ?N S" thus ?thesis using 2 by force qed qed subsection \Lemmata: Sufficient Conditions for Term Matching\ definition subst_var_inv::"('a,'b) subst \ 'b set \ ('a,'b) subst" where "subst_var_inv \ X \ (\x. if Var x \ \ ` X then Var ((inv_into X \) (Var x)) else Var x)" lemma subst_var_inv_subst_domain: assumes "x \ subst_domain (subst_var_inv \ X)" shows "Var x \ \ ` X" by (meson assms subst_dom_vars_in_subst subst_var_inv_def) lemma subst_var_inv_subst_domain': assumes "X \ subst_domain \" shows "x \ subst_domain (subst_var_inv \ X) \ Var x \ \ ` X" proof show "Var x \ \ ` X \ x \ subst_domain (subst_var_inv \ X)" by (metis (no_types, lifting) assms f_inv_into_f in_mono inv_into_into subst_domI subst_dom_vars_in_subst subst_var_inv_def term.inject(1)) qed (rule subst_var_inv_subst_domain) lemma subst_var_inv_Var_range: "subst_range (subst_var_inv \ X) \ range Var" unfolding subst_var_inv_def by auto text \Injective substitutions from variables to variables are invertible\ lemma inj_var_ran_subst_is_invertible: assumes \_inj_on_X: "inj_on \ X" and \_var_on_X: "\ ` X \ range Var" and fv_t: "fv t \ X" shows "t = t \ \ \\<^sub>s subst_var_inv \ X" proof - have "\ x \ subst_var_inv \ X = Var x" when x: "x \ X" for x proof - obtain y where y: "\ x = Var y" using x \_var_on_X fv_t by auto hence "Var y \ \ ` X" using x by simp thus ?thesis using y inv_into_f_eq[OF \_inj_on_X x y] unfolding subst_var_inv_def by simp qed thus ?thesis using fv_t by (simp add: subst_compose_def trm_subst_ident'' subset_eq) qed lemma inj_var_ran_subst_is_invertible': assumes \_inj_on_t: "inj_on \ (fv t)" and \_var_on_t: "\ ` fv t \ range Var" shows "t = t \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fast text \Sufficient conditions for matching unifiable terms\ lemma inj_var_ran_unifiable_has_subst_match: assumes "t \ \ = s \ \" "inj_on \ (fv t)" "\ ` fv t \ range Var" shows "t = s \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Parallel_Compositionality.thy b/thys/Stateful_Protocol_Composition_and_Typing/Parallel_Compositionality.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Parallel_Compositionality.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Parallel_Compositionality.thy @@ -1,1090 +1,1090 @@ (* Title: Parallel_Compositionality.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, The University of Sheffield SPDX-License-Identifier: BSD-3-Clause *) section \Parallel Compositionality of Security Protocols\ theory Parallel_Compositionality imports Typing_Result Labeled_Strands begin text\\label{sec:Parallel-Compositionality}\ subsection \Definitions: Labeled Typed Model Locale\ locale labeled_typed_model = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" + fixes label_witness1 and label_witness2::"'lbl" assumes at_least_2_labels: "label_witness1 \ label_witness2" begin text \The Ground Sub-Message Patterns (GSMP)\ definition GSMP::"('fun,'var) terms \ ('fun,'var) terms" where "GSMP P \ {t \ SMP P. fv t = {}}" definition typing_cond where "typing_cond \ \ wf\<^sub>s\<^sub>t {} \ \ fv\<^sub>s\<^sub>t \ \ bvars\<^sub>s\<^sub>t \ = {} \ tfr\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t \) \ Ana_invar_subst (ik\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>s\<^sub>t \)" subsection \Definitions: GSMP Disjointness and Parallel Composability\ definition GSMP_disjoint where "GSMP_disjoint P1 P2 Secrets \ GSMP P1 \ GSMP P2 \ Secrets \ {m. {} \\<^sub>c m}" definition declassified\<^sub>l\<^sub>s\<^sub>t where "declassified\<^sub>l\<^sub>s\<^sub>t (\::('fun,'var,'lbl) labeled_strand) \ \ {s. \{set ts | ts. (\, Receive ts) \ set (\ \\<^sub>l\<^sub>s\<^sub>t \)} \ s}" definition par_comp where "par_comp (\::('fun,'var,'lbl) labeled_strand) (Secrets::('fun,'var) terms) \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Secrets) \ (\s \ Secrets. \{} \\<^sub>c s) \ ground Secrets" definition strand_leaks\<^sub>l\<^sub>s\<^sub>t where "strand_leaks\<^sub>l\<^sub>s\<^sub>t \ Sec \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \ \. \l. (\ \ \proj_unl l \@[Send1 t]\))" subsection \Definitions: GSMP-Restricted Intruder Deduction Variant\ definition intruder_deduct_hom:: "('fun,'var) terms \ ('fun,'var,'lbl) labeled_strand \ ('fun,'var) term \ bool" ("\_;_\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P _" 50) where "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t \ \M; \t. t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)\ \\<^sub>r t" lemma intruder_deduct_hom_AxiomH[simp]: assumes "t \ M" shows "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" using intruder_deduct_restricted.AxiomR[of t M] assms unfolding intruder_deduct_hom_def by blast lemma intruder_deduct_hom_ComposeH[simp]: assumes "length X = arity f" "public f" "\x. x \ set X \ \M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P x" and "Fun f X \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" shows "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P Fun f X" using intruder_deduct_restricted.ComposeR[of X f M "\t. t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)"] assms unfolding intruder_deduct_hom_def by blast lemma intruder_deduct_hom_DecomposeH: assumes "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" "Ana t = (K, T)" "\k. k \ set K \ \M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P k" "t\<^sub>i \ set T" shows "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i" using intruder_deduct_restricted.DecomposeR[of M "\t. t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" t] assms unfolding intruder_deduct_hom_def by blast lemma intruder_deduct_hom_induct[consumes 1, case_names AxiomH ComposeH DecomposeH]: assumes "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" "\t. t \ M \ P M t" "\X f. \length X = arity f; public f; \x. x \ set X \ \M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P x; \x. x \ set X \ P M x; Fun f X \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \) \ \ P M (Fun f X)" "\t K T t\<^sub>i. \\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t; P M t; Ana t = (K, T); \k. k \ set K \ \M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P k; \k. k \ set K \ P M k; t\<^sub>i \ set T\ \ P M t\<^sub>i" shows "P M t" using intruder_deduct_restricted_induct[of M "\t. t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" t "\M Q t. P M t"] assms unfolding intruder_deduct_hom_def by blast lemma ideduct_hom_mono: "\\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t; M \ M'\ \ \M'; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" using ideduct_restricted_mono[of M _ t M'] unfolding intruder_deduct_hom_def by fast subsection \Lemmata: GSMP\ lemma GSMP_disjoint_empty[simp]: "GSMP_disjoint {} A Sec" "GSMP_disjoint A {} Sec" unfolding GSMP_disjoint_def GSMP_def by fastforce+ lemma GSMP_mono: assumes "N \ M" shows "GSMP N \ GSMP M" using SMP_mono[OF assms] unfolding GSMP_def by fast lemma GSMP_SMP_mono: assumes "SMP N \ SMP M" shows "GSMP N \ GSMP M" using assms unfolding GSMP_def by fast lemma GSMP_subterm: assumes "t \ GSMP M" "t' \ t" shows "t' \ GSMP M" using SMP.Subterm[of t M t'] ground_subterm[of t t'] assms unfolding GSMP_def by auto lemma GSMP_subterms: "subterms\<^sub>s\<^sub>e\<^sub>t (GSMP M) = GSMP M" using GSMP_subterm[of _ M] by blast lemma GSMP_Ana_key: assumes "t \ GSMP M" "Ana t = (K,T)" "k \ set K" shows "k \ GSMP M" using SMP.Ana[of t M K T k] Ana_keys_fv[of t K T] assms unfolding GSMP_def by auto lemma GSMP_union: "GSMP (A \ B) = GSMP A \ GSMP B" using SMP_union[of A B] unfolding GSMP_def by auto lemma GSMP_Union: "GSMP (trms\<^sub>l\<^sub>s\<^sub>t A) = (\l. GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l A))" proof - define P where "P \ (\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" define Q where "Q \ trms\<^sub>l\<^sub>s\<^sub>t A" have "SMP (\l. P l) = (\l. SMP (P l))" "Q = (\l. P l)" unfolding P_def Q_def by (metis SMP_Union, metis trms\<^sub>l\<^sub>s\<^sub>t_union) hence "GSMP Q = (\l. GSMP (P l))" unfolding GSMP_def by auto thus ?thesis unfolding P_def Q_def by metis qed lemma in_GSMP_in_proj: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t A) \ \n. t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A)" using GSMP_Union[of A] by blast lemma in_proj_in_GSMP: "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A) \ t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t A)" using GSMP_Union[of A] by blast lemma GSMP_disjointE: assumes A: "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t n A) (trms_proj\<^sub>l\<^sub>s\<^sub>t m A) Sec" shows "GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t m A) \ Sec \ {m. {} \\<^sub>c m}" using assms unfolding GSMP_disjoint_def by auto lemma GSMP_disjoint_term: assumes "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" shows "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) \ t \ Sec \ {} \\<^sub>c t" using assms unfolding GSMP_disjoint_def by blast lemma GSMP_wt_subst_subset: assumes "t \ GSMP (M \\<^sub>s\<^sub>e\<^sub>t \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "t \ GSMP M" using SMP_wt_subst_subset[OF _ assms(2,3), of t M] assms(1) unfolding GSMP_def by simp lemma GSMP_wt_substI: assumes "t \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" shows "t \ I \ GSMP M" proof - have "t \ SMP M" using assms(1) by auto hence *: "t \ I \ SMP M" using SMP.Substitution assms(2,3) wf_trm_subst_range_iff[of I] by simp moreover have "fv (t \ I) = {}" using assms(1) interpretation_grounds_all'[OF assms(4)] by auto ultimately show ?thesis unfolding GSMP_def by simp qed lemma GSMP_disjoint_subset: assumes "GSMP_disjoint L R S" "L' \ L" "R' \ R" shows "GSMP_disjoint L' R' S" using assms(1) SMP_mono[OF assms(2)] SMP_mono[OF assms(3)] by (auto simp add: GSMP_def GSMP_disjoint_def) subsection \Lemmata: Intruder Knowledge and Declassification\ lemma declassified\<^sub>l\<^sub>s\<^sub>t_alt_def: "declassified\<^sub>l\<^sub>s\<^sub>t \ \ = {s. (\{set ts | ts. (\, Receive ts) \ set \}) \\<^sub>s\<^sub>e\<^sub>t \ \ s}" proof - have 0: "(l, receive\ts\\<^sub>s\<^sub>t) \ set (\ \\<^sub>l\<^sub>s\<^sub>t \) = (\ts'. (l, receive\ts'\\<^sub>s\<^sub>t) \ set \ \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" (is "?A \ = ?B \") for ts l proof show "?A \ \ ?B \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?A \") case False hence "(l,receive\ts\\<^sub>s\<^sub>t) = a \\<^sub>l\<^sub>s\<^sub>t\<^sub>p \" using Cons.prems by auto thus ?thesis unfolding a by (cases b) auto qed (use Cons.IH in auto) qed simp show "?B \ \ ?A \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?B \") case False hence "\ts'. a = (l, receive\ts'\\<^sub>s\<^sub>t) \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Cons.prems by auto thus ?thesis unfolding a by (cases b) auto qed (use Cons.IH in auto) qed simp qed let ?M = "\A. \{set ts |ts. (\, receive\ts\\<^sub>s\<^sub>t) \ set A}" have 1: "?M (\ \\<^sub>l\<^sub>s\<^sub>t \) = ?M \ \\<^sub>s\<^sub>e\<^sub>t \" (is "?A = ?B") proof show "?A \ ?B" proof fix t assume t: "t \ ?A" then obtain ts where ts: "t \ set ts" "(\, receive\ts\\<^sub>s\<^sub>t) \ set (\ \\<^sub>l\<^sub>s\<^sub>t \)" by blast thus "t \ ?B" using 0[of \ ts] by fastforce qed show "?B \ ?A" proof fix t assume t: "t \ ?B" then obtain ts where ts: "t \ set ts \\<^sub>s\<^sub>e\<^sub>t \" "(\, receive\ts\\<^sub>s\<^sub>t) \ set \" by blast hence "(\, receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t) \ set (\ \\<^sub>l\<^sub>s\<^sub>t \)" using 0[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by blast thus "t \ ?A" using ts(1) by force qed qed show ?thesis using 1 unfolding declassified\<^sub>l\<^sub>s\<^sub>t_def by argo qed lemma declassified\<^sub>l\<^sub>s\<^sub>t_star_receive_supset: "{t | t ts. (\, Receive ts) \ set \ \ t \ set ts} \\<^sub>s\<^sub>e\<^sub>t \ \ declassified\<^sub>l\<^sub>s\<^sub>t \ \" unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def by (fastforce intro: intruder_deduct.Axiom) lemma ik_proj_subst_GSMP_subset: assumes I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" shows "ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A)" proof fix t assume "t \ ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I" hence *: "t \ trms_proj\<^sub>l\<^sub>s\<^sub>t n A \\<^sub>s\<^sub>e\<^sub>t I" by auto then obtain s where "s \ trms_proj\<^sub>l\<^sub>s\<^sub>t n A" "t = s \ I" by auto hence "t \ SMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A)" using SMP_I I(1,2) wf_trm_subst_range_iff[of I] by simp moreover have "fv t = {}" using * interpretation_grounds_all'[OF I(3)] by auto ultimately show "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A)" unfolding GSMP_def by simp qed lemma ik_proj_subst_subterms_GSMP_subset: assumes I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n A)" (is "?A \ ?B") proof fix t assume "t \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I" then obtain s where "s \ ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I" "t \ s" by fast thus "t \ ?B" using ik_proj_subst_GSMP_subset[OF I, of n A] ground_subterm[of s t] SMP.Subterm[of s "trms\<^sub>l\<^sub>s\<^sub>t (proj n A)" t] unfolding GSMP_def by blast qed lemma declassified_proj_eq: "declassified\<^sub>l\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>t (proj n A) I" unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def proj_def by auto lemma declassified_prefix_subset: assumes AB: "prefix A B" shows "declassified\<^sub>l\<^sub>s\<^sub>t A I \ declassified\<^sub>l\<^sub>s\<^sub>t B I" proof fix t assume t: "t \ declassified\<^sub>l\<^sub>s\<^sub>t A I" obtain C where C: "B = A@C" using prefixE[OF AB] by metis show "t \ declassified\<^sub>l\<^sub>s\<^sub>t B I" using t ideduct_mono[of "\{set ts |ts. (\, receive\ts\\<^sub>s\<^sub>t) \ set A} \\<^sub>s\<^sub>e\<^sub>t I" t "\{set ts |ts. (\, receive\ts\\<^sub>s\<^sub>t) \ set B} \\<^sub>s\<^sub>e\<^sub>t I"] unfolding C declassified\<^sub>l\<^sub>s\<^sub>t_alt_def by auto qed lemma declassified_proj_ik_subset: "declassified\<^sub>l\<^sub>s\<^sub>t A I \ {s. ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I \ s}" (is "?A A \ ?P A A") proof - have *: "\{set ts |ts. (\, receive\ts\\<^sub>s\<^sub>t) \ set A} \\<^sub>s\<^sub>e\<^sub>t I \ ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t I" using proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set by fastforce show ?thesis using ideduct_mono[OF _ *] unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def by blast qed lemma deduct_proj_priv_term_prefix_ex: assumes A: "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t I \ t" and t: "\{} \\<^sub>c t" shows "\B k s. (k = \ \ k = ln l) \ prefix (B@[(k,receive\s\\<^sub>s\<^sub>t)]) A \ declassified\<^sub>l\<^sub>s\<^sub>t ((B@[(k,receive\s\\<^sub>s\<^sub>t)])) I = declassified\<^sub>l\<^sub>s\<^sub>t A I \ ik\<^sub>s\<^sub>t (proj_unl l (B@[(k,receive\s\\<^sub>s\<^sub>t)])) = ik\<^sub>s\<^sub>t (proj_unl l A)" using A proof (induction A rule: List.rev_induct) case Nil have "ik\<^sub>s\<^sub>t (proj_unl l []) \\<^sub>s\<^sub>e\<^sub>t I = {}" by auto thus ?case using Nil t deducts_eq_if_empty_ik[of t] by argo next case (snoc a A) obtain k b where a: "a = (k,b)" by (metis surj_pair) let ?P = "k = \ \ k = (ln l)" let ?Q = "\ts. b = receive\ts\\<^sub>s\<^sub>t" have 0: "ik\<^sub>s\<^sub>t (proj_unl l (A@[a])) = ik\<^sub>s\<^sub>t (proj_unl l A)" when "?P \ \?Q" using that ik\<^sub>s\<^sub>t_snoc_no_receive_eq[OF that, of I "proj_unl l A"] unfolding ik\<^sub>s\<^sub>t_is_rcv_set a by (cases "k = \ \ k = (ln l)") auto have 1: "declassified\<^sub>l\<^sub>s\<^sub>t (A@[a]) I = declassified\<^sub>l\<^sub>s\<^sub>t A I" when "?P \ \?Q" using that snoc.prems unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def a by (metis (no_types, lifting) UnCI UnE empty_iff insert_iff list.set prod.inject set_append) note 2 = snoc.prems snoc.IH 0 1 show ?case proof (cases ?P) case True note T = this thus ?thesis proof (cases ?Q) case True thus ?thesis using T unfolding a by blast qed (use 2 in auto) qed (use 2 in auto) qed subsection \Lemmata: Homogeneous and Heterogeneous Terms (Deprecated Theory)\ text \The following theory is no longer needed for the compositionality result\ context begin private definition proj_specific where "proj_specific n t \ Secrets \ t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n \) - (Secrets \ {m. {} \\<^sub>c m})" private definition heterogeneous\<^sub>l\<^sub>s\<^sub>t where "heterogeneous\<^sub>l\<^sub>s\<^sub>t t \ Secrets \ ( (\l1 l2. \s1 \ subterms t. \s2 \ subterms t. l1 \ l2 \ proj_specific l1 s1 \ Secrets \ proj_specific l2 s2 \ Secrets))" private abbreviation homogeneous\<^sub>l\<^sub>s\<^sub>t where "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Secrets \ \heterogeneous\<^sub>l\<^sub>s\<^sub>t t \ Secrets" private definition intruder_deduct_hom':: "('fun,'var) terms \ ('fun,'var,'lbl) labeled_strand \ ('fun,'var) terms \ ('fun,'var) term \ bool" ("\_;_;_\ \\<^sub>h\<^sub>o\<^sub>m _" 50) where "\M; \; Sec\ \\<^sub>h\<^sub>o\<^sub>m t \ \M; \t. homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec \ t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)\ \\<^sub>r t" private lemma GSMP_disjoint_fst_specific_not_snd_specific: assumes "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" "l \ l'" and "proj_specific l m \ Sec" shows "\proj_specific l' m \ Sec" using assms by (fastforce simp add: GSMP_disjoint_def proj_specific_def) private lemma GSMP_disjoint_snd_specific_not_fst_specific: assumes "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and "proj_specific l' m \ Sec" shows "\proj_specific l m \ Sec" using assms by (auto simp add: GSMP_disjoint_def proj_specific_def) private lemma GSMP_disjoint_intersection_not_specific: assumes "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and "t \ Sec \ {} \\<^sub>c t" shows "\proj_specific l t \ Sec" "\proj_specific l t \ Sec" using assms by (auto simp add: GSMP_disjoint_def proj_specific_def) private lemma proj_specific_secrets_anti_mono: assumes "proj_specific l t \ Sec" "Sec' \ Sec" shows "proj_specific l t \ Sec'" using assms unfolding proj_specific_def by fast private lemma heterogeneous_secrets_anti_mono: assumes "heterogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" "Sec' \ Sec" shows "heterogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec'" using assms proj_specific_secrets_anti_mono unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by metis private lemma homogeneous_secrets_mono: assumes "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec'" "Sec' \ Sec" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" using assms heterogeneous_secrets_anti_mono by blast private lemma heterogeneous_supterm: assumes "heterogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" "t \ t'" shows "heterogeneous\<^sub>l\<^sub>s\<^sub>t t' \ Sec" proof - obtain l1 l2 s1 s2 where *: "l1 \ l2" "s1 \ t" "proj_specific l1 s1 \ Sec" "s2 \ t" "proj_specific l2 s2 \ Sec" - using assms(1) unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by moura + using assms(1) unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by fast thus ?thesis using term.order_trans[OF *(2) assms(2)] term.order_trans[OF *(4) assms(2)] by (auto simp add: heterogeneous\<^sub>l\<^sub>s\<^sub>t_def) qed private lemma homogeneous_subterm: assumes "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" "t' \ t" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t' \ Sec" by (metis assms heterogeneous_supterm) private lemma proj_specific_subterm: assumes "t \ t'" "proj_specific l t' \ Sec" shows "proj_specific l t \ Sec \ t \ Sec \ {} \\<^sub>c t" using GSMP_subterm[OF _ assms(1)] assms(2) by (auto simp add: proj_specific_def) private lemma heterogeneous_term_is_Fun: assumes "heterogeneous\<^sub>l\<^sub>s\<^sub>t t A S" shows "\f T. t = Fun f T" using assms by (cases t) (auto simp add: GSMP_def heterogeneous\<^sub>l\<^sub>s\<^sub>t_def proj_specific_def) private lemma proj_specific_is_homogeneous: assumes \: "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and t: "proj_specific l m \ Sec" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t m \ Sec" proof assume "heterogeneous\<^sub>l\<^sub>s\<^sub>t m \ Sec" then obtain s l' where s: "s \ subterms m" "proj_specific l' s \ Sec" "l \ l'" - unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by moura + unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by atomize_elim auto hence "s \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" "s \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)" using t by (auto simp add: GSMP_def proj_specific_def) hence "s \ Sec \ {} \\<^sub>c s" using \ s(3) by (auto simp add: GSMP_disjoint_def) thus False using s(2) by (auto simp add: proj_specific_def) qed private lemma deduct_synth_homogeneous: assumes "{} \\<^sub>c t" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" proof - have "\s \ subterms t. {} \\<^sub>c s" using deduct_synth_subterm[OF assms] by auto thus ?thesis unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def proj_specific_def by auto qed private lemma GSMP_proj_is_homogeneous: assumes "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l A) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' A) Sec" and "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" "t \ Sec" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t A Sec" proof assume "heterogeneous\<^sub>l\<^sub>s\<^sub>t t A Sec" then obtain s l' where s: "s \ subterms t" "proj_specific l' s A Sec" "l \ l'" - unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by moura + unfolding heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by atomize_elim auto hence "s \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" "s \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' A)" using assms by (auto simp add: GSMP_def proj_specific_def) hence "s \ Sec \ {} \\<^sub>c s" using assms(1) s(3) by (auto simp add: GSMP_disjoint_def) thus False using s(2) by (auto simp add: proj_specific_def) qed private lemma homogeneous_is_not_proj_specific: assumes "homogeneous\<^sub>l\<^sub>s\<^sub>t m \ Sec" shows "\l::'lbl. \proj_specific l m \ Sec" proof - let ?P = "\l s. proj_specific l s \ Sec" have "\l1 l2. \s1\subterms m. \s2\subterms m. (l1 \ l2 \ (\?P l1 s1 \ \?P l2 s2))" using assms heterogeneous\<^sub>l\<^sub>s\<^sub>t_def by metis then obtain l1 l2 where "l1 \ l2" "\?P l1 m \ \?P l2 m" by (metis term.order_refl at_least_2_labels) thus ?thesis by metis qed private lemma secrets_are_homogeneous: assumes "\s \ Sec. P s \ (\s' \ subterms s. {} \\<^sub>c s' \ s' \ Sec)" "s \ Sec" "P s" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t s \ Sec" using assms by (auto simp add: heterogeneous\<^sub>l\<^sub>s\<^sub>t_def proj_specific_def) private lemma GSMP_is_homogeneous: assumes \: "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and t: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" "t \ Sec" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" proof - - obtain n where n: "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n \)" using in_GSMP_in_proj[OF t(1)] by moura + obtain n where n: "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n \)" using in_GSMP_in_proj[OF t(1)] by atomize_elim auto show ?thesis using GSMP_proj_is_homogeneous[OF \ n t(2)] by metis qed private lemma GSMP_intersection_is_homogeneous: assumes \: "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and t: "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)" "l \ l'" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" proof - define M where "M \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" define M' where "M' \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)" have t_in: "t \ M \ M'" "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using t(1) in_proj_in_GSMP[of t _ \] unfolding M_def M'_def by blast+ have "M \ M' \ Sec \ {m. {} \\<^sub>c m}" using \ GSMP_disjointE[of l \ l' Sec] t(2) unfolding M_def M'_def by presburger moreover have "subterms\<^sub>s\<^sub>e\<^sub>t (M \ M') = M \ M'" using GSMP_subterms unfolding M_def M'_def by blast ultimately have *: "subterms\<^sub>s\<^sub>e\<^sub>t (M \ M') \ Sec \ {m. {} \\<^sub>c m}" by blast show ?thesis proof (cases "t \ Sec") case True thus ?thesis using * secrets_are_homogeneous[of Sec "\t. t \ M \ M'", OF _ _ t_in(1)] by fast qed (metis GSMP_is_homogeneous[OF \ t_in(2)]) qed private lemma GSMP_is_homogeneous': assumes \: "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and t: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" "t \ Sec - \{GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) | l1 l2. l1 \ l2}" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" using GSMP_is_homogeneous[OF \ t(1)] GSMP_intersection_is_homogeneous[OF \] t(2) by blast private lemma Ana_keys_homogeneous: assumes \: "\l l'. l \ l' \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) Sec" and t: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" and k: "Ana t = (K,T)" "k \ set K" "k \ Sec - \{GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) | l1 l2. l1 \ l2}" shows "homogeneous\<^sub>l\<^sub>s\<^sub>t k \ Sec" proof (cases "k \ \{GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) | l1 l2. l1 \ l2}") case False hence "k \ Sec" using k(3) by fast moreover have "k \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using t SMP.Ana[OF _ k(1,2)] Ana_keys_fv[OF k(1)] k(2) unfolding GSMP_def by auto ultimately show ?thesis using GSMP_is_homogeneous[OF \, of k] by metis qed (use GSMP_intersection_is_homogeneous[OF \] in blast) end subsection \Lemmata: Intruder Deduction Equivalences\ lemma deduct_if_hom_deduct: "\M;A\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P m \ M \ m" using deduct_if_restricted_deduct unfolding intruder_deduct_hom_def by blast lemma hom_deduct_if_hom_ik: assumes "\M;A\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P m" "\m \ M. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t A)" shows "m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t A)" proof - let ?Q = "\m. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t A)" have "?Q t'" when "?Q t" "t' \ t" for t t' using GSMP_subterm[OF _ that(2)] that(1) by blast thus ?thesis using assms(1) restricted_deduct_if_restricted_ik[OF _ assms(2)] unfolding intruder_deduct_hom_def by blast qed lemma deduct_hom_if_synth: assumes hom: "m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" and m: "M \\<^sub>c m" shows "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P m" proof - let ?Q = "\m. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" have "?Q t'" when "?Q t" "t' \ t" for t t' using GSMP_subterm[OF _ that(2)] that(1) by blast thus ?thesis using assms deduct_restricted_if_synth[of ?Q] unfolding intruder_deduct_hom_def by blast qed lemma hom_deduct_if_deduct: assumes M: "\m \ M. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" and m: "M \ m" "m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" shows "\M; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P m" proof - let ?P = "\x. x \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" (* have GSMP_hom: "homogeneous\<^sub>l\<^sub>s\<^sub>t t \ Sec" when "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" for t using \ GSMP_is_homogeneous[of \ Sec t] secrets_are_homogeneous[of Sec "\x. True" t \] that unfolding par_comp_def by blast *) have P_Ana: "?P k" when "?P t" "Ana t = (K, T)" "k \ set K" for t K T k using GSMP_Ana_key[OF _ that(2,3), of "trms\<^sub>l\<^sub>s\<^sub>t \"] that (* GSMP_hom *) by presburger have P_subterm: "?P t'" when "?P t" "t' \ t" for t t' using GSMP_subterm[of _ "trms\<^sub>l\<^sub>s\<^sub>t \"] (* homogeneous_subterm[of _ \ Sec] *) that by blast have P_m: "?P m" using (* GSMP_hom[OF m(2)] *) m(2) by metis show ?thesis using restricted_deduct_if_deduct'[OF M _ _ m(1) P_m] P_Ana P_subterm unfolding intruder_deduct_hom_def by fast qed subsection \Lemmata: Deduction Reduction of Parallel Composable Constraints\ lemma par_comp_hom_deduct: assumes \: "par_comp \ Sec" and M: "\l. M l \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" "\l. Discl \ {s. M l \ s}" and Sec: "\l. \s \ Sec - Discl. \(\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P s)" and t: "\\l. M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" shows "t \ Sec - Discl" (is ?A) "\l. t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ \M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" (is ?B) proof - have M': "\l. \m \ M l. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" proof (intro allI ballI) fix l m show "m \ M l \ m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using M(1) in_proj_in_GSMP[of m l \] by blast qed have Discl_hom_deduct: "\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P u" when u: "u \ Discl" "u \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" for l u proof- have "M l \ u" using M(2) u by auto thus ?thesis using hom_deduct_if_deduct[of "M l" \ u] M(1) M' u by auto qed show ?A ?B using t proof (induction t rule: intruder_deduct_hom_induct) case (AxiomH t) - then obtain lt where t_in_proj_ik: "t \ M lt" by moura + then obtain lt where t_in_proj_ik: "t \ M lt" by atomize_elim auto show t_not_Sec: "t \ Sec - Discl" proof assume "t \ Sec - Discl" hence "\l. \(\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t)" using Sec by auto thus False using intruder_deduct_hom_AxiomH[OF t_in_proj_ik] by metis qed have 1: "\l. t \ M l \ t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" using M(1,2) AxiomH by auto have 3: "{} \\<^sub>c t \ t \ Discl" when "l1 \ l2" "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \)" for l1 l2 using \ t_not_Sec that by (auto simp add: par_comp_def GSMP_disjoint_def) have 4: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using M(1) M' t_in_proj_ik by auto show "\l. t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ \M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" by (metis (lifting) Int_iff empty_subsetI 1 3 4 Discl_hom_deduct t_in_proj_ik intruder_deduct_hom_AxiomH[of t _ \] deduct_hom_if_synth[of t \ "{}"] ideduct_hom_mono[of "{}" \ t]) next case (ComposeH T f) show "\l. Fun f T \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ \M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P Fun f T" proof (intro allI impI) fix l assume "Fun f T \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" hence "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" when "t \ set T" for t using that GSMP_subterm[OF _ subtermeqI''] by auto thus "\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P Fun f T" using ComposeH.IH(2) intruder_deduct_hom_ComposeH[OF ComposeH.hyps(1,2) _ ComposeH.hyps(4)] by simp qed thus "Fun f T \ Sec - Discl" using Sec ComposeH.hyps(4) trms\<^sub>l\<^sub>s\<^sub>t_union[of \] GSMP_Union[of \] by blast next case (DecomposeH t K T t\<^sub>i) have ti_subt: "t\<^sub>i \ t" using Ana_subterm[OF DecomposeH.hyps(2)] \t\<^sub>i \ set T\ by auto have t: "t \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using DecomposeH.hyps(1) hom_deduct_if_hom_ik M(1) M' by auto have ti: "t\<^sub>i \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" using intruder_deduct_hom_DecomposeH[OF DecomposeH.hyps] hom_deduct_if_hom_ik M(1) M' by auto obtain l where l: "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" using in_GSMP_in_proj[of _ \] ti t by presburger have K_IH: "\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P k" when "k \ set K" for k using that GSMP_Ana_key[OF _ DecomposeH.hyps(2)] DecomposeH.IH(4) l by auto have ti_IH: "\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i" using K_IH DecomposeH.IH(2) l intruder_deduct_hom_DecomposeH[OF _ DecomposeH.hyps(2) _ \t\<^sub>i \ set T\] by blast thus ti_not_Sec: "t\<^sub>i \ Sec - Discl" using Sec by blast have "{} \\<^sub>c t\<^sub>i \ t\<^sub>i \ Discl" when "t\<^sub>i \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)" "l' \ l" for l' proof - have "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) Sec" using that(2) \ by (simp add: par_comp_def) thus ?thesis using ti_not_Sec GSMP_subterm[OF l ti_subt] that(1) by (auto simp add: GSMP_disjoint_def) qed hence "\M l'; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i" when "t\<^sub>i \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)" "l' \ l" for l' using that Discl_hom_deduct[OF _ ti] deduct_hom_if_synth[OF ti, THEN ideduct_hom_mono[OF _ empty_subsetI]] by (cases "t\<^sub>i \ Discl") simp_all thus "\l. t\<^sub>i \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \) \ \M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i" using ti_IH by blast qed qed lemma par_comp_deduct_proj: assumes \: "par_comp \ Sec" and M: "\l. M l \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" "\l. Discl \ {s. M l \ s}" and t: "(\l. M l) \ t" "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" shows "M l \ t \ (\s \ Sec - Discl. \l. M l \ s)" using t proof (induction t rule: intruder_deduct_induct) case (Axiom t) - then obtain l' where t_in_ik_proj: "t \ M l'" by moura + then obtain l' where t_in_ik_proj: "t \ M l'" by atomize_elim auto show ?case proof (cases "t \ Sec - Discl \ {} \\<^sub>c t") case True thus ?thesis by (cases "t \ Sec - Discl") (metis intruder_deduct.Axiom[OF t_in_ik_proj], use ideduct_mono[of "{}" t "M l"] in blast) next case False hence "t \ Sec - Discl" "\{} \\<^sub>c t" "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" using Axiom by auto hence "(\l'. l \ l' \ t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)) \ t \ Discl" using \ unfolding GSMP_disjoint_def par_comp_def by auto hence "(\l'. l \ l' \ t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l' \)) \ M l \ t \ {} \\<^sub>c t" using M by blast thus ?thesis by (cases "\s \ M l. t \ s \ {s} \ t") (blast intro: ideduct_mono[of _ t "M l"], metis (no_types, lifting) False M(1) intruder_deduct.Axiom subsetCE t_in_ik_proj) qed next case (Compose T f) hence "Fun f T \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" using Compose.prems by auto hence "t \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" when "t \ set T" for t using that unfolding GSMP_def by auto hence IH: "M l \ t \ (\s \ Sec - Discl. \l. M l \ s)" when "t \ set T" for t using that Compose.IH by auto show ?case by (cases "\t \ set T. M l \ t") (metis intruder_deduct.Compose[OF Compose.hyps(1,2)], metis IH) next case (Decompose t K T t\<^sub>i) have hom_ik: "m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" when m: "m \ M l" for m l using in_proj_in_GSMP[of m l \] M(1) m by blast have "\\l. M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i" using intruder_deduct.Decompose[OF Decompose.hyps] hom_deduct_if_deduct[of "\l. M l"] hom_ik in_proj_in_GSMP[OF Decompose.prems(1)] by blast hence "(\M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t\<^sub>i) \ (\s \ Sec-Discl. \l. \M l; \\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P s)" using par_comp_hom_deduct(2)[OF \ M] Decompose.prems(1) by blast thus ?case using deduct_if_hom_deduct[of _ \] by auto qed subsection \Theorem: Parallel Compositionality for Labeled Constraints\ lemma par_comp_prefix: assumes "par_comp (A@B) M" shows "par_comp A M" proof - let ?L = "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A \ trms_proj\<^sub>l\<^sub>s\<^sub>t l B" have "GSMP_disjoint (?L l1) (?L l2) M" when "l1 \ l2" for l1 l2 using that assms unfolding par_comp_def by (metis trms\<^sub>s\<^sub>t_append proj_append(2) unlabel_append) hence "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 A) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 A) M" when "l1 \ l2" for l1 l2 using that SMP_union by (auto simp add: GSMP_def GSMP_disjoint_def) thus ?thesis using assms unfolding par_comp_def by blast qed theorem par_comp_constr_typed: assumes \: "par_comp \ Sec" and \: "\ \ \unlabel \\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "(\l. (\ \ \proj_unl l \\)) \ (\\' l' t. prefix \' \ \ suffix [(l', receive\t\\<^sub>s\<^sub>t)] \' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \' Sec \))" proof - let ?sem = "\\. \{}; \\\<^sub>d \" let ?Q = "\\. \l. ?sem (proj_unl l \)" let ?L = "\\'. \t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \. \l. ?sem (proj_unl l \'@[send\[t]\\<^sub>s\<^sub>t])" let ?P = "\\ \' l' ts. prefix (\'@[(l',receive\ts\\<^sub>s\<^sub>t)]) \ \ ?L (\'@[(l',receive\ts\\<^sub>s\<^sub>t)])" have "\{}; unlabel \\\<^sub>d \" using \ by (simp add: constr_sem_d_def) with \ have aux: "?Q \ \ (\\'. prefix \' \ \ ?L \')" proof (induction "unlabel \" arbitrary: \ rule: List.rev_induct) case Nil hence "\ = []" using unlabel_nil_only_if_nil by simp thus ?case by auto next case (snoc b B \) hence disj: "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" when "l1 \ l2" for l1 l2 using that by (auto simp add: par_comp_def) obtain a A n where a: "\ = A@[a]" "a = (ln n, b) \ a = (\, b)" - using unlabel_snoc_inv[OF snoc.hyps(2)[symmetric]] by moura + using unlabel_snoc_inv[OF snoc.hyps(2)[symmetric]] by atomize_elim auto hence A: "\ = A@[(ln n, b)] \ \ = A@[(\, b)]" by metis have 1: "B = unlabel A" using a snoc.hyps(2) unlabel_append[of A "[a]"] by auto have 2: "par_comp A Sec" using par_comp_prefix snoc.prems(1) a by metis have 3: "\{}; unlabel A\\<^sub>d \" by (metis 1 snoc.prems(2) snoc.hyps(2) strand_sem_split(3)) have IH: "(\l. \{}; proj_unl l A\\<^sub>d \) \ (\\'. prefix \' A \ ?L \')" by (rule snoc.hyps(1)[OF 1 2 3]) show ?case proof (cases "\l. \{}; proj_unl l A\\<^sub>d \") case False then obtain \' where \': "prefix \' A" "?L \'" by (metis IH) hence "prefix \' (A@[a])" using a prefix_prefix[of _ A "[a]"] by simp thus ?thesis using \'(2) a by auto next case True note IH' = True show ?thesis proof (cases b) case (Send ts) hence "\t \ set ts. ik\<^sub>s\<^sub>t (unlabel A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using a \\{}; unlabel \\\<^sub>d \\ strand_sem_split(2)[of "{}" "unlabel A" "unlabel [a]" \] unlabel_append[of A "[a]"] by auto hence *: "\t \ set ts. (\l. (ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \)) \ t \ \" using proj_ik_union_is_unlabel_ik image_UN by metis have "ik\<^sub>s\<^sub>t (proj_unl l \) = ik\<^sub>s\<^sub>t (proj_unl l A)" for l using Send A by (metis append_Nil2 ik\<^sub>s\<^sub>t.simps(3) proj_unl_cons(3) proj_nil(2) singleton_lst_proj(1,2) proj_ik_append) hence **: "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" for l using ik_proj_subst_GSMP_subset[OF \(3,4,2), of _ \] by auto note Discl = declassified_proj_ik_subset(1)[of A \] have Sec: "ground Sec" using \ by (auto simp add: par_comp_def) (* have Sec_hom: "homogeneous\<^sub>l\<^sub>s\<^sub>t s \ Sec" when "s \ Sec" for s using that secrets_are_homogeneous[of Sec "\_. True" s \] snoc.prems(1) unfolding par_comp_def by auto *) have "\m\ik\<^sub>s\<^sub>t (proj_unl l \) \\<^sub>s\<^sub>e\<^sub>t \. m \ GSMP (trms\<^sub>l\<^sub>s\<^sub>t \)" "ik\<^sub>s\<^sub>t (proj_unl l \) \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" for l using ik_proj_subst_GSMP_subset[OF \(3,4,2), of _ \] GSMP_Union[of \] by auto moreover have "ik\<^sub>s\<^sub>t (proj_unl l [a]) = {}" for l using Send proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set[of _ "[a]"] a(2) by auto ultimately have M: "\l. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)" using a(1) proj_ik_append[of _ A "[a]"] by auto have prefix_A: "prefix A \" using A by auto have "s \ \ = s" when "s \ Sec" for s using that Sec by auto hence leakage_case: "\{}; proj_unl l A@[Send1 s]\\<^sub>d \" when "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \" "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ s" for l s using that strand_sem_append(2) IH' by auto have proj_deduct_case_n: "\m. m \ n \ \{}; proj_unl m (A@[a])\\<^sub>d \" "\t \ set ts. ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \{}; proj_unl n (A@[a])\\<^sub>d \" when "a = (ln n, Send ts)" using that IH' proj_append(2)[of _ A] by auto have proj_deduct_case_star: "\{}; proj_unl l (A@[a])\\<^sub>d \" when "a = (\, Send ts)" "\t \ set ts. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" for l using that IH' proj_append(2)[of _ A] by auto show ?thesis proof (cases "\l. \m \ ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \. m \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \") case True then obtain l s where ls: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \" "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ s" using intruder_deduct.Axiom by metis thus ?thesis using leakage_case prefix_A by blast next case False have A_decl_subset: "\l. declassified\<^sub>l\<^sub>s\<^sub>t A \ \ {s. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ s}" using Discl unfolding a(1) by auto note deduct_proj_lemma = par_comp_deduct_proj[OF snoc.prems(1) M A_decl_subset] from a(2) show ?thesis proof assume "a = (ln n, b)" hence "a = (ln n, Send ts)" "\t \ set ts. t \ \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n \)" using Send a(1) trms_proj\<^sub>l\<^sub>s\<^sub>t_append[of n A "[a]"] GSMP_wt_substI[OF _ \(3,4,2)] by (metis, force) hence "a = (ln n, Send ts)" "\m. m \ n \ \{}; proj_unl m (A@[a])\\<^sub>d \" "\t \ set ts. ik\<^sub>s\<^sub>t (proj_unl n A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \{}; proj_unl n (A@[a])\\<^sub>d \" "\t \ set ts. t \ \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t n \)" using proj_deduct_case_n by auto hence "(\l. \{}; proj_unl l \\\<^sub>d \) \ (\s \ Sec-declassified\<^sub>l\<^sub>s\<^sub>t A \. \l. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ s)" using deduct_proj_lemma * unfolding a(1) list_all_iff by metis thus ?thesis using leakage_case prefix_A by metis next assume "a = (\, b)" hence ***: "a = (\, Send ts)" "list_all (\t. t \ \ \ GSMP (trms_proj\<^sub>l\<^sub>s\<^sub>t l \)) ts" for l using Send a(1) GSMP_wt_substI[OF _ \(3,4,2)] unfolding list_all_iff by (metis, force) hence "t \ \ \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \ \ t \ \ \ declassified\<^sub>l\<^sub>s\<^sub>t A \ \ t \ \ \ {m. {} \\<^sub>c m}" when "t \ set ts" for t using that snoc.prems(1) a(1) at_least_2_labels unfolding par_comp_def GSMP_disjoint_def list_all_iff by blast hence "(\t \ set ts. t \ \ \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \) \ (\t \ set ts. t \ \ \ declassified\<^sub>l\<^sub>s\<^sub>t A \ \ t \ \ \ {m. {} \\<^sub>c m})" by blast thus ?thesis proof assume "\t \ set ts. t \ \ \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \" then obtain t where t: "t \ set ts" "t \ \ \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \" "(\l. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \) \ t \ \" using * unfolding list_all_iff by blast have "\s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t A \. \l. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ s" using t(1,2) deduct_proj_lemma[OF t(3)] ***(2) A a Discl unfolding list_all_iff by blast thus ?thesis using prefix_A leakage_case by blast next assume t: "\t \ set ts. t \ \ \ declassified\<^sub>l\<^sub>s\<^sub>t A \ \ t \ \ \ {m. {} \\<^sub>c m}" moreover { fix t l assume "t \ set ts" "t \ \ \ declassified\<^sub>l\<^sub>s\<^sub>t A \" hence "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using intruder_deduct.Axiom Discl(1)[of l] ideduct_mono[of _ "t \ \" "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \"] by blast } moreover { fix t l assume "t \ set ts" "t \ \ \ {m. {} \\<^sub>c m}" hence "ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using ideduct_mono[OF deduct_if_synth] by blast } ultimately have "\t \ set ts. ik\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" for l by blast thus ?thesis using proj_deduct_case_star[OF ***(1)] a(1) by fast qed qed qed next case (Receive t) hence "\{}; proj_unl l \\\<^sub>d \" for l using IH' a proj_append(2)[of l A "[a]"] unfolding unlabel_def proj_def by auto thus ?thesis by metis next case (Equality ac t t') hence *: "\M; [Equality ac t t']\\<^sub>d \" for M using a \\{}; unlabel \\\<^sub>d \\ unlabel_append[of A "[a]"] by auto show ?thesis using a proj_append(2)[of _ A "[a]"] Equality strand_sem_append(2)[OF _ *] IH' unfolding unlabel_def proj_def by auto next case (Inequality X F) hence *: "\M; [Inequality X F]\\<^sub>d \" for M using a \\{}; unlabel \\\<^sub>d \\ unlabel_append[of A "[a]"] by auto show ?thesis using a proj_append(2)[of _ A "[a]"] Inequality strand_sem_append(2)[OF _ *] IH' unfolding unlabel_def proj_def by auto qed qed qed from aux have "?Q \ \ (\\' l' t. ?P \ \' l' t)" proof assume "\\'. prefix \' \ \ ?L \'" then obtain \' t l where \': "prefix \' \" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \" "?sem (proj_unl l \'@[send\[t]\\<^sub>s\<^sub>t])" by blast have *: "ik\<^sub>s\<^sub>t (proj_unl l \') \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" "\{} \\<^sub>c t \ \" using \'(2) \ subst_ground_ident[of t \] strand_sem_split(4)[OF \'(3)] unfolding par_comp_def by (simp, fastforce) obtain B k s where B: "k = \ \ k = ln l" "prefix (B@[(k, receive\s\\<^sub>s\<^sub>t)]) \'" "declassified\<^sub>l\<^sub>s\<^sub>t (B@[(k, receive\s\\<^sub>s\<^sub>t)]) \ = declassified\<^sub>l\<^sub>s\<^sub>t \' \" "ik\<^sub>s\<^sub>t (proj_unl l (B@[(k, receive\s\\<^sub>s\<^sub>t)])) = ik\<^sub>s\<^sub>t (proj_unl l \')" using deduct_proj_priv_term_prefix_ex[OF *] by force have "prefix (B@[(k, receive\s\\<^sub>s\<^sub>t)]) \" using B(2) \'(1) unfolding prefix_def by force moreover have "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t (B@[(k, receive\s\\<^sub>s\<^sub>t)]) \" using B(3) \'(2) by blast moreover have "?sem (proj_unl l (B@[(k, receive\s\\<^sub>s\<^sub>t)])@[send\[t]\\<^sub>s\<^sub>t])" using *(1)[unfolded B(4)[symmetric]] prefix_proj(2)[OF B(2), of l, unfolded prefix_def] strand_sem_split(3)[OF \'(3)] strand_sem_append(2)[ of _ _ \ "[send\[t]\\<^sub>s\<^sub>t]", OF strand_sem_split(3)[of "{}" "proj_unl l (B@[(k, receive\s\\<^sub>s\<^sub>t)])"]] by force ultimately show ?thesis by blast qed simp thus ?thesis using \(1) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def suffix_def constr_sem_d_def by blast qed end locale labeled_typing = labeled_typed_model arity public Ana \ label_witness1 label_witness2 + typing_result arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and label_witness1::"'lbl" and label_witness2::"'lbl" begin theorem par_comp_constr: assumes \: "par_comp \ Sec" "typing_cond (unlabel \)" and \: "\ \ \unlabel \\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ (\\<^sub>\ \ \unlabel \\) \ ((\l. (\\<^sub>\ \ \proj_unl l \\)) \ (\\' l' t. prefix \' \ \ suffix [(l', receive\t\\<^sub>s\<^sub>t)] \' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \' Sec \\<^sub>\)))" proof - from \(2) have *: "wf\<^sub>s\<^sub>t {} (unlabel \)" "fv\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>t (unlabel \) = {}" "tfr\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (unlabel \))" "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel \) \ assignment_rhs\<^sub>s\<^sub>t (unlabel \))" unfolding typing_cond_def tfr\<^sub>s\<^sub>t_def by metis+ obtain \\<^sub>\ where \\<^sub>\: "\\<^sub>\ \ \unlabel \\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_attack_if_tfr_attack_d[OF * \(2,1)] by metis show ?thesis using par_comp_constr_typed[OF \(1) \\<^sub>\] \\<^sub>\ by auto qed end subsection \Automated GSMP Disjointness\ locale labeled_typed_model' = typed_model' arity public Ana \ + labeled_typed_model arity public Ana \ label_witness1 label_witness2 for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and label_witness1 label_witness2::'lbl begin lemma GSMP_disjointI: fixes A' A B B'::"('fun, ('fun, 'atom) term \ nat) terms" defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "\ \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t A))" assumes A'_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity A'" and B'_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity B'" and A_inst: "has_all_wt_instances_of \ A' A" and B_inst: "has_all_wt_instances_of \ B' (B \\<^sub>s\<^sub>e\<^sub>t \)" and A_SMP_repr: "finite_SMP_representation arity Ana \ A" and B_SMP_repr: "finite_SMP_representation arity Ana \ (B \\<^sub>s\<^sub>e\<^sub>t \)" and AB_trms_disj: "\t \ A. \s \ B \\<^sub>s\<^sub>e\<^sub>t \. \ t = \ s \ mgu t s \ None \ (intruder_synth' public arity {} t) \ ((\u \ Sec. is_wt_instance_of_cond \ t u))" and Sec_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s Sec" shows "GSMP_disjoint A' B' ((f Sec) - {m. {} \\<^sub>c m})" proof - have A_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" and B_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (B \\<^sub>s\<^sub>e\<^sub>t \)" and A'_wf': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A'" and B'_wf': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B'" and A_finite: "finite A" and B_finite: "finite (B \\<^sub>s\<^sub>e\<^sub>t \)" using finite_SMP_representationD[OF A_SMP_repr] finite_SMP_representationD[OF B_SMP_repr] A'_wf B'_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] list_all_iff by blast+ have AB_fv_disj: "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t (B \\<^sub>s\<^sub>e\<^sub>t \) = {}" using var_rename_fv_set_disjoint'[of A B, unfolded \_def[symmetric]] A_finite by simp have "GSMP_disjoint A (B \\<^sub>s\<^sub>e\<^sub>t \) ((f Sec) - {m. {} \\<^sub>c m})" using ground_SMP_disjointI[OF AB_fv_disj A_SMP_repr B_SMP_repr Sec_wf AB_trms_disj] unfolding GSMP_def GSMP_disjoint_def f_def by blast moreover have "SMP A' \ SMP A" "SMP B' \ SMP (B \\<^sub>s\<^sub>e\<^sub>t \)" using SMP_I'[OF A'_wf' A_wf A_inst] SMP_SMP_subset[of A' A] SMP_I'[OF B'_wf' B_wf B_inst] SMP_SMP_subset[of B' "B \\<^sub>s\<^sub>e\<^sub>t \"] by (blast, blast) ultimately show ?thesis unfolding GSMP_def GSMP_disjoint_def by auto qed end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy @@ -1,3795 +1,3795 @@ (* Title: Stateful_Compositionality.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Stateful Protocol Compositionality\ theory Stateful_Compositionality imports Stateful_Typing Parallel_Compositionality Labeled_Stateful_Strands begin text\\label{sec:Stateful-Compositionality}\ subsection \Small Lemmata\ lemma (in typed_model) wt_subst_sstp_vars_type_subset: fixes a::"('fun,'var) stateful_strand_step" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?A) and "\ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" (is ?B) and "\ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?C) proof - show ?A proof fix \ assume \: "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" - then obtain x where x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura + then obtain x where x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by atomize_elim auto show "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] False x(1) by fastforce+ obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] NegChecks *(2) by fastforce thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ - then obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura + then obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by atomize_elim auto hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed show ?B by (metis bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst) show ?C proof fix \ assume \: "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" - then obtain x where x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura + then obtain x where x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by atomize_elim auto show "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] False x(1) by (fastforce, blast) obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] NegChecks by blast thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ - then obtain y where y: "y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura + then obtain y where y: "y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by atomize_elim auto hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed qed lemma (in typed_model) wt_subst_lsst_vars_type_subset: fixes A::"('fun,'var,'a) labeled_stateful_strand" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?A) and "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?B) and "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?C) proof - have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that unlabel_Cons(1)[of l b A] unlabel_subst[of "a#A" \] subst_lsst_cons[of a A \] subst_sst_cons[of b "unlabel A" \] subst_apply_labeled_stateful_strand_step.simps(1)[of l b \] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by simp_all hence *: "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that by fast+ have "?A \ ?B \ ?C" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH wt_subst_sstp_vars_type_subset[OF assms, of b] *[OF a, of A] by (metis Un_mono) qed simp thus ?A ?B ?C by metis+ qed lemma (in stateful_typed_model) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset: assumes "d \ set D" shows "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using assms unfolding pair_def by (induct D) (auto simp add: unlabel_def) lemma (in stateful_typed_model) labeled_sat_ineq_lift: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?R1 D") and "\(j,p) \ {(i,t,s)} \ set D \ set Di. \(k,q) \ {(i,t,s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?R2 D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" using assms proof (induction D) case (Cons dl D) obtain d l where dl: "dl = (l,d)" by (metis surj_pair) have 1: "?R1 D" proof (cases "i = l") case True thus ?thesis using Cons.prems(1) dl by (cases "dl \ set Di") (auto simp add: dbproj_def) next case False thus ?thesis using Cons.prems(1) dl by (auto simp add: dbproj_def) qed have "set D \ set (dl#D)" by auto hence 2: "?R2 D" using Cons.prems(2) by blast have "i \ l \ dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems(1) dl by (auto simp add: ineq_model_def dbproj_def) moreover have "\\. Unifier \ (pair (t,s)) (pair d) \ i = l" using Cons.prems(2) dl by force ultimately have 3: "dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using strand_sem_not_unif_is_sat_ineq[of "pair (t,s)" "pair d"] dl by fastforce show ?case using Cons.IH[OF 1 2] 3 dl by auto qed simp lemma (in stateful_typed_model) labeled_sat_ineq_dbproj: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" (is "?P D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?Q D") using assms proof (induction D) case (Cons di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "?P D" using Cons.prems by (cases "di \ set Di") auto hence IH: "?Q D" by (metis Cons.IH) show ?case using di IH proof (cases "i = j \ di \ set Di") case True have 1: "\M; [\X\\\: [(pair (t,s), pair (snd di))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems True by auto have 2: "dbproj i (di#D) = di#dbproj i D" using True dbproj_Cons(1) di by auto show ?thesis using 1 2 IH by auto qed (auto simp add: dbproj_def) qed (simp add: dbproj_def) lemma (in stateful_typed_model) labeled_sat_ineq_dbproj_sem_equiv: assumes "\(j,p) \ ((\(t, s). (i, t, s)) ` set F') \ set D. \(k,q) \ ((\(t, s). (i, t, s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" shows "\M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))\\<^sub>d \ \ \M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))\\<^sub>d \" proof - let ?A = "set (map snd D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?B = "set (map snd (dbproj i D)) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?C = "set (map snd D) - set (map snd (dbproj i D))" let ?F = "(\(t, s). (i, t, s)) ` set F'" let ?P = "\\. subst_domain \ = set X \ ground (subst_range \)" have 1: "\(t, t') \ set (map snd D). (fv t \ fv t') \ set X = {}" "\(t, t') \ set (map snd (dbproj i D)). (fv t \ fv t') \ set X = {}" using assms(2) dbproj_subset[of i D] unfolding unlabel_def by force+ have 2: "?B \ ?A" unfolding dbproj_def by auto have 3: "\Unifier \ (pair f) (pair d)" when f: "f \ set F'" and d: "d \ set (map snd D) - set (map snd (dbproj i D))" for f d and \::"('fun,'var) subst" proof - obtain k where k: "(k,d) \ set D - set (dbproj i D)" using d by force have "(i,f) \ ((\(t, s). (i, t, s)) ` set F') \ set D" "(k,d) \ ((\(t, s). (i, t, s)) ` set F') \ set D" using f k by auto hence "i = k" when "Unifier \ (pair f) (pair d)" for \ using assms(1) that by blast moreover have "k \ i" using k d unfolding dbproj_def by simp ultimately show ?thesis by metis qed have "f \\<^sub>p \ \ d \\<^sub>p \" when "f \ set F'" "d \ ?C" for f d and \::"('fun,'var) subst" by (metis fun_pair_eq_subst 3[OF that]) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \)" when "f \ set F'" for f and \::"('fun,'var) subst" using that by blast moreover have "?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "?P \" for \ using assms(2) that pairs_substI[of \ "(set (map snd D) - set (map snd (dbproj i D)))"] by blast ultimately have 4: "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "f \ set F'" "?P \" for f and \::"('fun,'var) subst" by (metis that subst_pairs_compose) { fix f and \::"('fun,'var) subst" assume "f \ set F'" "?P \" hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by (metis 4) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by force } hence 5: "\f\set F'. \\. ?P \ \ f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by metis show ?thesis using negchecks_model_db_subset[OF 2] negchecks_model_db_supset[OF 2 5] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(1)] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(2)] tr_NegChecks_constr_iff(1) strand_sem_eq_defs(2) by (metis (no_types, lifting)) qed lemma (in stateful_typed_model) labeled_sat_eqs_list_all: assumes "\(j, p) \ {(i,t,s)} \ set D. \(k,q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) D\\<^sub>d \" (is "?Q D") shows "list_all (\d. fst d = i) D" using assms proof (induction D rule: List.rev_induct) case (snoc di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "pair (t,s) \ \ = pair d \ \" using di snoc.prems(2) by auto hence "\\. Unifier \ (pair (t,s)) (pair d)" by auto hence 1: "i = j" using snoc.prems(1) di by fastforce have "set D \ set (D@[di])" by auto hence 2: "?P D" using snoc.prems(1) by blast have 3: "?Q D" using snoc.prems(2) by auto show ?case using di 1 snoc.IH[OF 2 3] by simp qed simp lemma (in stateful_typed_model) labeled_sat_eqs_subseqs: assumes "Di \ set (subseqs D)" and "\(j, p) \ {(i,t,s)} \ set D. \(k, q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di\\<^sub>d \" shows "Di \ set (subseqs (dbproj i D))" proof - have "set Di \ set D" by (rule subseqs_subset[OF assms(1)]) hence "?P Di" using assms(2) by blast thus ?thesis using labeled_sat_eqs_list_all[OF _ assms(3)] subseqs_mem_dbproj[OF assms(1)] by simp qed lemma (in stateful_typing_result) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" shows "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" using assms proof (induction S) case (Cons a S) have prems: "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" using Cons.prems unlabel_Cons(2)[of a S] by simp_all hence IH: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" by (metis Cons.IH) obtain l b where a: "a = (l,b)" by (metis surj_pair) with Cons show ?case proof (cases b) case (Equality c t t') hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(3) a) thus ?thesis using a IH prems by fastforce next case (NegChecks X F G) hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(7) a) thus ?thesis using a IH prems by fastforce qed auto qed simp lemma (in stateful_typed_model) setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "setops\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)) = setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed simp subsection \Locale Setup and Definitions\ locale labeled_stateful_typed_model = stateful_typed_model arity public Ana \ Pair + labeled_typed_model arity public Ana \ label_witness1 label_witness2 for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin definition lpair where "lpair lp \ case lp of (i,p) \ (i,pair p)" lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_pair_image[simp]: "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,send\ts\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,receive\ts\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ t'\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,insert\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,delete\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\X\\\: F \\: F'\)) = ((\(t,s). (i, pair (t,s))) ` set F')" unfolding lpair_def by force+ definition par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\::('fun,'var,'lbl) labeled_stateful_strand) (Secrets::('fun,'var) terms) \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Secrets) \ (\s \ Secrets. \{} \\<^sub>c s) \ ground Secrets \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j)" definition declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ {s. \{set ts | ts. \\, receive\ts\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)} \ s}" definition strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t ("_ leaks _ under _") where "(\::('fun,'var,'lbl) labeled_stateful_strand) leaks Secrets under \ \ (\t \ Secrets - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \. \n. \ \\<^sub>s (proj_unl n \@[send\[t]\]))" type_synonym ('a,'b,'c) labeleddbstate = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) set" type_synonym ('a,'b,'c) labeleddbstatelist = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) list" definition typing_cond\<^sub>s\<^sub>s\<^sub>t where "typing_cond\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t \) \ tfr\<^sub>s\<^sub>s\<^sub>t \" text \ For proving the compositionality theorem for stateful constraints the idea is to first define a variant of the reduction technique that was used to establish the stateful typing result. This variant performs database-state projections, and it allows us to reduce the compositionality problem for stateful constraints to ordinary constraints. \ fun tr\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr\<^sub>p\<^sub>c [] D = [[]]" | "tr\<^sub>p\<^sub>c ((i,send\ts\)#A) D = map ((#) (i,send\ts\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,receive\ts\)#A) D = map ((#) (i,receive\ts\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@B) (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs (dbproj i D))))" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) (dbproj i D)) (tr\<^sub>p\<^sub>c A D))" | "tr\<^sub>p\<^sub>c ((i,\X\\\: F \\: F' \)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))) (tr\<^sub>p\<^sub>c A D)" end locale labeled_stateful_typing = labeled_stateful_typed_model arity public Ana \ Pair label_witness1 label_witness2 + stateful_typing_result arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_typing by unfold_locales end subsection \Small Lemmata\ context labeled_stateful_typed_model begin lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ = {s. \{set ts | ts. \\, receive\ts\\ \ set \} \\<^sub>s\<^sub>e\<^sub>t \ \ s}" proof - have 0: "(l, receive\ts\) \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (\ts'. (l, receive\ts'\) \ set \ \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" (is "?A \ = ?B \") for ts l proof show "?A \ \ ?B \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?A \") case False hence "(l,receive\ts\) = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using Cons.prems subst_lsst_cons[of a \ \] by auto thus ?thesis unfolding a by (cases b) auto qed (use Cons.IH in auto) qed simp show "?B \ \ ?A \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?B \") case False hence "\ts'. a = (l, receive\ts'\) \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Cons.prems by auto thus ?thesis using subst_lsst_cons[of a \ \] unfolding a by (cases b) auto qed (use Cons.IH subst_lsst_cons[of a \ \] in auto) qed simp qed let ?M = "\A. \{set ts |ts. \\, receive\ts\\ \ set A}" have 1: "?M (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = ?M \ \\<^sub>s\<^sub>e\<^sub>t \" (is "?A = ?B") proof show "?A \ ?B" proof fix t assume t: "t \ ?A" then obtain ts where ts: "t \ set ts" "\\, receive\ts\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast thus "t \ ?B" using 0[of \ ts] by fastforce qed show "?B \ ?A" proof fix t assume t: "t \ ?B" then obtain ts where ts: "t \ set ts \\<^sub>s\<^sub>e\<^sub>t \" "\\, receive\ts\\ \ set \" by blast hence "\\, receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 0[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by blast thus "t \ ?A" using ts(1) by force qed qed show ?thesis using 1 unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by argo qed lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_prefix_subset: assumes AB: "prefix A B" shows "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" proof fix t assume t: "t \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" obtain C where C: "B = A@C" using prefixE[OF AB] by metis show "t \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" using t ideduct_mono[of "\{set ts |ts. (\, receive\ts\) \ set A} \\<^sub>s\<^sub>e\<^sub>t I" t "\{set ts |ts. (\, receive\ts\) \ set B} \\<^sub>s\<^sub>e\<^sub>t I"] unfolding C declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by auto qed lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_star_receive_supset: "{t | t ts. \\, receive\ts\\ \ set \ \ t \ set ts} \\<^sub>s\<^sub>e\<^sub>t \ \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \" unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by (fastforce intro: intruder_deduct.Axiom) lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) I" using proj_mem_iff(2)[of _ A] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil: assumes "ground Sec" "\s \ Sec. \s'\subterms s. {} \\<^sub>c s' \ s' \ Sec" "\s \ Sec. \{} \\<^sub>c s" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] Sec" using assms unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" and BA: "set B \ set A" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" proof - let ?L = "\n A. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" have "?L n B \ ?L n A" for n using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] setops\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] by blast hence "GSMP_disjoint (?L m B) (?L n B) Sec" when nm: "m \ n" for n m::'lbl using GSMP_disjoint_subset[of "?L m A" "?L n A" Sec "?L m B" "?L n B"] A nm unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp thus "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using A setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono[OF BA] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp_all lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \{} \\<^sub>c s" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?M = "\l B. (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B))" let ?P = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1 B) (?M l2 B) S" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair (snd p)) (pair (snd q))) \ fst p = fst q" have "?P A" "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold by blast+ thus "?P (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "?Q (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (metis setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t, metis setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \{} \\<^sub>c s" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?N = "\l B. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B)" define M where "M \ \l (B::('fun,'var,'lbl) labeled_stateful_strand). ?N l B" let ?P = "\p q. \\. Unifier \ (pair (snd p)) (pair (snd q))" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. ?P p q \ fst p = fst q" let ?R = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?N l1 B) (?N l2 B) S" have d: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ subst_domain \ = {}" for l using \(3) unfolding proj_def bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto have "GSMP_disjoint (M l1 A) (M l2 A) S" when l: "l1 \ l2" for l1 l2 using l A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def M_def by presburger moreover have "M l (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (M l A) \\<^sub>s\<^sub>e\<^sub>t \" for l using fun_pair_subst_set[of \ "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)", symmetric] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] setops\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] proj_subst[of l A \] unfolding M_def unlabel_subst by auto ultimately have "GSMP_disjoint (M l1 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) (M l2 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) S" when l: "l1 \ l2" for l1 l2 using l GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l1 A"] GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l2 A"] unfolding GSMP_disjoint_def by fastforce thus "?R (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding M_def by blast have "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force thus "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using \(3) proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have 0: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by simp have "?Q A" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using Cons.prems 0 unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto hence IH: "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast have 1: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q using a p q pq by (cases b) auto have 2: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q proof - obtain p' X where p': "p' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst p'" "X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd p = snd p' \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] 0 by blast obtain q' Y where q': "q' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst q = fst q'" "Y \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd q = snd q' \\<^sub>p rm_vars Y \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst[OF q] 0 by blast have "pair (snd p) = pair (snd p') \ \" "pair (snd q) = pair (snd q') \ \" using fun_pair_subst[of "snd p'" "rm_vars X \"] fun_pair_subst[of "snd q'" "rm_vars Y \"] p'(3,4) q'(3,4) Cons.prems(2) rm_vars_apply'[of \ X] rm_vars_apply'[of \ Y] by fastforce+ hence "\\. Unifier \ (pair (snd p')) (pair (snd q'))" using pq Unifier_comp' by metis thus ?thesis using Cons.prems p'(1,2) q'(1,2) by simp qed show ?case by (metis 1 2 IH Un_iff setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons subst_lsst_cons) qed simp qed lemma wf_pair_negchecks_map': assumes "wf\<^sub>s\<^sub>t X (unlabel A)" shows "wf\<^sub>s\<^sub>t X (unlabel (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) M@A))" using assms by (induct M) auto lemma wf_pair_eqs_ineqs_map': fixes A::"('fun,'var,'lbl) labeled_strand" assumes "wf\<^sub>s\<^sub>t X (unlabel A)" "Di \ set (subseqs (dbproj i D))" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" shows "wf\<^sub>s\<^sub>t X (unlabel ( (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@A))" proof - let ?f = "[d\dbproj i D. d \ set Di]" define c1 where c1: "c1 = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" define c2 where c2: "c2 = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) ?f" define c3 where c3: "c3 = map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) (unlabel Di)" define c4 where c4: "c4 = map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) (unlabel ?f)" have ci_eqs: "c3 = unlabel c1" "c4 = unlabel c2" unfolding c1 c2 c3 c4 unlabel_def by auto have 1: "wf\<^sub>s\<^sub>t X (unlabel (c2@A))" using wf_fun_pair_ineqs_map[OF assms(1)] ci_eqs(2) unlabel_append[of c2 A] c4 by metis have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ X" using assms(3) subseqs_set_subset(1)[OF assms(2)] unfolding unlabel_def dbproj_def by fastforce { fix B::"('fun,'var) strand" assume "wf\<^sub>s\<^sub>t X B" hence "wf\<^sub>s\<^sub>t X (unlabel c1@B)" using 2 unfolding c1 unlabel_def by (induct Di) auto } thus ?thesis using 1 unfolding c1 c2 unlabel_def by simp qed lemma trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: defines "M \ \A. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\t \ M B. \s \ M A. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix t assume "t \ M B" then obtain b where b: "b \ set B" "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b)" unfolding M_def unfolding unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson note \' = wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] have "t \ M (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) a unfolding M_def subst_apply_labeled_stateful_strand_def unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'[OF that] \' unfolding M_def by blast moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" proof - obtain p where p: "p \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" "t = pair p" using t by blast then obtain q X where q: "q \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "p = q \\<^sub>p rm_vars (set X) \" using setops\<^sub>s\<^sub>s\<^sub>t_subst'[OF p(1)] by blast hence "t = pair q \ rm_vars (set X) \" using fun_pair_subst[of q "rm_vars (set X) \"] p(2) by presburger thus ?thesis using \'[of "set X"] q(1) unfolding M_def by blast qed ultimately show "\s \ M A. \\. t = s \ \ \ ?P \" unfolding M_def unlabel_subst by fast qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix p assume "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" then obtain b where b: "b \ set B" "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b" unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson hence p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto obtain X q where q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst q" "snd p = snd q \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] by blast show "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ ?P \" using q wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] by blast qed lemma deduct_proj_priv_term_prefix_ex_stateful: assumes A: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t I \ t" and t: "\{} \\<^sub>c t" shows "\B k s. (k = \ \ k = ln l) \ prefix (B@[(k,receive\s\)]) A \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B@[(k,receive\s\)])) I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I \ ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l (B@[(k,receive\s\)])) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" using A proof (induction A rule: List.rev_induct) case Nil have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l []) \\<^sub>s\<^sub>e\<^sub>t I = {}" by auto thus ?case using Nil t deducts_eq_if_empty_ik[of t] by argo next case (snoc a A) obtain k b where a: "a = (k,b)" by (metis surj_pair) let ?P = "k = \ \ k = (ln l)" let ?Q = "\s. b = receive\s\" have 0: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l (A@[a])) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" when "?P \ \?Q" using that ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_eq[OF that, of I "proj_unl l A"] unfolding ik\<^sub>s\<^sub>s\<^sub>t_def a by (cases "k = \ \ k = (ln l)") auto have 1: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a]) I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" when "?P \ \?Q" using that snoc.prems unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def a by (metis (no_types, lifting) UnCI UnE empty_iff insert_iff list.set prod.inject set_append) note 2 = snoc.prems snoc.IH 0 1 show ?case proof (cases ?P) case True note T = this thus ?thesis proof (cases ?Q) case True thus ?thesis using T unfolding a by blast qed (use 2 in auto) qed (use 2 in auto) qed lemma constr_sem_stateful_proj_priv_term_prefix_obtain: assumes \': "prefix \' \" "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[t]\])" and t: "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "\{} \\<^sub>c t" "t \ \\<^sub>\ = t" obtains B k' s where "k' = \ \ k' = ln n" "prefix B \'" "suffix [(k', receive\s\)] B" "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>\ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \')" "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \') \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t" using \'(2) t(3) strand_sem_append_stateful[of "{}" "{}" "proj_unl n \'" "[send\[t]\]" \\<^sub>\] by simp then obtain B k' s where B: "k' = \ \ k' = ln n" "prefix B \'" "suffix [(k', receive\s\)] B" "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>\ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \')" using deduct_proj_priv_term_prefix_ex_stateful[OF _ t(2), of \\<^sub>\ n \'] unfolding suffix_def by blast have B': "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" using B(5) \'(2) strand_sem_append_stateful[of "{}" "{}" "proj_unl n \'" "[send\[t]\]" \\<^sub>\] strand_sem_append_stateful[of "{}" "{}" "proj_unl n B" _ \\<^sub>\] prefix_proj(2)[OF B(2), of n] by (metis (no_types, lifting) append_Nil2 prefix_def strand_sem_stateful.simps(2)) have B'': "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using \' t B(1-4) declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq[of B \\<^sub>\ n] unfolding suffix_def prefix_def proj_def by auto show ?thesis by (rule that[OF B B' B'']) qed lemma constr_sem_stateful_star_proj_no_leakage: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \ s. prefix \ \ \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ \\<^sub>\ \\<^sub>s (unlabel \@[send\[s]\])" assumes Sec: "ground Sec" and \: "\(l,a) \ set \. l = \" shows "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain I B s where B: "prefix B \" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" "I \\<^sub>s (unlabel B@[send\[s]\])" unfolding no_leakage_def by blast have 1: "\(\{set ts | ts. \\, receive\ts\\ \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)} \ s)" using B(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast have 2: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" using B(2,3) Sec strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[s]\]" I] subst_apply_term_ident[of s I] unlabel_subst[of B] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel B"] by force have "l = \" when "(l,c) \ set B" for l c using that \ B(1) set_mono_prefix by blast hence "l = \" when "(l,c) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" for l c using that unfolding subst_apply_labeled_stateful_strand_def by auto hence 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) = (\{set ts | ts. \\, receive\ts\\ \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)})" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of _ "B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I"] unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto show False using 1 2 3 by force qed end subsection \Lemmata: Properties of the Constraint Translation Function\ context labeled_stateful_typed_model begin lemma tr_par_labeled_rcv_iff: "B \ set (tr\<^sub>p\<^sub>c A D) \ (i, receive\t\\<^sub>s\<^sub>t) \ set B \ (i, receive\t\) \ set A" by (induct A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) auto lemma tr_par_declassified_eq: "B \ set (tr\<^sub>p\<^sub>c A D) \ declassified\<^sub>l\<^sub>s\<^sub>t B I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" using tr_par_labeled_rcv_iff unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by simp lemma tr_par_ik_eq: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) = ik\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof - have "{t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set B} = {t. \i. (i, receive\t\) \ set A}" using tr_par_labeled_rcv_iff[OF assms] by simp moreover have "\C. {t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set C} = {t. receive\t\\<^sub>s\<^sub>t \ set (unlabel C)}" "\C. {t. \i. (i, receive\t\) \ set C} = {t. receive\t\ \ set (unlabel C)}" unfolding unlabel_def by force+ ultimately show ?thesis unfolding ik\<^sub>s\<^sub>s\<^sub>t_def ik\<^sub>s\<^sub>t_is_rcv_set by fast qed lemma tr_par_deduct_iff: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) \\<^sub>s\<^sub>e\<^sub>t I \ t \ ik\<^sub>s\<^sub>s\<^sub>t (unlabel A) \\<^sub>s\<^sub>e\<^sub>t I \ t" using tr_par_ik_eq[OF assms] by metis lemma tr_par_vars_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" (is ?P) and "bvars\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A)" (is ?Q) proof - show ?P using assms proof (induction "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) case (ConsIn A' D ac t s AA A A') then obtain i B where iB: "A = (i,\ac: t \ s\)#B" "AA = unlabel B" - unfolding unlabel_def by moura + unfolding unlabel_def by atomize_elim auto then obtain A'' d where *: "d \ set (dbproj i D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c B D)" - using ConsIn.prems(1) by moura + using ConsIn.prems(1) by atomize_elim force hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel B) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" apply (metis ConsIn.hyps(1)[OF iB(2)]) using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[OF dbproj_subset[of i D]] fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset[OF *(1)] by blast thus ?case using * iB unfolding pair_def by auto next case (ConsDel A' D t s AA A A') then obtain i B where iB: "A = (i,delete\t,s\)#B" "AA = unlabel B" - unfolding unlabel_def by moura + unfolding unlabel_def by atomize_elim auto define fltD1 where "fltD1 = (\Di. filter (\d. d \ set Di) D)" define fltD2 where "fltD2 = (\Di. filter (\d. d \ set Di) (dbproj i D))" define constr where "constr = (\Di. (map (\d. (i, \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i, \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (fltD2 Di)))" from iB obtain A'' Di where *: "Di \ set (subseqs (dbproj i D))" "A' = (constr Di)@A''" "A'' \ set (tr\<^sub>p\<^sub>c B (fltD1 Di))" - using ConsDel.prems(1) unfolding constr_def fltD1_def fltD2_def by moura + using ConsDel.prems(1) unfolding constr_def fltD1_def fltD2_def by atomize_elim auto hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding constr_def fltD1_def by (metis ConsDel.hyps(1) iB(2)) hence 1: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[of "unlabel (fltD1 Di)" "unlabel D"] unfolding unlabel_def fltD1_def by force have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using subseqs_set_subset(1)[OF *(1)] unfolding fltD1_def unlabel_def dbproj_def by auto have 5: "fv\<^sub>l\<^sub>s\<^sub>t A' = fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using * unfolding unlabel_def by force have "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding unlabel_def constr_def fltD1_def fltD2_def pair_def dbproj_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 2 by blast have 4: "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) = fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t AA" using iB by auto have "fv\<^sub>s\<^sub>t (unlabel A') \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 1 3 4 5 by blast thus ?case by metis next case (ConsNegChecks A' D X F F' AA A A') then obtain i B where iB: "A = (i,NegChecks X F F')#B" "AA = unlabel B" - unfolding unlabel_def by moura + unfolding unlabel_def by atomize_elim auto define D' where "D' \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel (dbproj i D))))" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" from iB obtain A'' where *: "A'' \ set (tr\<^sub>p\<^sub>c B D)" "A' = constr@A''" - using ConsNegChecks.prems(1) unfolding constr_def by moura + using ConsNegChecks.prems(1) unfolding constr_def by atomize_elim auto hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (metis ConsNegChecks.hyps(1) iB(2)) hence **: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by auto have 1: "fv\<^sub>l\<^sub>s\<^sub>t constr \ (D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" unfolding D'_def constr_def unlabel_def by auto have "set (unlabel (dbproj i D)) \ set (unlabel D)" unfolding unlabel_def dbproj_def by auto hence 2: "D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset'[of F' "unlabel (dbproj i D)"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono unfolding D'_def by blast have 3: "fv\<^sub>l\<^sub>s\<^sub>t A' \ ((fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using 1 2 *(2) unfolding unlabel_def by fastforce have 4: "fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by (metis ConsNegChecks.hyps(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset) have 5: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using ConsNegChecks.hyps(2) unfolding unlabel_def by force show ?case using ** 3 4 5 by blast qed (fastforce simp add: unlabel_def)+ show ?Q using assms apply (induct "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) by (fastforce simp add: unlabel_def)+ qed lemma tr_par_vars_disj: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using assms tr_par_vars_subset by fast lemma tr_par_trms_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) - then obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura + then obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by atomize_elim auto hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "2.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (3 i t A D) then obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" - by moura + by atomize_elim auto hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "3.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (4 i ac t t' A D) then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" - by moura + by atomize_elim auto hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "4.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (5 i t s A D) hence "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp hence "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set (List.insert (i,t,s) D)" by (metis "5.IH") thus ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (6 i t s A D) from 6 obtain Di A'' B C where A'': "Di \ set (subseqs (dbproj i D))" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "A' = (B@C)@A''" "B = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" "C = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]" - by moura + by atomize_elim auto hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set [d\D. d \ set Di]" by (metis "6.IH") moreover have "set [d\D. d \ set Di] \ set D" using set_filter by auto ultimately have "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by blast hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` snd ` set D" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset trms\<^sub>s\<^sub>s\<^sub>t_cons by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "set Di \ set D" "set [d\dbproj i D . d \ set Di] \ set D" using subseqs_set_subset[OF A''(1)] unfolding dbproj_def by auto hence "trms\<^sub>s\<^sub>t (unlabel B) \ insert (pair (t, s)) (pair ` snd ` set D)" "trms\<^sub>s\<^sub>t (unlabel C) \ insert (pair (t, s)) (pair ` snd ` set D)" using A''(4,5) unfolding unlabel_def by auto hence "trms\<^sub>s\<^sub>t (unlabel (B@C)) \ insert (pair (t,s)) (pair ` snd ` set D)" using unlabel_append[of B C] by auto moreover have "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#unlabel A)" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(3) trms\<^sub>s\<^sub>t_append[of "unlabel (B@C)" "unlabel A'"] unlabel_append[of "B@C" A''] by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (7 i ac t s A D) from 7 obtain d A'' where A'': "d \ set (dbproj i D)" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" - by moura + by atomize_elim force hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "7.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel A') = {pair (t,s), pair (snd d)} \ trms\<^sub>s\<^sub>t (unlabel A'')" using A''(1,3) by auto ultimately show ?case using A''(1) unfolding dbproj_def by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (8 i X F F' A D) define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define B where "B \ \(trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" from 8 obtain A'' where A'': "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = constr@A''" - unfolding constr_def by moura + unfolding constr_def by atomize_elim auto have "trms\<^sub>s\<^sub>t (unlabel A'') \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set D" by (metis A''(1) "8.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel constr) \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` snd ` set D" unfolding unlabel_def constr_def B_def by auto ultimately have "trms\<^sub>s\<^sub>t (unlabel A') \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using A'' unlabel_append[of constr A''] by auto moreover have "set (dbproj i D) \ set D" unfolding dbproj_def by auto hence "B \ pair ` set F' \ pair ` snd ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset'[of F' "map snd (dbproj i D)"] unfolding B_def by force moreover have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i, \X\\\: F \\: F'\)#A)) = pair ` set F' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by auto ultimately show ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed lemma tr_par_wf_trms: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" using tr_par_trms_subset[OF assms(1)] setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2)[OF assms(2)] by auto lemma tr_par_wf': assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" and "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "wf\<^sub>l\<^sub>s\<^sub>t X A'" proof - define P where "P = (\(D::('fun,'var,'lbl) labeleddbstatelist) (A::('fun,'var,'lbl) labeled_stateful_strand). (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {})" have "P D A" using assms(1,4) by (simp add: P_def) with assms(5,3,2) show ?thesis proof (induction A arbitrary: X A' D) case Nil thus ?case by simp next case (Cons a A) obtain i s where i: "a = (i,s)" by (metis surj_pair) note prems = Cons.prems note IH = Cons.IH show ?case proof (cases s) case (Receive ts) note si = Receive i then obtain A'' where A'': "A' = (i,receive\ts\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" - using prems unlabel_Cons(1)[of i s A] by moura + using prems unlabel_Cons(1)[of i s A] by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1,3) by simp next case (Send ts) note si = Send i then obtain A'' where A'': "A' = (i,send\ts\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" - using prems by moura + using prems by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1) by simp next case (Equality ac t t') note si = Equality i then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "ac = Assign \ fv t' \ X" - using prems unlabel_Cons(1)[of i s] by moura + using prems unlabel_Cons(1)[of i s] by atomize_elim force have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t) (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1,3) by (cases ac) simp_all next case (Insert t t') note si = Insert i hence A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,t') D))" "fv t \ X" "fv t' \ X" using prems by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (List.insert (i,t,t') D)) \ X" using prems si by (auto simp add: unlabel_def) have **: "P (List.insert (i,t,t') D) A" using prems(4) si unfolding P_def unlabel_def by fastforce show ?thesis using IH[OF A'(1) * **] A'(2,3) by simp next case (Delete t t') note si = Delete i define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,t'), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" from prems si obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "Di \ set (subseqs (dbproj i D))" unfolding constr_def by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ X" using prems si apply simp using prems si by (fastforce simp add: unlabel_def) have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (auto simp add: unlabel_def) hence **: "P [d\D. d \ set Di] A" using prems si unfolding P_def by fastforce have ***: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_eqs_ineqs_map'[OF _ A''(3) ***] unfolding constr_def by simp next case (InSet ac t t') note si = InSet i then obtain d A'' where A'': "A' = (i,\ac: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set D" using prems by (auto simp add: dbproj_def) have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t \ fv t') (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t \ fv t'" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce have **: "fv (pair (snd d)) \ X" using A''(3) prems(3) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset by fast have ***: "fv (pair (t,t')) = fv t \ fv t'" unfolding pair_def by auto show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1) ** *** by (cases ac) (simp_all add: Un_assoc) next case (NegChecks Y F F') note si = NegChecks i then obtain A'' where A'': "A' = (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" - using prems by moura + using prems by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto have "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" by auto hence **: "P D A" using prems si unfolding P_def by blast show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_negchecks_map' by simp qed qed qed lemma tr_par_wf: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "wf\<^sub>l\<^sub>s\<^sub>t {} A'" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" and "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using tr_par_wf'[OF _ _ _ _ assms(1)] tr_par_wf_trms[OF assms(1,3)] tr_par_vars_disj[OF assms(1)] assms(2) by fastforce+ lemma tr_par_proj: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "proj n B \ set (tr\<^sub>p\<^sub>c (proj n A) (proj n D))" using assms proof (induction A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have IH': "proj n B \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n (List.insert (i,t,s) D)))" using prems IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True thus ?thesis using IH' proj_list_insert(1,2)[of n "(t,s)" D] proj_list_Cons(1,2)[of n _ S] by auto next case False then obtain m where "i = ln m" "n \ m" by (cases i) simp_all thus ?thesis using IH' proj_list_insert(3)[of n _ "(t,s)" D] proj_list_Cons(3)[of n _ "insert\t,s\" S] by auto qed next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di D. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" obtain Di B' where B': "B = constr Di D@B'" "Di \ set (subseqs (dbproj i D))" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems constr_def by fastforce hence "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n [d\D. d \ set Di]))" using IH by simp hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) [d\proj n D. d \ set Di])" by (metis proj_filter) show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr Di D@proj n B'" "Di \ set (subseqs (dbproj i (proj n D)))" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto moreover have "constr Di (proj n D) = constr Di D" using True proj_dbproj(1,2)[of n D] unfolding constr_def by presburger ultimately have "proj n B \ set (tr\<^sub>p\<^sub>c ((i, delete\t,s\)#proj n S) (proj n D))" using IH' unfolding constr_def by force thus ?thesis by (metis proj_list_Cons(1,2) True) next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence *: "(ln n) \ i" by simp have "proj n B = proj n B'" using B'(1) False unfolding constr_def proj_def by auto moreover have "[d\proj n D. d \ set Di] = proj n D" using proj_subseq[OF _ m(2)[symmetric]] m(1) B'(2) by simp ultimately show ?thesis using m(1) IH' proj_list_Cons(3)[OF m(2), of _ S] by auto qed next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::'lbl strand_label \ ('fun,'var) term \ ('fun,'var) term. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain d B' where B': "B = constr d#B'" "d \ set (dbproj i D)" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr d#proj n B'" "d \ set (dbproj i (proj n D))" using B' proj_list_Cons(1,2)[of n _ B'] unfolding constr_def by (force, metis proj_dbproj(1,2)) hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, InSet ac t s)#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) proj_list_Cons(3) unfolding constr_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "InSet ac t s" S] unfolding constr_def by auto qed next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = (\D. map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" obtain B' where B': "B = constr D@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr (proj n D)@proj n B'" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, NegChecks X F F')#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) unfolding constr_def proj_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "NegChecks X F F'" S] unfolding constr_def by auto qed qed (force simp add: proj_def)+ lemma tr_par_preserves_par_comp: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "par_comp A' Sec" proof - let ?M = "\l. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" let ?N = "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A'" have 0: "\l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1) (?M l2) Sec" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp_all { fix l1 l2::'lbl assume *: "l1 \ l2" hence "GSMP_disjoint (?M l1) (?M l2) Sec" using 0(1) by metis moreover have "pair ` snd ` set (proj n []) = {}" for n::'lbl unfolding proj_def by simp hence "?N l1 \ ?M l1" "?N l2 \ ?M l2" using tr_par_trms_subset[OF tr_par_proj[OF assms(2)]] by (metis Un_empty_right)+ ultimately have "GSMP_disjoint (?N l1) (?N l2) Sec" using GSMP_disjoint_subset by presburger } hence 1: "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 A') (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 A') Sec" using 0(1) by metis have 2: "ground Sec" "\s \ Sec. \{} \\<^sub>c s" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis+ show ?thesis using 1 2 unfolding par_comp_def by metis qed lemma tr_preserves_receives: assumes "E \ set (tr\<^sub>p\<^sub>c F D)" "(l, receive\t\) \ set F" shows "(l, receive\t\\<^sub>s\<^sub>t) \ set E" using assms by (induct F D arbitrary: E rule: tr\<^sub>p\<^sub>c.induct) auto lemma tr_preserves_last_receive: assumes "E \ set (tr\<^sub>p\<^sub>c F D)" "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E" shows "\G. suffix ((l, receive\t\)#G) F \ list_all (Not \ is_Receive \ snd) G" (is "\G. ?P G F \ ?Q G") using assms proof (induction F D arbitrary: E rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t' S D) note prems = "2.prems" note IH = "2.IH" obtain E' where E': "E = (i,send\t'\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto obtain G where G: "?P G S" "?Q G" using suffix_Cons'[OF prems(2)[unfolded E'(1)]] IH[OF E'(2)] by blast show ?case by (metis suffix_Cons G) next case (3 i t' S D) note prems = "3.prems" note IH = "3.IH" obtain E' where E': "E = (i,receive\t'\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto show ?case using suffix_Cons'[OF prems(2)[unfolded E'(1)]] proof assume "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" then obtain G where G: "?P G S" "?Q G" using IH[OF E'(2)] by blast show ?thesis by (metis suffix_Cons G) next assume "(i, receive\t'\\<^sub>s\<^sub>t) = (l, receive\t\\<^sub>s\<^sub>t) \ E' = []" hence *: "i = l" "t' = t" "E' = []" by simp_all show ?thesis using tr_preserves_receives[OF E'(2)] unfolding * list_all_iff is_Receive_def by fastforce qed next case (4 i ac t' t'' S D) note prems = "4.prems" note IH = "4.IH" obtain E' where E': "E = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto obtain G where G: "?P G S" "?Q G" using suffix_Cons'[OF prems(2)[unfolded E'(1)]] IH[OF E'(2)] by blast show ?case by (metis suffix_Cons G) next case (5 i t' s S D) note prems = "5.prems" note IH = "5.IH" have "E \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t',s) D))" using prems(1) by auto thus ?case by (metis IH[OF _ prems(2)] suffix_Cons) next case (6 i t' s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t',s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" obtain E' Di where E': "E = constr Di@E'" "E' \ set (tr\<^sub>p\<^sub>c S (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" using prems(1) unfolding constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd ` set (constr Di)" unfolding constr_def by force hence "\suffix [(l, receive\t\\<^sub>s\<^sub>t)] (constr Di)" unfolding suffix_def by auto hence "1 \ length E'" using prems(2) E'(1) by (cases E') auto hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using suffix_length_suffix[OF prems(2) suffixI[OF E'(1)]] by simp thus ?case by (metis IH[OF E'(3,2)] suffix_Cons) next case (7 i ac t' s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::(('lbl strand_label \ ('fun,'var) term \ ('fun,'var) term)). (i,\ac: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain E' d where E': "E = constr d#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(1) unfolding constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd (constr d)" unfolding constr_def by force hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using prems(2) E'(1) suffix_Cons' by fastforce thus ?case by (metis IH[OF E'(2)] suffix_Cons) next case (8 i X G G' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\H. (i,\X\\\: (G@H)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G' (map snd (dbproj i D)))" obtain E' where E': "E = constr@E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd ` set constr" unfolding constr_def by force hence "\suffix [(l, receive\t\\<^sub>s\<^sub>t)] constr" unfolding suffix_def by auto hence "1 \ length E'" using prems(2) E'(1) by (cases E') auto hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using suffix_length_suffix[OF prems(2) suffixI[OF E'(1)]] by simp thus ?case by (metis IH[OF E'(2)] suffix_Cons) qed lemma tr_leaking_prefix_exists: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "prefix B A'" "ik\<^sub>s\<^sub>t (proj_unl n B) \\<^sub>s\<^sub>e\<^sub>t \ \ t" shows "\C D. prefix C B \ prefix D A \ C \ set (tr\<^sub>p\<^sub>c D []) \ (ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t) \ (\{} \\<^sub>c t \ (\l s G. suffix ((l, receive\s\)#G) D \ list_all (Not \ is_Receive \ snd) G))" proof - let ?P = "\B C C'. B = C@C' \ (\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C') \ (C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C))" have "\C C'. ?P B C C'" proof (induction B) case (Cons b B) - then obtain C C' n s where *: "?P B C C'" "b = (n,s)" by moura + then obtain C C' n s where *: "?P B C C'" "b = (n,s)" by atomize_elim auto show ?case proof (cases "C = []") case True note T = True show ?thesis proof (cases "\t. s = receive\t\\<^sub>s\<^sub>t") case True hence "?P (b#B) [b] C'" using * T by auto thus ?thesis by metis next case False hence "?P (b#B) [] (b#C')" using * T by auto thus ?thesis by metis qed next case False hence "?P (b#B) (b#C) C'" using * unfolding suffix_def by auto thus ?thesis by metis qed qed simp then obtain C C' where C: "B = C@C'" "\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C'" "C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C)" - by moura + by atomize_elim auto hence 1: "prefix C B" by simp hence 2: "prefix C A'" using assms(2) by simp have "\m t. (m,receive\t\\<^sub>s\<^sub>t) \ set B \ (m,receive\t\\<^sub>s\<^sub>t) \ set C" using C by auto hence "\t. receive\t\\<^sub>s\<^sub>t \ set (proj_unl n B) \ receive\t\\<^sub>s\<^sub>t \ set (proj_unl n C)" unfolding unlabel_def proj_def by force hence "ik\<^sub>s\<^sub>t (proj_unl n B) \ ik\<^sub>s\<^sub>t (proj_unl n C)" using ik\<^sub>s\<^sub>t_is_rcv_set by blast hence 3: "ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t" by (metis ideduct_mono[OF assms(3)] subst_all_mono) have "\F. prefix F A \ E \ set (tr\<^sub>p\<^sub>c F D)" when "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" "A' \ set (tr\<^sub>p\<^sub>c A D)" for D E m t using that proof (induction A D arbitrary: A' E rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t' S D) note prems = "2.prems" note IH = "2.IH" obtain A'' where *: "A' = (i,send\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,send\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Send t')#F) ((i,Send t')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Send t')#F) D)" using ** by auto thus ?case by metis next case (3 i t' S D) note prems = "3.prems" note IH = "3.IH" obtain A'' where *: "A' = (i,receive\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,receive\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto show ?case proof (cases "(m, receive\t\\<^sub>s\<^sub>t) = (i, receive\t'\\<^sub>s\<^sub>t)") case True note T = True show ?thesis proof (cases "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'") case True hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis next case False hence "E' = []" using **(1) T prems(1) suffix_Cons[of "[(m, receive\t\\<^sub>s\<^sub>t)]" "(m, receive\t\\<^sub>s\<^sub>t)" E'] by auto hence "prefix [(i,receive\t'\)] ((i,receive\t'\) # S) \ E \ set (tr\<^sub>p\<^sub>c [(i,receive\t'\)] D)" using * ** prems by auto thus ?thesis by metis qed next case False hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis qed next case (4 i ac t' t'' S D) note prems = "4.prems" note IH = "4.IH" obtain A'' where *: "A' = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Equality ac t' t'')#F) ((i,Equality ac t' t'')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t' t'')#F) D)" using ** by auto thus ?case by metis next case (5 i t' s S D) note prems = "5.prems" note IH = "5.IH" have *: "A' \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t',s) D))" using prems(3) by auto have "E \ []" using prems(1) by auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" using *(1) prems(1,2) suffix_Cons[of _ _ E] by auto then obtain F where "prefix F S" "E \ set (tr\<^sub>p\<^sub>c F (List.insert (i,t',s) D))" using * IH by metis hence "prefix ((i,insert\t',s\)#F) ((i,insert\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,insert\t',s\)#F) D)" by auto thus ?case by metis next case (6 i t' s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t',s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" obtain A'' Di where *: "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c S (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set (constr Di)" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr Di@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" "constr Di" E'] *** by (metis (no_types, opaque_lifting) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F (filter (\d. d \ set Di) D))" using *(2,3) ** IH by metis hence "prefix ((i,delete\t',s\)#F) ((i,delete\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,delete\t',s\)#F) D)" using *(3) ** constr_def by auto thus ?case by metis next case (7 i ac t' s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::(('lbl strand_label \ ('fun,'var) term \ ('fun,'var) term)). (i,\ac: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain A'' d where *: "A' = constr d#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(3) constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr d#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] using constr_def by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,InSet ac t' s)#F) ((i,InSet ac t' s)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t' s)#F) D)" using *(3) ** unfolding constr_def by auto thus ?case by metis next case (8 i X G G' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\H. (i,\X\\\: (G@H)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G' (map snd (dbproj i D)))" obtain A'' where *: "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set constr" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" constr E'] *** by (metis (no_types, opaque_lifting) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,NegChecks X G G')#F) ((i,NegChecks X G G')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X G G')#F) D)" using ** constr_def by auto thus ?case by metis qed moreover have "prefix [] A" "[] \ set (tr\<^sub>p\<^sub>c [] [])" by auto moreover have "{} \\<^sub>c t" when "C = []" using 3 by (simp add: deducts_eq_if_empty_ik that) ultimately have 4: "\D. prefix D A \ C \ set (tr\<^sub>p\<^sub>c D []) \ (\{} \\<^sub>c t \ (\l s G. suffix ((l, receive\s\)#G) D \ list_all (Not \ is_Receive \ snd) G))" using C(3) assms(1) 2 by (meson tr_preserves_last_receive) show ?thesis by (metis 1 3 4) qed end context labeled_stateful_typing begin lemma tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P0 A D") and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P1 A D") and "\t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. \t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. (\\. Unifier \ t t') \ \ t = \ t'" (is "?P3 A D") shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" proof - have sublmm: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" "?P0 A D" "?P1 A D" "?P3 A D" when p: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (a#A))" "?P0 (a#A) D" "?P1 (a#A) D" "?P3 (a#A) D" for a A D proof - show "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using p(1) by (simp add: unlabel_def tfr\<^sub>s\<^sub>s\<^sub>t_def) show "?P0 A D" using p(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce show "?P1 A D" using p(3) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce have "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (a#A))" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by auto thus "?P3 A D" using p(4) by blast qed show ?thesis using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) note prems = "2.prems" note IH = "2.IH" - from prems(1) obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura + from prems(1) obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (3 i t A D) note prems = "3.prems" note IH = "3.IH" - from prems(1) obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura + from prems(1) obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (4 i ac t t' A D) note prems = "4.prems" note IH = "4.IH" - from prems(1) obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura + from prems(1) obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) prems(2) by simp next case (5 i t s A D) note prems = "5.prems" note IH = "5.IH" from prems(1) have A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp have 1: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using sublmm[OF prems(2,3,4,5)] by simp have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,insert\t,s\)#A)) \ pair`snd`set D = pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set (List.insert (i,t,s) D)" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence 3: "?P3 A (List.insert (i,t,s) D)" using prems(5) by metis moreover have "?P1 A (List.insert (i,t,s) D)" using prems(3,4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "insert\t,s\"] unfolding unlabel_def by fastforce ultimately have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using IH[OF A' sublmm(1,2)[OF prems(2,3,4,5)] _ 3] by metis thus ?case using A'(1) by auto next case (6 i t s A D) note prems = "6.prems" note IH = "6.IH" define constr where constr: "constr \ (\Di. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" from prems(1) obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" unfolding constr by fastforce define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set [d\D. d \ set Di] \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair`snd`set D" using subseqs_set_subset[OF A''(3)] by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "\a\M. \b\M. P a b" when "M \ N" "\a\N. \b\N. P a b" for M N::"('fun, 'var) terms" and P using that by blast ultimately have *: "?P3 A (filter (\d. d \ set Di) D)" using prems(5) by presburger have **: "?P1 A (filter (\d. d \ set Di) D)" using prems(4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "delete\t,s\"] unfolding unlabel_def by fastforce have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(3,2) sublmm(1,2)[OF prems(2,3,4,5)] ** *] by metis have 2: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" when "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" for ac u u' using that A''(1) unfolding constr unlabel_def by force have 3: "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u X)" when "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" for X u using that A''(1) unfolding Q2_def constr unlabel_def dbproj_def by force have 4: "\d\set D. (\\. Unifier \ (pair (t,s)) (pair (snd d))) \ \ (pair (t,s)) = \ (pair (snd d))" using prems(5) by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) { fix ac u u' assume a: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" "\\. Unifier \ u u'" hence "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" using 2 by metis moreover { assume "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'')" hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using 1 Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by fast } moreover { fix d assume "d \ set Di" "u = pair (t,s)" "u' = pair (snd d)" hence "\\. Unifier \ u u' \ \ u = \ u'" using 4 dbproj_subseq_subset A''(3) by fast hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by simp hence "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] using a(2) unfolding unlabel_def by auto } ultimately have "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] a(2) unfolding unlabel_def by auto } moreover { fix u U assume "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" hence "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u U)" using 3 by metis hence "Q1 u U \ Q2 u U" using 1 4 subseqs_set_subset[OF A''(3)] tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] unfolding Q1_def Q2_def by blast } ultimately show ?case using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A'"] unfolding Q1_def Q2_def unlabel_def by blast next case (7 i ac t s A D) note prems = "7.prems" note IH = "7.IH" from prems(1) obtain d A'' where A'': "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set (dbproj i D)" - by moura + by atomize_elim force have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "\ (pair (t,s)) = \ (pair (snd d))" when "\\. Unifier \ (pair (t,s)) (pair (snd d))" using that prems(2,5) A''(3) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def dbproj_def) show ?case using A''(1) 1 2 by fastforce next case (8 i X F F' A D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(M::('fun,'var) terms) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M \ T = [] \ (\s \ set T. s \ Var ` set X))" have Q2_subset: "Q2 M' X" when "M' \ M" "Q2 M X" for X M M' using that unfolding Q2_def by auto have Q2_supset: "Q2 (M \ M') X" when "Q2 M X" "Q2 M' X" for X M M' using that unfolding Q2_def by auto from prems obtain A'' where A'': "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" - using constr_def by moura + using constr_def by atomize_elim auto have 0: "constr = [(i,\X\\\: F\\<^sub>s\<^sub>t)]" when "F' = []" using that unfolding constr_def by simp have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "(F' = [] \ Q1 F X) \ Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" using prems(2) unfolding Q1_def Q2_def by simp have 3: "F' = [] \ Q1 F X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 0 2 tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel constr"] unfolding Q1_def by auto { fix c assume "c \ set (unlabel constr)" hence "\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))). c = \X\\\: (F@G)\\<^sub>s\<^sub>t" unfolding constr_def unlabel_def by force } moreover { fix G assume G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" and c: "\X\\\: (F@G)\\<^sub>s\<^sub>t \ set (unlabel constr)" and e: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" have d_Q2: "Q2 (pair ` set (map snd D)) X" unfolding Q2_def proof (intro allI impI) fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (pair ` set (map snd D))" then obtain d where d: "d \ set (map snd D)" "Fun f T \ subterms (pair d)" by force hence "fv (pair d) \ set X = {}" using prems(4) unfolding pair_def by (force simp add: unlabel_def) thus "T = [] \ (\s \ set T. s \ Var ` set X)" by (metis fv_disj_Fun_subterm_param_cases d(2)) qed have "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F' \ pair ` set (map snd D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset[OF G] unfolding dbproj_def by force hence "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G)) X" using Q2_subset[OF _ Q2_supset[OF e d_Q2]] by metis hence "tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: (F@G)\\<^sub>s\<^sub>t)" by (metis Q2_def tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)) } ultimately have 4: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using Ball_set by blast have 5: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 2 3 4 by metis show ?case using 1 5 A''(1) by (simp add: unlabel_def) qed qed lemma tr_par_tfr: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "tfr\<^sub>s\<^sub>t (unlabel A')" proof - have *: "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using tr_par_trms_subset[OF assms(1)] by simp hence "SMP (trms\<^sub>l\<^sub>s\<^sub>t A') \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using SMP_mono by simp moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>t A')" by (metis tfr_subset(2)[OF _ *]) have **: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) - Var`\" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs unfolding pair_def by auto hence "\ t = \ t'" when "\\. Unifier \ t t'" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" for t t' using that assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel []) = {}" "pair ` snd ` set [] = {}" by auto ultimately have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF assms(1) ** assms(3)] by simp show ?thesis by (metis 1 2 tfr\<^sub>s\<^sub>t_def) qed lemma tr_par_preserves_typing_cond: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "typing_cond (unlabel A')" proof - have "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all hence 1: "wf\<^sub>s\<^sub>t {} (unlabel A')" "fv\<^sub>s\<^sub>t (unlabel A') \ bvars\<^sub>s\<^sub>t (unlabel A') = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (unlabel A'))" "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel A') \ assignment_rhs\<^sub>s\<^sub>t (unlabel A'))" using tr_par_wf[OF assms(3)] Ana_invar_subst' by metis+ have 2: "tfr\<^sub>s\<^sub>t (unlabel A')" by (metis tr_par_tfr assms(2,3) typing_cond\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis by (metis 1 2 typing_cond_def) qed end subsection \Theorem: Semantic Equivalence of Translation\ context labeled_stateful_typed_model begin context begin text \ An alternative version of the translation that does not perform database-state projections. It is used as an intermediate step in the proof of semantic equivalence/correctness. \ private fun tr'\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr'\<^sub>p\<^sub>c [] D = [[]]" | "tr'\<^sub>p\<^sub>c ((i,send\ts\)#A) D = map ((#) (i,send\ts\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,receive\ts\)#A) D = map ((#) (i,receive\ts\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr'\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])@B) (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs D)))" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) D) (tr'\<^sub>p\<^sub>c A D))" | "tr'\<^sub>p\<^sub>c ((i,\X\\\: F \\: F'\)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D)))) (tr'\<^sub>p\<^sub>c A D)" subsubsection \Part 1\ private lemma tr'_par_iff_unlabel_tr: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" shows "(\C \ set (tr'\<^sub>p\<^sub>c A D). B = unlabel C) \ B \ set (tr (unlabel A) (unlabel D))" (is "?A \ ?B") proof { fix C have "C \ set (tr'\<^sub>p\<^sub>c A D) \ unlabel C \ set (tr (unlabel A) (unlabel D))" using assms proof (induction A D arbitrary: C rule: tr'\<^sub>p\<^sub>c.induct) case (5 i t s S D) hence "unlabel C \ set (tr (unlabel S) (unlabel (List.insert (i, t, s) D)))" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) moreover have "insert (i,t,s) (set D) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,insert\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j,p) \ insert (i,t,s) (set D). \(k,q) \ insert (i,t,s) (set D). p = q \ j = k" using "5.prems"(2) by blast hence "unlabel (List.insert (i, t, s) D) = (List.insert (t, s) (unlabel D))" using map_snd_list_insert_distrib[of "(i,t,s)" D] unfolding unlabel_def by simp ultimately show ?case by auto next case (6 i t s S D) let ?f1 = "\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t" let ?g1 = "\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t" let ?f2 = "\d. (i, ?f1 (snd d))" let ?g2 = "\d. (i, ?g1 (snd d))" define constr1 where "constr1 = (\Di. (map ?f1 Di)@(map ?g1 [d\unlabel D. d \ set Di]))" define constr2 where "constr2 = (\Di. (map ?f2 Di)@(map ?g2 [d\D. d \ set Di]))" obtain C' Di where C': "Di \ set (subseqs D)" "C = constr2 Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" - using "6.prems"(1) unfolding constr2_def by moura + using "6.prems"(1) unfolding constr2_def by atomize_elim auto have 0: "set [d\D. d \ set Di] \ set D" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. p = q \ j = k" using "6.prems"(2) by blast have "\(i,p) \ set D \ set Di. \(j,q) \ set D \ set Di. p = q \ i = j" using "6.prems"(2) subseqs_set_subset(1)[OF C'(1)] by blast hence 2: "unlabel [d\D. d \ set Di] = [d\unlabel D. d \ set (unlabel Di)]" using unlabel_filter_eq[of D "set Di"] unfolding unlabel_def by simp have 3: "\f g::('a \ 'a \ 'c). \A B::(('b \ 'a \ 'a) list). map snd ((map (\d. (i, f (snd d))) A)@(map (\d. (i, g (snd d))) B)) = map f (map snd A)@map g (map snd B)" by simp have "unlabel (constr2 Di) = constr1 (unlabel Di)" using 2 3[of ?f1 Di ?g1 "[d\D. d \ set Di]"] by (simp add: constr1_def constr2_def unlabel_def) hence 4: "unlabel C = constr1 (unlabel Di)@unlabel C'" using C'(2) unlabel_append by metis have "unlabel Di \ set (map unlabel (subseqs D))" using C'(1) unfolding unlabel_def by simp hence 5: "unlabel Di \ set (subseqs (unlabel D))" using map_subseqs[of snd D] unfolding unlabel_def by simp show ?case using "6.IH"[OF C'(1,3) 1] 2 4 5 unfolding constr1_def by auto next case (7 i ac t s S D) obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" - using "7.prems"(1) by moura + using "7.prems"(1) by atomize_elim force have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "7.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "7.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by force next case (8 i X F F' S D) obtain C' where C': "C = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" - using "8.prems"(1) by moura + using "8.prems"(1) by atomize_elim auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "8.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "8.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by auto qed (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } thus "?A \ ?B" by blast show "?B \ ?A" using assms proof (induction A arbitrary: B D) case (Cons a A) - obtain ia sa where a: "a = (ia,sa)" by moura + obtain ia sa where a: "a = (ia,sa)" by atomize_elim auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" using a by (cases sa) (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ j = k" using Cons.prems(2) by blast show ?case proof (cases sa) case (Send t) then obtain B' where B': "B = send\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Send by auto next case (Receive t) then obtain B' where B': "B = receive\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Receive by auto next case (Equality ac t t') then obtain B' where B': "B = \ac: t \ t'\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Equality by auto next case (Insert t s) hence B: "B \ set (tr (unlabel A) (List.insert (t,s) (unlabel D)))" using Cons.prems(1) a by auto let ?P = "\i. List.insert (t,s) (unlabel D) = unlabel (List.insert (i,t,s) D)" { obtain j where j: "?P j" "j = ia \ (j,t,s) \ set D" - using labeled_list_insert_eq_ex_cases[of "(t,s)" D ia] by moura + using labeled_list_insert_eq_ex_cases[of "(t,s)" D ia] by atomize_elim auto hence "j = ia" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "?P ia" using j(1) by metis } hence j: "?P ia" by metis have 2: "\(k1, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). \(k2, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). p = q \ k1 = k2" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis using Cons.IH[OF _ 2] j(1) B Insert a by auto next case (Delete t s) define c where "c \ (\(i::'lbl strand_label) Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di@ map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])" define d where "d \ (\Di. map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di@ map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\unlabel D. d \ set Di])" obtain B' Di where B': "B = d Di@B'" "Di \ set (subseqs (unlabel D))" "B' \ set (tr (unlabel A) [d\unlabel D. d \ set Di])" using Cons.prems(1) a Delete unfolding d_def by auto obtain Di' where Di': "Di' \ set (subseqs D)" "unlabel Di' = Di" - using unlabel_subseqsD[OF B'(2)] by moura + using unlabel_subseqsD[OF B'(2)] by atomize_elim auto have 2: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. p = q \ j = k" using 1 subseqs_subset[OF Di'(1)] filter_is_subset[of "\d. d \ set Di'"] by blast have "set Di' \ set D" by (rule subseqs_subset[OF Di'(1)]) hence "\(j, p)\set D \ set Di'. \(k, q)\set D \ set Di'. p = q \ j = k" using Cons.prems(2) by blast hence 3: "[d\unlabel D. d \ set Di] = unlabel [d\D. d \ set Di']" using Di'(2) unlabel_filter_eq[of D "set Di'"] unfolding unlabel_def by auto obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di'])" "B' = unlabel C" using 3 Cons.IH[OF _ 2] B'(3) by auto hence 4: "c ia Di'@C \ set (tr'\<^sub>p\<^sub>c (a#A) D)" using Di'(1) a Delete unfolding c_def by auto have "unlabel (c ia Di') = d Di" using Di' 3 unfolding c_def d_def unlabel_def by auto hence 5: "B = unlabel (c ia Di'@C)" using B'(1) C(2) unlabel_append[of "c ia Di'" C] by simp show ?thesis using 4 5 by blast next case (InSet ac t s) then obtain B' d where B': "B = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" "d \ set (unlabel D)" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a InSet unfolding unlabel_def by auto next case (NegChecks X F F') then obtain B' where B': "B = map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel D))@B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a NegChecks unfolding unlabel_def by auto qed qed simp qed subsubsection \Part 2\ private lemma tr_par_iff_tr'_par: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R3 A D") and "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R4 A D") and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R5 A D") shows "(\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \) \ (\C \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel C\\<^sub>d \)" (is "?P \ ?Q") proof { fix B assume "B \ set (tr\<^sub>p\<^sub>c A D)" "\M; unlabel B\\<^sub>d \" hence ?Q using assms proof (induction A D arbitrary: B M rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i ts S D) note prems = "2.prems" note IH = "2.IH" obtain B' where B': "B = (i,send\ts\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" - using prems(1) by moura + using prems(1) by atomize_elim auto have 1: "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "\t \ set ts. M \ t \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" - using IH[OF B'(2) 1 4 5 6] by moura + using IH[OF B'(2) 1 4 5 6] by atomize_elim auto hence "((i,send\ts\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Send ts)#S) D)" "\M; unlabel ((i,send\ts\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i ts S D) note prems = "3.prems" note IH = "3.IH" - obtain B' where B': "B = (i,receive\ts\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura + obtain B' where B': "B = (i,receive\ts\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by atomize_elim auto have 1: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel C\\<^sub>d \" - using IH[OF B'(2) 1 4 5 6] by moura + using IH[OF B'(2) 1 4 5 6] by atomize_elim auto hence "((i,receive\ts\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,receive\ts\)#S) D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel ((i,receive\ts\\<^sub>s\<^sub>t)#C)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain B' where B': "B = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" - using prems(1) by moura + using prems(1) by atomize_elim auto have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" - using IH[OF B'(2) 1 4 5 6] by moura + using IH[OF B'(2) 1 4 5 6] by atomize_elim auto hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have B: "B \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp have 1: "\M; unlabel B\\<^sub>d \ " using prems(2) B(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S D" using prems(5) by force show ?case using IH[OF B(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain B' Di where B': "Di \ set (subseqs (dbproj i D))" "B = c Di@B'" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" - using prems(1) c by moura + using prems(1) c by atomize_elim auto have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S [d\D . d \ set Di])" "\M; unlabel C\\<^sub>d \" - using IH[OF B'(1,3) 1 4 5 6] by moura + using IH[OF B'(1,3) 1 4 5 6] by atomize_elim auto have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel B'"] unfolding c unlabel_def by auto have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) moreover { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D \ set Di" "(k, q) \ {(i, t, s)} \ set D \ set Di" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D \ set Di. \(k, q) \ {(i, t, s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast ultimately have "\M; ?du2 Di\\<^sub>d \" using labeled_sat_ineq_lift by simp hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 8: "\M; unlabel (d Di@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto have 9: "d Di@C \ set (tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using C(1) dbproj_subseq_in_subseqs[OF B'(1)] unfolding d unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain B' d where B': "B = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" - using prems(1) by moura + using prems(1) by atomize_elim force have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" - using IH[OF B'(2) 1 4 5 6] by moura + using IH[OF B'(2) 1 4 5 6] by atomize_elim auto hence "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 B'(3) unfolding dbproj_def by auto thus ?case by metis next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" - obtain B' where B': "B = c@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) c by moura + obtain B' where B': "B = c@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) c by atomize_elim auto have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" - using IH[OF B'(2) 1 4 5 6] by moura + using IH[OF B'(2) 1 4 5 6] by atomize_elim auto have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel B'"] unfolding c unlabel_def by auto have 8: "d@C \ set (tr'\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using C(1) unfolding d unlabel_def by auto have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?P \ ?Q" by metis { fix C assume "C \ set (tr'\<^sub>p\<^sub>c A D)" "\M; unlabel C\\<^sub>d \" hence ?P using assms proof (induction A D arbitrary: C M rule: tr'\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i ts S D) note prems = "2.prems" note IH = "2.IH" obtain C' where C': "C = (i,send\ts\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" - using prems(1) by moura + using prems(1) by atomize_elim auto have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "\t \ set ts. M \ t \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" - using IH[OF C'(2) 1 4 5 6] by moura + using IH[OF C'(2) 1 4 5 6] by atomize_elim auto hence "((i,send\ts\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Send ts)#S) D)" "\M; unlabel ((i,send\ts\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i ts S D) note prems = "3.prems" note IH = "3.IH" obtain C' where C': "C = (i,receive\ts\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" - using prems(1) by moura + using prems(1) by atomize_elim auto have 1: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel B\\<^sub>d \" - using IH[OF C'(2) 1 4 5 6] by moura + using IH[OF C'(2) 1 4 5 6] by atomize_elim auto hence "((i,receive\ts\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,receive\ts\)#S) D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel ((i,receive\ts\\<^sub>s\<^sub>t)#B)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain C' where C': "C = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" - using prems(1) by moura + using prems(1) by atomize_elim auto have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" - using IH[OF C'(2) 1 4 5 6] by moura + using IH[OF C'(2) 1 4 5 6] by atomize_elim auto hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have C: "C \ set (tr'\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp have 1: "\M; unlabel C\\<^sub>d \ " using prems(2) C(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S (List.insert (i,t,s) D)" using prems(5) by force show ?case using IH[OF C(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain C' Di where C': "Di \ set (subseqs D)" "C = c Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" - using prems(1) c by moura + using prems(1) c by atomize_elim auto have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" "\M; unlabel B\\<^sub>d \" - using IH[OF C'(1,3) 1 4 5 6] by moura + using IH[OF C'(1,3) 1 4 5 6] by atomize_elim auto have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel C'"] unfolding c unlabel_def by auto { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D" "(k, q) \ {(i, t, s)} \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D. \(k, q) \ {(i, t, s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_append by auto hence "\M; ?cu1 Di\\<^sub>d \" by (metis 0(3)) ultimately have *: "Di \ set (subseqs (dbproj i D))" using labeled_sat_eqs_subseqs[OF C'(1)] by simp hence 8: "d Di@B \ set (tr\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using B(1) unfolding d unlabel_def by auto have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) hence "\M; ?du2 Di\\<^sub>d \" by (metis labeled_sat_ineq_dbproj) hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d Di@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" - using prems(1) by moura + using prems(1) by atomize_elim force have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" - using IH[OF C'(2) 1 4 5 6] by moura + using IH[OF C'(2) 1 4 5 6] by atomize_elim auto have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) C'(1) by simp have "(i,t,s) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(fst d, snd d) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" using C'(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\\. Unifier \ (pair (t,s)) (pair (snd d)) \ i = fst d" using prems(3) by blast hence "fst d = i" using 7 by auto hence 8: "d \ set (dbproj i D)" using C'(3) unfolding dbproj_def by auto have 9: "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" using B 8 by auto have 10: "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using B 7 8 by auto show ?case by (metis 9 10) next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" - obtain C' where C': "C = c@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) c by moura + obtain C' where C': "C = c@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) c by atomize_elim auto have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" - using IH[OF C'(2) 1 4 5 6] by moura + using IH[OF C'(2) 1 4 5 6] by atomize_elim auto have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel C'"] unfolding c unlabel_def by auto have 8: "d@B \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using B(1) unfolding d unlabel_def by auto have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?Q \ ?P" by metis qed subsubsection \Part 3\ private lemma tr'_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R A D") and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") proof - have 1: "\(t,s) \ set (unlabel D). (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" using assms(1) unfolding unlabel_def by force have 2: "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" using assms(4) subst_apply_term_empty by blast show ?thesis by (metis tr_sem_equiv'[OF 1 assms(2,3) \] tr'_par_iff_unlabel_tr[OF 2]) qed subsubsection \Part 4\ lemma tr_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") using tr_par_iff_tr'_par[OF assms(4,1,2), of M \] tr'_par_sem_equiv[OF assms] by metis end end subsection \Theorem: The Stateful Compositionality Result, on the Constraint Level\ theorem (in labeled_stateful_typed_model) par_comp_constr_stateful_typed: assumes \: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" and \: "\ \\<^sub>s unlabel \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "(\n. \ \\<^sub>s proj_unl n \) \ (\\' l' ts. prefix \' \ \ suffix [(l', receive\ts\)] \' \ (\' leaks Sec under \))" proof - let ?P = "\n A D. \(i, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. \(j, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" have 1: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "ground {}" using \(2) by simp_all have 2: "\n. \(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" "\n. fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" using 1 sst_vars_proj_subset[of _ \] by fast+ note 3 = par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj[OF \(1)] have 4: "\{}; set (unlabel []) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \'; unlabel \\\<^sub>s \' \ (\B\set (tr\<^sub>p\<^sub>c \ []). \{}; unlabel B\\<^sub>d \')" when \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" for \' using tr_par_sem_equiv[OF 1 _ \'] \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def constr_sem_d_def by auto obtain \' where \': "\' \ set (tr\<^sub>p\<^sub>c \ [])" "\ \ \unlabel \'\" - using 4[OF \(2)] \(1) unfolding constr_sem_d_def by moura + using 4[OF \(2)] \(1) unfolding constr_sem_d_def by atomize_elim auto have \': "(\n. (\ \ \proj_unl n \'\)) \ (\\'' l' ts. prefix \'' \' \ suffix [(l', receive\ts\\<^sub>s\<^sub>t)] \'' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \))" using par_comp_constr_typed[OF tr_par_preserves_par_comp[OF \(1) \'(1)] \'(2) \(2-)] by blast show ?thesis proof (cases "\n. (\ \ \proj_unl n \'\)") case True { fix n assume "\ \ \proj_unl n \'\" hence "\{}; {}; unlabel (proj n \)\\<^sub>s \" using tr_par_proj[OF \'(1), of n] tr_par_sem_equiv[OF 2(1,2) 1(3) _ \(2), of n] 3(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def constr_sem_d_def by force } thus ?thesis using True \(1,2,3) \(1) by metis next case False then obtain \''::"('fun,'var,'lbl) labeled_strand" where \'': "prefix \'' \'" "strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \" using \' by blast have "\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \. \l. (\ \ \unlabel (proj l \'')\) \ ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof - obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \" "\{}; proj_unl m \''@[send\[s]\\<^sub>s\<^sub>t]\\<^sub>d \" using \'' unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def constr_sem_d_def by blast show ?thesis using strand_sem_split(3,4)[OF sm(2)] sm(1) unfolding constr_sem_d_def by auto qed then obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \" "\ \ \unlabel (proj m \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj m \'')) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" - by moura + by atomize_elim auto hence s': "\{} \\<^sub>c s \ \" "s \ \ = s" using \(1) subst_ground_ident[of s \] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto \ \ We now need to show that there is some prefix \B\ of \\''\ that also leaks and where \B \ set (tr C D)\ for some prefix \C\ of \\\ \ obtain B::"('fun,'var,'lbl) labeled_strand" and C G::"('fun,'var,'lbl) labeled_stateful_strand" where BC: "prefix B \'" "prefix C \" "B \ set (tr\<^sub>p\<^sub>c C [])" "ik\<^sub>s\<^sub>t (unlabel (proj m B)) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" "prefix B \''" "\l t. suffix ((l, receive\t\)#G) C" and G: "list_all (Not \ is_Receive \ snd) G" using tr_leaking_prefix_exists[OF \'(1) \''(1) sm(3)] prefix_order.order_trans[OF _ \''(1)] s' by blast obtain C' where C': "C = C'@G" "\l t. suffix [(l, receive\t\)] C'" using BC(6) unfolding suffix_def by (metis append_Cons append_assoc append_self_conv2) have "\{}; unlabel (proj m B)\\<^sub>d \" using sm(2) BC(5) unfolding prefix_def unlabel_def proj_def constr_sem_d_def by auto hence BC': "\ \ \proj_unl m B@[send\[s]\\<^sub>s\<^sub>t]\" using BC(4) unfolding constr_sem_d_def by auto have BC'': "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t B \" using BC(5) sm(1) declassified_prefix_subset by auto have "\n. \ \\<^sub>s (proj_unl n C'@[Send [s]])" proof - have 5: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n C) Sec" for n using \(1) BC(2) par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split(1)[THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj] unfolding prefix_def by auto have "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \(2) BC(2) sst_vars_append_subset(1,2)[of "unlabel C"] unfolding prefix_def unlabel_def by auto hence "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" for n using sst_vars_proj_subset[of _ C] sst_vars_proj_subset[of _ \] by blast hence 6: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "ground {}" for n using 2 by auto have 7: "?P n C []" for n using 5 unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp obtain n where n: "\ \\<^sub>s proj_unl n C" "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using s'(2) tr_par_proj[OF BC(3), of m] BC'(1) tr_par_sem_equiv[OF 6 7 \(2), of m] tr_par_deduct_iff[OF tr_par_proj(1)[OF BC(3)], of \ m s] unfolding proj_def constr_sem_d_def by auto have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C')" using C'(1) G unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def proj_def list_all_iff by fastforce hence 8: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C') \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using n(2) by simp have 9: "\ \\<^sub>s proj_unl n C'" using n(1) C'(1) strand_sem_append_stateful by simp show ?thesis using 8 9 strand_sem_append_stateful by auto qed moreover have "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \" by (metis tr_par_declassified_eq BC(3) BC'') hence "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C' \" using ideduct_mono[of "\{set ts |ts. \\, receive\ts\\ \ set C'} \\<^sub>s\<^sub>e\<^sub>t \" _ "\{set ts |ts. \\, receive\ts\\ \ set (C'@G)} \\<^sub>s\<^sub>e\<^sub>t \"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def C'(1) by auto moreover have "prefix C' \" using BC(2) C' unfolding prefix_def by auto ultimately show ?thesis using C'(2) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by meson qed qed theorem (in labeled_stateful_typing) par_comp_constr_stateful: assumes \: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" and \: "\ \\<^sub>s unlabel \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ (\\<^sub>\ \\<^sub>s unlabel \) \ ((\n. \\<^sub>\ \\<^sub>s proj_unl n \) \ (\\' l' ts. prefix \' \ \ suffix [(l', receive\ts\)] \' \ (\' leaks Sec under \\<^sub>\)))" proof - let ?P = "\n A D. \(i, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. \(j, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" have 1: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "ground {}" using \(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all have 2: "\n. \(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" "\n. fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" using 1(1,2) sst_vars_proj_subset[of _ \] by fast+ have 3: "\n. par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj[OF \(1)] by metis have 4: "\{}; set (unlabel []) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \'; unlabel \\\<^sub>s \' \ (\B\set (tr\<^sub>p\<^sub>c \ []). \{}; unlabel B\\<^sub>d \')" when \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" for \' using tr_par_sem_equiv[OF 1 _ \'] \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def constr_sem_d_def by auto obtain \' where \': "\' \ set (tr\<^sub>p\<^sub>c \ [])" "\ \ \unlabel \'\" - using 4[OF \(2)] \(1) unfolding constr_sem_d_def by moura + using 4[OF \(2)] \(1) unfolding constr_sem_d_def by atomize_elim auto obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\<^sub>\ \ \unlabel \'\" "(\n. (\\<^sub>\ \ \proj_unl n \'\)) \ (\\'' l' ts. prefix \'' \' \ suffix [(l', receive\ts\\<^sub>s\<^sub>t)] \'' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\))" using par_comp_constr[OF tr_par_preserves_par_comp[OF \(1) \'(1)] tr_par_preserves_typing_cond[OF \ \'(1)] \'(2) \(2)] - by moura + by atomize_elim auto have \\<^sub>\': "\\<^sub>\ \\<^sub>s unlabel \" using 4[OF \\<^sub>\(1)] \'(1) \\<^sub>\(4) unfolding constr_sem_d_def by auto show ?thesis proof (cases "\n. (\\<^sub>\ \ \proj_unl n \'\)") case True { fix n assume "\\<^sub>\ \ \proj_unl n \'\" hence "\{}; {}; unlabel (proj n \)\\<^sub>s \\<^sub>\" using tr_par_proj[OF \'(1), of n] tr_par_sem_equiv[OF 2(1,2) 1(3) _ \\<^sub>\(1), of n] 3(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def constr_sem_d_def by force } thus ?thesis using True \\<^sub>\(1,2,3) \\<^sub>\' by metis next case False then obtain \''::"('fun,'var,'lbl) labeled_strand" where \'': "prefix \'' \'" "strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\" using \\<^sub>\ by blast moreover { fix ts l assume *: "\{}; unlabel (proj l \'')@[send\ts\\<^sub>s\<^sub>t]\\<^sub>d \\<^sub>\" have "\\<^sub>\ \ \unlabel (proj l \'')\" "\t \ set ts. ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" using strand_sem_split(3,4)[OF *] unfolding constr_sem_d_def by auto } ultimately have "\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\. \l. (\\<^sub>\ \ \unlabel (proj l \'')\) \ ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def constr_sem_d_def by force then obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\" "\\<^sub>\ \ \unlabel (proj m \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj m \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" - by moura + by atomize_elim auto hence s': "\{} \\<^sub>c s \ \\<^sub>\" "s \ \\<^sub>\ = s" using \(1) subst_ground_ident[of s \\<^sub>\] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto \ \ We now need to show that there is some prefix \B\ of \\''\ that also leaks and where \B \ set (tr C D)\ for some prefix \C\ of \\\ \ obtain B::"('fun,'var,'lbl) labeled_strand" and C G::"('fun,'var,'lbl) labeled_stateful_strand" where BC: "prefix B \'" "prefix C \" "B \ set (tr\<^sub>p\<^sub>c C [])" "ik\<^sub>s\<^sub>t (unlabel (proj m B)) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" "prefix B \''" "\l t. suffix ((l, receive\t\)#G) C" and G: "list_all (Not \ is_Receive \ snd) G" using tr_leaking_prefix_exists[OF \'(1) \''(1) sm(3)] prefix_order.order_trans[OF _ \''(1)] s' by blast obtain C' where C': "C = C'@G" "\l t. suffix [(l, receive\t\)] C'" using BC(6) unfolding suffix_def by (metis append_Cons append_assoc append_self_conv2) have "\{}; unlabel (proj m B)\\<^sub>d \\<^sub>\" using sm(2) BC(5) unfolding prefix_def unlabel_def proj_def constr_sem_d_def by auto hence BC': "\\<^sub>\ \ \proj_unl m B@[send\[s]\\<^sub>s\<^sub>t]\" using BC(4) unfolding constr_sem_d_def by auto have BC'': "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t B \\<^sub>\" using BC(5) sm(1) declassified_prefix_subset by auto have "\n. \\<^sub>\ \\<^sub>s (proj_unl n C'@[Send [s]])" proof - have 5: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n C) Sec" for n using \(1) BC(2) par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split(1)[THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj] unfolding prefix_def by auto have "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \(2) BC(2) sst_vars_append_subset(1,2)[of "unlabel C"] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def prefix_def unlabel_def by auto hence "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" for n using sst_vars_proj_subset[of _ C] sst_vars_proj_subset[of _ \] by blast hence 6: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "ground {}" for n using 2 by auto have 7: "?P n C []" for n using 5 unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp obtain n where n: "\\<^sub>\ \\<^sub>s proj_unl n C" "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" using s'(2) tr_par_proj[OF BC(3), of m] BC'(1) tr_par_sem_equiv[OF 6 7 \\<^sub>\(1), of m] tr_par_deduct_iff[OF tr_par_proj(1)[OF BC(3)], of \\<^sub>\ m s] unfolding proj_def constr_sem_d_def by auto have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C')" using C'(1) G unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def proj_def list_all_iff by fastforce hence 8: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C') \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" using n(2) by simp have 9: "\\<^sub>\ \\<^sub>s proj_unl n C'" using n(1) C'(1) strand_sem_append_stateful by simp show ?thesis using 8 9 strand_sem_append_stateful by auto qed moreover have "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \\<^sub>\" by (metis tr_par_declassified_eq BC(3) BC'') hence "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C' \\<^sub>\" using ideduct_mono[of "\{set ts |ts. \\, receive\ts\\ \ set C'} \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\" _ "\{set ts |ts. \\, receive\ts\\ \ set (C'@G)} \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def C'(1) by auto moreover have "prefix C' \" using BC(2) C' unfolding prefix_def by auto ultimately show ?thesis using \\<^sub>\(1,2,3) \\<^sub>\' C'(2) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by meson qed qed subsection \Theorem: The Stateful Compositionality Result, on the Protocol Level\ context labeled_stateful_typing begin context begin subsubsection \Definitions: Labeled Protocols\ text \ We state our result on the level of protocol traces (i.e., the constraints reachable in a symbolic execution of the actual protocol). Hence, we do not need to convert protocol strands to intruder constraints in the following well-formedness definitions. \ private definition wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s::"('fun,'var,'lbl) labeled_strand set \ bool" where "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s \ \ (\\ \ \. wf\<^sub>l\<^sub>s\<^sub>t {} \) \ (\\ \ \. \\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \' = {})" private definition wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s':: "('fun,'var,'lbl) labeled_strand set \ ('fun,'var,'lbl) labeled_strand \ bool" where "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\\' \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t \) (unlabel \')) \ (\\' \ \. \\'' \ \. fv\<^sub>l\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>t \'' = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>t \ = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \' = {})" private definition typing_cond_prot where "typing_cond_prot \

\ wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s \

\ tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) \ (\\ \ \

. list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel \)) \ Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` unlabel ` \

) \ \(assignment_rhs\<^sub>s\<^sub>t ` unlabel ` \

))" private definition par_comp_prot where "par_comp_prot \

Sec \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec) \ ground Sec \ (\s \ Sec. \{} \\<^sub>c s) \ typing_cond_prot \

" subsubsection \Lemmata: Labeled Protocols\ private lemma wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_eqs_wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s S = wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s' S []" unfolding wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s'_def unlabel_def by auto private lemma par_comp_prot_impl_par_comp: assumes "par_comp_prot \

Sec" "\ \ \

" shows "par_comp \ Sec" proof - have *: "\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using assms(1) unfolding par_comp_prot_def by metis { fix l1 l2::'lbl assume **: "l1 \ l2" hence ***: "GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using * by auto have "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using GSMP_disjoint_subset[OF ***] assms(2) by auto } hence "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" by metis thus ?thesis using assms unfolding par_comp_prot_def par_comp_def by metis qed private lemma typing_cond_prot_impl_typing_cond: assumes "typing_cond_prot \

" "\ \ \

" shows "typing_cond (unlabel \)" proof - have 1: "wf\<^sub>s\<^sub>t {} (unlabel \)" "fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \ = {}" using assms unfolding typing_cond_prot_def wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_def by auto have "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

))" "trms\<^sub>l\<^sub>s\<^sub>t \ \ \(trms\<^sub>l\<^sub>s\<^sub>t ` \

)" "SMP (trms\<^sub>l\<^sub>s\<^sub>t \) - Var`\ \ SMP (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) - Var`\" using assms SMP_mono[of "trms\<^sub>l\<^sub>s\<^sub>t \" "\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)"] unfolding typing_cond_prot_def by (metis, metis, auto) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>t \)" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t \)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetD)+ have 4: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel \)" using assms unfolding typing_cond_prot_def by auto have "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (unlabel \) \ assignment_rhs\<^sub>s\<^sub>t (unlabel \)) \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t ` unlabel ` \

) \ \(assignment_rhs\<^sub>s\<^sub>t ` unlabel ` \

))" using assms(2) by auto hence 5: "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel \) \ assignment_rhs\<^sub>s\<^sub>t (unlabel \))" using assms SMP_mono unfolding typing_cond_prot_def Ana_invar_subst_def by (meson subsetD) show ?thesis using 1 2 3 4 5 unfolding typing_cond_def tfr\<^sub>s\<^sub>t_def by blast qed subsubsection \Theorem: Parallel Compositionality for Labeled Protocols\ private definition component_prot where "component_prot n P \ (\l \ P. \s \ set l. has_LabelN n s \ has_LabelS s)" private definition composed_prot where "composed_prot \

\<^sub>i \ {\. \n. proj n \ \ \

\<^sub>i n}" private definition component_secure_prot where "component_secure_prot n P Sec attack \ (\\ \ P. suffix [(ln n, Send1 (Fun attack []))] \ \ (\\\<^sub>\. (interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)) \ \(\\<^sub>\ \ \proj_unl n \\) \ (\\'. prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))))" private definition component_leaks where "component_leaks n \ Sec \ (\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. (\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))" private definition unsat where "unsat \ \ (\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \(\ \ \unlabel \\))" private theorem par_comp_constr_prot: assumes P: "P = composed_prot Pi" "par_comp_prot P Sec" "\n. component_prot n (Pi n)" and left_secure: "component_secure_prot n (Pi n) Sec attack" shows "\\ \ P. suffix [(ln n, Send1 (Fun attack []))] \ \ unsat \ \ (\m. n \ m \ component_leaks m \ Sec)" proof - { fix \ \' assume \: "\ = \'@[(ln n, Send1 (Fun attack []))]" "\ \ P" let ?P = "\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \ \proj_unl m \'@[Send1 t]\))" have tcp: "typing_cond_prot P" using P(2) unfolding par_comp_prot_def by simp have par_comp: "par_comp \ Sec" "typing_cond (unlabel \)" using par_comp_prot_impl_par_comp[OF P(2) \(2)] typing_cond_prot_impl_typing_cond[OF tcp \(2)] by metis+ have "unlabel (proj n \) = proj_unl n \" "proj_unl n \ = proj_unl n (proj n \)" "\A. A \ Pi n \ proj n A = A" "proj n \ = (proj n \')@[(ln n, Send1 (Fun attack []))]" using P(1,3) \ by (auto simp add: proj_def unlabel_def component_prot_def composed_prot_def) moreover have "proj n \ \ Pi n" using P(1) \ unfolding composed_prot_def by blast moreover { fix A assume "prefix A \" hence *: "prefix (proj n A) (proj n \)" unfolding proj_def prefix_def by force hence "proj_unl n A = proj_unl n (proj n A)" "\I. declassified\<^sub>l\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>t (proj n A) I" unfolding proj_def declassified\<^sub>l\<^sub>s\<^sub>t_alt_def by auto hence "\B. prefix B (proj n \) \ proj_unl n A = proj_unl n B \ (\I. declassified\<^sub>l\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>t B I)" using * by metis } ultimately have *: "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ \(\\<^sub>\ \ \proj_unl n \\) \ (\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))" using left_secure unfolding component_secure_prot_def composed_prot_def suffix_def by metis { fix \ assume \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \unlabel \\" obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\'. prefix \' \ \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \' Sec \\<^sub>\)" - using par_comp_constr[OF par_comp \(2,1)] * by moura + using par_comp_constr[OF par_comp \(2,1)] * by atomize_elim auto hence "\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \ \proj_unl m \'@[Send1 t]\))" using \\<^sub>\(4) * unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def by metis hence ?P using \\<^sub>\(1,2,3) by auto } hence "unsat \ \ (\m. n \ m \ component_leaks m \ Sec)" by (metis unsat_def component_leaks_def) } thus ?thesis unfolding suffix_def by metis qed subsubsection \Theorem: Parallel Compositionality for Stateful Protocols\ private abbreviation wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t V \ \ wf'\<^sub>s\<^sub>s\<^sub>t V (unlabel \)" text \ We state our result on the level of protocol traces (i.e., the constraints reachable in a symbolic execution of the actual protocol). Hence, we do not need to convert protocol strands to intruder constraints in the following well-formedness definitions. \ private definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s::"('fun,'var,'lbl) labeled_stateful_strand set \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s \ \ (\\ \ \. wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t {} \) \ (\\ \ \. \\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" private definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s':: "('fun,'var,'lbl) labeled_stateful_strand set \ ('fun,'var,'lbl) labeled_stateful_strand \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\\' \ \. wf'\<^sub>s\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (unlabel \')) \ (\\' \ \. \\'' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'' = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" private definition typing_cond_prot_stateful where "typing_cond_prot_stateful \

\ wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s \

\ tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

)) \ (\S \ \

. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S))" private definition par_comp_prot_stateful where "par_comp_prot_stateful \

Sec \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec) \ ground Sec \ (\s \ Sec. \{} \\<^sub>c s) \ (\(i,p) \ \\ \ \

. setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ \\ \ \

. setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j) \ typing_cond_prot_stateful \

" private definition component_secure_prot_stateful where "component_secure_prot_stateful n P Sec attack \ (\\ \ P. suffix [(ln n, Send [Fun attack []])] \ \ (\\\<^sub>\. (interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)) \ \(\\<^sub>\ \\<^sub>s (proj_unl n \)) \ (\\'. prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))))" private definition component_leaks_stateful where "component_leaks_stateful n \ Sec \ (\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. (\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))" private definition unsat_stateful where "unsat_stateful \ \ (\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \(\ \\<^sub>s unlabel \))" private lemma wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_eqs_wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s S = wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s' S []" unfolding wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s'_def unlabel_def wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp private lemma par_comp_prot_impl_par_comp_stateful: assumes "par_comp_prot_stateful \

Sec" "\ \ \

" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" proof - have *: "\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using assms(1) unfolding par_comp_prot_stateful_def by argo { fix l1 l2::'lbl assume **: "l1 \ l2" hence ***: "GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using * by auto have "GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using GSMP_disjoint_subset[OF ***] assms(2) by auto } hence "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" by metis moreover have "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j" using assms(1,2) unfolding par_comp_prot_stateful_def by blast ultimately show ?thesis using assms unfolding par_comp_prot_stateful_def par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast qed private lemma typing_cond_prot_impl_typing_cond_stateful: assumes "typing_cond_prot_stateful \

" "\ \ \

" shows "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" proof - have 1: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel \)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using assms unfolding typing_cond_prot_stateful_def wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_def by auto have "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

))" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

)" "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)) - Var`\ \ SMP (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)) - Var`\" using assms SMP_mono[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)"] unfolding typing_cond_prot_stateful_def by (metis, metis, auto) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetD)+ have 4: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel \)" using assms unfolding typing_cond_prot_stateful_def by auto show ?thesis using 1 2 3 4 unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by blast qed private theorem par_comp_constr_prot_stateful: assumes P: "P = composed_prot Pi" "par_comp_prot_stateful P Sec" "\n. component_prot n (Pi n)" and left_secure: "component_secure_prot_stateful n (Pi n) Sec attack" shows "\\ \ P. suffix [(ln n, Send [Fun attack []])] \ \ unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" proof - { fix \ \' assume \: "\ = \'@[(ln n, Send [Fun attack []])]" "\ \ P" let ?P = "\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send [t]])))" have tcp: "typing_cond_prot_stateful P" using P(2) unfolding par_comp_prot_stateful_def by simp have par_comp: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using par_comp_prot_impl_par_comp_stateful[OF P(2) \(2)] typing_cond_prot_impl_typing_cond_stateful[OF tcp \(2)] by metis+ have "unlabel (proj n \) = proj_unl n \" "proj_unl n \ = proj_unl n (proj n \)" "\A. A \ Pi n \ proj n A = A" "proj n \ = (proj n \')@[(ln n, Send [Fun attack []])]" using P(1,3) \ by (auto simp add: proj_def unlabel_def component_prot_def composed_prot_def) moreover have "proj n \ \ Pi n" using P(1) \ unfolding composed_prot_def by blast moreover { fix A assume "prefix A \" hence *: "prefix (proj n A) (proj n \)" unfolding proj_def prefix_def by force hence "proj_unl n A = proj_unl n (proj n A)" "\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) I" by (simp, metis declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq) hence "\B. prefix B (proj n \) \ proj_unl n A = proj_unl n B \ (\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I)" using * by metis } ultimately have *: "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ \(\\<^sub>\ \\<^sub>s (proj_unl n \)) \ (\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))" using left_secure unfolding component_secure_prot_stateful_def composed_prot_def suffix_def by metis { fix \ assume \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \\<^sub>s unlabel \" obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\'. prefix \' \ \ (\' leaks Sec under \\<^sub>\)" - using par_comp_constr_stateful[OF par_comp \(2,1)] * by moura + using par_comp_constr_stateful[OF par_comp \(2,1)] * by atomize_elim auto hence "\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send [t]])))" using \\<^sub>\(4) * unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis hence ?P using \\<^sub>\(1,2,3) by auto } hence "unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" by (metis unsat_stateful_def component_leaks_stateful_def) } thus ?thesis unfolding suffix_def by metis qed end end subsection \Automated Compositionality Conditions\ definition comp_GSMP_disjoint where "comp_GSMP_disjoint public arity Ana \ A' B' A B C \ let B\ = B \\<^sub>s\<^sub>e\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t A)) in has_all_wt_instances_of \ A' A \ has_all_wt_instances_of \ B' B\ \ finite_SMP_representation arity Ana \ A \ finite_SMP_representation arity Ana \ B\ \ (\t \ A. \s \ B\. \ t = \ s \ mgu t s \ None \ (intruder_synth' public arity {} t) \ (\u \ C. is_wt_instance_of_cond \ t u))" definition comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. if i = j then True else (let s = pair' pair_fun p; t = pair' pair_fun q in mgu s (t \ var_rename (max_var s)) = None))" definition comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ let L = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A)); MP0 = \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B; pr = \l. MP0 (proj_unl l A) in length L > 1 \ comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A)) \ (\i \ set L. \j \ set L. if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C)" lemma comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI: fixes pair_fun A MP0 pr defines "MP0 \ \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B" and "pr \ \l. MP0 (proj_unl l A)" assumes L_def: "L = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" and L_gt: "length L > 1" and cpc': "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" and MP0_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A))" and GSMP_disj: "\i \ set L. \j \ set L. if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C" shows "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" using assms unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by presburger lemma comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI': fixes pair_fun A MP0 pr Ms defines "MP0 \ \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B" and "pr \ \l. MP0 (proj_unl l A)" and "M \ \l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}" assumes L_def: "map fst Ms = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" and L_gt: "length (map fst Ms) > 1" and cpc': "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" and MP0_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A))" and GSMP_disj: "\i \ set (map fst Ms). \j \ set (map fst Ms). if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C" shows "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" by (rule comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI[OF L_def L_gt cpc' MP0_wf[unfolded MP0_def] GSMP_disj[unfolded pr_def MP0_def]]) locale labeled_stateful_typed_model' = labeled_typed_model' arity public Ana \ label_witness1 label_witness2 + stateful_typed_model' arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_stateful_typed_model by unfold_locales lemma GSMP_disjoint_if_comp_GSMP_disjoint: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes AB'_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity A'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity B'" and C_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" and AB'_disj: "comp_GSMP_disjoint public arity Ana \ A' B' A B C" shows "GSMP_disjoint A' B' (f C - {m. {} \\<^sub>c m})" using GSMP_disjointI[of A' B' A B] AB'_wf AB'_disj C_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def comp_GSMP_disjoint_def f_def wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff Let_def by blast lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes A: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A (f C - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) let ?Sec = "f C - {m. {} \\<^sub>c m}" let ?L = "remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" let ?N1 = "\B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' Pair) ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?N2 = "\B. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?pr = "\l. ?N1 (proj_unl l A)" let ?\ = "\p. var_rename (max_var (pair p))" note defs = pair_code wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t have 0: "length ?L > 1" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (?N1 (unlabel A))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" (* "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t C) C" *) (* "is_TComp_var_instance_closed \ C" *) "\i \ set ?L. \j \ set ?L. i \ j \ comp_GSMP_disjoint public arity Ana \ (?pr i) (?pr j) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ ?\ p) = None" using A unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code by meson+ have L_in_iff: "l \ set ?L \ (\a \ set A. has_LabelN l a)" for l by force have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using 0(2) unfolding defs by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A))" for l using trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] by blast hence A_proj_wf_trms': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (?N1 (proj_unl l A))" for l unfolding defs by auto note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] (* note 1 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF C_wf_trms] C_wf_trms 0(4)] *) have 2: "GSMP (?N2 (proj_unl l A)) \ GSMP (?N2 (proj_unl l' A))" when "l \ set ?L" for l l' using that L_in_iff GSMP_mono[of "?N2 (proj_unl l A)" "?N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff by fast have 3: "GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" when "l1 \ set ?L" "l2 \ set ?L" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (?N1 (proj_unl l1 A)) (?N1 (proj_unl l2 A)) ?Sec" using 0(4) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t by simp qed obtain a1 a2 where a: "a1 \ set ?L" "a2 \ set ?L" "a1 \ a2" - using remdups_ex2[OF 0(1)] by moura + using remdups_ex2[OF 0(1)] by atomize_elim auto show "ground ?Sec" unfolding f_def by fastforce { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and pq: "\\. Unifier \ (pair p) (pair q)" have "\\. Unifier \ (pair p) (pair q \ ?\ p)" using pq vars_term_disjoint_imp_unifier[OF var_rename_fv_disjoint[of "pair p"], of _ "pair q"] by (metis (no_types, lifting) subst_subst_compose var_rename_inv_comp) hence "i = j" using 0(5) mgu_None_is_subst_neq[of "pair p" "pair q \ ?\ p"] p q by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast show "\l1 l2. l1 \ l2 \ GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" using 2 3 3[OF a] unfolding GSMP_disjoint_def by blast (* show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed *) show "\s \ ?Sec. \{} \\<^sub>c s" by simp qed end locale labeled_stateful_typing' = labeled_stateful_typed_model' arity public Ana \ Pair label_witness1 label_witness2 + stateful_typing_result' arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_stateful_typing by unfold_locales lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t': defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes a: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" and B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" (is "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \") shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B (f C - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) define N1 where "N1 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' Pair) ` setops\<^sub>s\<^sub>s\<^sub>t B" define N2 where "N2 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" define L where "L \ \A::('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand. remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" define \ where "\ \ \p. var_rename (max_var (pair p::('fun, ('fun,'atom) term_type \ nat) term)) ::('fun, ('fun,'atom) term_type \ nat) subst" let ?Sec = "f C - {m. {} \\<^sub>c m}" have 0: "length (L A) > 1" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (N1 (unlabel A))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" (* "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t C) C" *) (* "is_TComp_var_instance_closed \ C" *) "\i \ set (L A). \j \ set (L A). i \ j \ comp_GSMP_disjoint public arity Ana \ (N1 (proj_unl i A)) (N1 (proj_unl j A)) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ \ p) = None" using a unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code L_def N1_def \_def by meson+ note 1 = trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) have N1_iff_N2: "N1 A = N2 A" for A unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t N1_def N2_def by simp have N2_proj_subset: "N2 (proj_unl l A) \ N2 (unlabel A)" for l::'lbl and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using 1(1)[of l A] image_mono[OF 1(2)[of l A], of pair] unfolding N2_def by blast have L_in_iff: "l \ set (L A) \ (\a \ set A. has_LabelN l a)" for l A unfolding L_def by force have L_B_subset_A: "l \ set (L A)" when l: "l \ set (L B)" for l using L_in_iff[of l B] L_in_iff[of l A] B l by fastforce note B_setops = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] have B_proj: "\b \ set (proj l B). \a \ set (proj l A). \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \" for l using proj_instance_ex[OF B] by fast have B': "\t \ N2 (unlabel B). \s \ N2 (unlabel A). \\. t = s \ \ \ ?D \" using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] unfolding N2_def by blast have B'_proj: "\t \ N2 (proj_unl l B). \s \ N2 (proj_unl l A). \\. t = s \ \ \ ?D \" for l using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B_proj] unfolding N2_def by presburger have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (unlabel A))" using N1_iff_N2[of "unlabel A"] 0(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (proj_unl l A))" for l using 1[of l] unfolding N2_def by blast hence A_proj_wf_trms': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (N1 (proj_unl l A))" for l using N1_iff_N2[of "proj_unl l A"] unfolding wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by presburger note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] have 2: "GSMP (N2 (proj_unl l A)) \ GSMP (N2 (proj_unl l' A))" when "l \ set (L A)" for l l' and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using that L_in_iff[of _ A] GSMP_mono[of "N2 (proj_unl l A)" "N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff N2_def by fast have 3: "GSMP (N2 (proj_unl l B)) \ GSMP (N2 (proj_unl l A))" (is "?X \ ?Y") for l proof fix t assume "t \ ?X" hence t: "t \ SMP (N2 (proj_unl l B))" "fv t = {}" unfolding GSMP_def by simp_all have "t \ SMP (N2 (proj_unl l A))" using t(1) B'_proj[of l] SMP_wt_instances_subset[of "N2 (proj_unl l B)" "N2 (proj_unl l A)"] by metis thus "t \ ?Y" using t(2) unfolding GSMP_def by fast qed have "GSMP_disjoint (N2 (proj_unl l1 A)) (N2 (proj_unl l2 A)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (N1 (proj_unl l1 A)) (N1 (proj_unl l2 A)) ?Sec" using 0(4) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis using N1_iff_N2 by simp qed hence 4: "GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 using that 3 unfolding GSMP_disjoint_def by blast { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and pq: "\\. Unifier \ (pair p) (pair q)" obtain p' \p where p': "(i,p') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "p = p' \\<^sub>p \p" "pair p = pair p' \ \p" using p B_setops unfolding pair_def by auto obtain q' \q where q': "(j,q') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "q = q' \\<^sub>p \q" "pair q = pair q' \ \q" using q B_setops unfolding pair_def by auto obtain \ where "Unifier \ (pair p) (pair q)" using pq by blast hence "\\. Unifier \ (pair p') (pair q' \ \ p')" using p'(3) q'(3) var_rename_inv_comp[of "pair q'"] subst_subst_compose vars_term_disjoint_imp_unifier[ OF var_rename_fv_disjoint[of "pair p'"], of "\p \\<^sub>s \" "pair q'" "var_rename_inv (max_var_set (fv (pair p'))) \\<^sub>s \q \\<^sub>s \"] unfolding \_def by fastforce hence "i = j" using mgu_None_is_subst_neq[of "pair p'" "pair q' \ \ p'"] p'(1) q'(1) 0(5) unfolding \_def by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast obtain a1 a2 where a: "a1 \ set (L A)" "a2 \ set (L A)" "a1 \ a2" - using remdups_ex2[OF 0(1)[unfolded L_def]] unfolding L_def by moura + using remdups_ex2[OF 0(1)[unfolded L_def]] unfolding L_def by atomize_elim auto show "\l1 l2. l1 \ l2 \ GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" using 2[of _ B] 4 4[OF a] L_B_subset_A unfolding GSMP_disjoint_def by blast show "ground ?Sec" unfolding f_def by fastforce (* show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed *) show "\s \ ?Sec. \{} \\<^sub>c s" by simp qed end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy @@ -1,2814 +1,2814 @@ (* Title: Stateful_Strands.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Stateful Strands\ theory Stateful_Strands imports Strands_and_Constraints begin subsection \Stateful Constraints\ datatype (funs\<^sub>s\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: 'b) stateful_strand_step = Send (the_msgs: "('a,'b) term list") ("send\_\" 80) | Receive (the_msgs: "('a,'b) term list") ("receive\_\" 80) | Equality (the_check: poscheckvariant) (the_lhs: "('a,'b) term") (the_rhs: "('a,'b) term") ("\_: _ \ _\" [80,80]) | Insert (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("insert\_,_\" 80) | Delete (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("delete\_,_\" 80) | InSet (the_check: poscheckvariant) (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("\_: _ \ _\" [80,80]) | NegChecks (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: "'b list") (the_eqs: "(('a,'b) term \ ('a,'b) term) list") (the_ins: "(('a,'b) term \ ('a,'b) term) list") ("\_\\\: _ \\: _\" [80,80]) where "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ _ _) = []" type_synonym ('a,'b) stateful_strand = "('a,'b) stateful_strand_step list" type_synonym ('a,'b) dbstatelist = "(('a,'b) term \ ('a,'b) term) list" type_synonym ('a,'b) dbstate = "(('a,'b) term \ ('a,'b) term) set" abbreviation "is_Assignment x \ (is_Equality x \ is_InSet x) \ the_check x = Assign" abbreviation "is_Check x \ ((is_Equality x \ is_InSet x) \ the_check x = Check) \ is_NegChecks x" abbreviation "is_Check_or_Assignment x \ is_Equality x \ is_InSet x \ is_NegChecks x" abbreviation "is_Update x \ is_Insert x \ is_Delete x" abbreviation InSet_select ("select\_,_\") where "select\t,s\ \ InSet Assign t s" abbreviation InSet_check ("\_ in _\") where "\t in s\ \ InSet Check t s" abbreviation Equality_assign ("\_ := _\") where "\t := s\ \ Equality Assign t s" abbreviation Equality_check ("\_ == _\") where "\t == s\ \ Equality Check t s" abbreviation NegChecks_Inequality1 ("\_ != _\") where "\t != s\ \ NegChecks [] [(t,s)] []" abbreviation NegChecks_Inequality2 ("\_\_ != _\") where "\x\t != s\ \ NegChecks [x] [(t,s)] []" abbreviation NegChecks_Inequality3 ("\_,_\_ != _\") where "\x,y\t != s\ \ NegChecks [x,y] [(t,s)] []" abbreviation NegChecks_Inequality4 ("\_,_,_\_ != _\") where "\x,y,z\t != s\ \ NegChecks [x,y,z] [(t,s)] []" abbreviation NegChecks_NotInSet1 ("\_ not in _\") where "\t not in s\ \ NegChecks [] [] [(t,s)]" abbreviation NegChecks_NotInSet2 ("\_\_ not in _\") where "\x\t not in s\ \ NegChecks [x] [] [(t,s)]" abbreviation NegChecks_NotInSet3 ("\_,_\_ not in _\") where "\x,y\t not in s\ \ NegChecks [x,y] [] [(t,s)]" abbreviation NegChecks_NotInSet4 ("\_,_,_\_ not in _\") where "\x,y,z\t not in s\ \ NegChecks [x,y,z] [] [(t,s)]" fun trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = set ts" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = set ts" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ F F') = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F'" definition trms\<^sub>s\<^sub>s\<^sub>t where "trms\<^sub>s\<^sub>s\<^sub>t S \ \(trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)" declare trms\<^sub>s\<^sub>s\<^sub>t_def[simp] fun trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = ts" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = ts" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ F F') = concat (map (\(t,t'). [t,t']) (F@F'))" definition trms_list\<^sub>s\<^sub>s\<^sub>t where "trms_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" definition ik\<^sub>s\<^sub>s\<^sub>t where "ik\<^sub>s\<^sub>s\<^sub>t A \ {t | t ts. Receive ts \ set A \ t \ set ts}" definition bvars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "bvars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map (set \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p) S))" fun fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b) stateful_strand_step \ 'b set" where "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks X F F') = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X" definition fv\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "fv\<^sub>s\<^sub>s\<^sub>t S \ \(set (map fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" fun fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) = concat (map fv_list ts)" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) = concat (map fv_list ts)" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\_: t \ s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\_: t \ s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: F'\) = filter (\x. x \ set X) (fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@F'))" definition fv_list\<^sub>s\<^sub>s\<^sub>t where "fv_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" declare bvars\<^sub>s\<^sub>s\<^sub>t_def[simp] declare fv\<^sub>s\<^sub>s\<^sub>t_def[simp] definition vars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "vars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b) stateful_strand_step \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ case x of NegChecks _ _ _ \ {} | Equality Check _ _ \ {} | InSet Check _ _ \ {} | Delete _ _ \ {} | _ \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x" definition wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ case x of Send ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign s t \ fv s | InSet Assign s t \ fv s \ fv t | _ \ {}" definition wfvarsoccs\<^sub>s\<^sub>s\<^sub>t where "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ \(set (map wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" fun wf'\<^sub>s\<^sub>s\<^sub>t::"'b set \ ('a,'b) stateful_strand \ bool" where "wf'\<^sub>s\<^sub>s\<^sub>t V [] = True" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Receive ts#S) = (fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Send ts#S) = wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Equality Assign t t'#S) = (fv t' \ V \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Equality Check _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Insert t s#S) = (fv t \ V \ fv s \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Delete _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (InSet Assign t s#S) = wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv s) S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (InSet Check _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (NegChecks _ _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" abbreviation "wf\<^sub>s\<^sub>s\<^sub>t S \ wf'\<^sub>s\<^sub>s\<^sub>t {} S \ fv\<^sub>s\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" fun subst_apply_stateful_strand_step:: "('a,'b) stateful_strand_step \ ('a,'b) subst \ ('a,'b) stateful_strand_step" (infix "\\<^sub>s\<^sub>s\<^sub>t\<^sub>p" 51) where "send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" | "receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" | "\a: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \a: (t \ \) \ (s \ \)\" | "\a: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \a: (t \ \) \ (s \ \)\" | "insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = insert\t \ \, s \ \\" | "delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = delete\t \ \, s \ \\" | "\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\: (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\" definition subst_apply_stateful_strand:: "('a,'b) stateful_strand \ ('a,'b) subst \ ('a,'b) stateful_strand" (infix "\\<^sub>s\<^sub>s\<^sub>t" 51) where "S \\<^sub>s\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) S" fun dbupd\<^sub>s\<^sub>s\<^sub>t::"('f,'v) stateful_strand \ ('f,'v) subst \ ('f,'v) dbstate \ ('f,'v) dbstate" where "dbupd\<^sub>s\<^sub>s\<^sub>t [] I D = D" | "dbupd\<^sub>s\<^sub>s\<^sub>t (Insert t s#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I (insert ((t,s) \\<^sub>p I) D)" | "dbupd\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I (D - {((t,s) \\<^sub>p I)})" | "dbupd\<^sub>s\<^sub>s\<^sub>t (_#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I D" fun db'\<^sub>s\<^sub>s\<^sub>t::"('f,'v) stateful_strand \ ('f,'v) subst \ ('f,'v) dbstatelist \ ('f,'v) dbstatelist" where "db'\<^sub>s\<^sub>s\<^sub>t [] I D = D" | "db'\<^sub>s\<^sub>s\<^sub>t (Insert t s#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I (List.insert ((t,s) \\<^sub>p I) D)" | "db'\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I (List.removeAll ((t,s) \\<^sub>p I) D)" | "db'\<^sub>s\<^sub>s\<^sub>t (_#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I D" definition db\<^sub>s\<^sub>s\<^sub>t where "db\<^sub>s\<^sub>s\<^sub>t S I \ db'\<^sub>s\<^sub>s\<^sub>t S I []" fun setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ _ F') = set F'" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = {}" text \The set-operations of a stateful strand\ definition setops\<^sub>s\<^sub>s\<^sub>t where "setops\<^sub>s\<^sub>s\<^sub>t S \ \(setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)" fun setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ _ F') = F'" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = []" text \The set-operations of a stateful strand (list variant)\ definition setops_list\<^sub>s\<^sub>s\<^sub>t where "setops_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" subsection \Small Lemmata\ lemma is_Check_or_Assignment_iff[simp]: "is_Check x \ is_Assignment x \ is_Check_or_Assignment x" by (cases x) (blast intro: poscheckvariant.exhaust)+ lemma subst_apply_stateful_strand_step_Inequality[simp]: "\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t \ \ != s \ \\" "\x\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x\t \ rm_vars {x} \ != s \ rm_vars {x} \\" "\x,y\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y\t \ rm_vars {x,y} \ != s \ rm_vars {x,y} \\" "\x,y,z\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y,z\t \ rm_vars {x,y,z} \ != s \ rm_vars {x,y,z} \\" by simp_all lemma subst_apply_stateful_strand_step_NotInSet[simp]: "\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t \ \ not in s \ \\" "\x\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x\t \ rm_vars {x} \ not in s \ rm_vars {x} \\" "\x,y\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y\t \ rm_vars {x,y} \ not in s \ rm_vars {x,y} \\" "\x,y,z\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y,z\t \ rm_vars {x,y,z} \ not in s \ rm_vars {x,y,z} \\" by simp_all lemma trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t: "trms\<^sub>s\<^sub>s\<^sub>t S = set (trms_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding trms\<^sub>s\<^sub>t_def trms_list\<^sub>s\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t: "setops\<^sub>s\<^sub>s\<^sub>t S = set (setops_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding setops\<^sub>s\<^sub>s\<^sub>t_def setops_list\<^sub>s\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = set (fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" proof (cases a) case (NegChecks X F G) thus ?thesis using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append[of F G] fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append[of F G] fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of "F@G"] by auto qed (simp_all add: fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s fv_list_is_fv) lemma fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: "fv\<^sub>s\<^sub>s\<^sub>t S = set (fv_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def fv_list\<^sub>s\<^sub>s\<^sub>t_def by (induct S) (simp_all add: fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p) lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma trms\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (trms\<^sub>s\<^sub>s\<^sub>t S)" using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma vars\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (vars\<^sub>s\<^sub>s\<^sub>t S)" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma fv\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (fv\<^sub>s\<^sub>s\<^sub>t S)" using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x))" by (rule finite_set) lemma bvars\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (bvars\<^sub>s\<^sub>s\<^sub>t S)" using bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma subst_sst_nil[simp]: "[] \\<^sub>s\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_stateful_strand_def) lemma db\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "db\<^sub>s\<^sub>s\<^sub>t [] \ = []" by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "ik\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma in_ik\<^sub>s\<^sub>s\<^sub>t_iff: "t \ ik\<^sub>s\<^sub>s\<^sub>t A \ (\ts. receive\ts\ \ set A \ t \ set ts)" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by blast lemma ik\<^sub>s\<^sub>s\<^sub>t_append[simp]: "ik\<^sub>s\<^sub>s\<^sub>t (A@B) = ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_concat: "ik\<^sub>s\<^sub>s\<^sub>t (concat xs) = \(ik\<^sub>s\<^sub>s\<^sub>t ` set xs)" by (induct xs) auto lemma ik\<^sub>s\<^sub>s\<^sub>t_subst: "ik\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" proof (induction A) case (Cons a A) have "ik\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Receive ts) thus ?thesis using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "[a]"] in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "[a] \\<^sub>s\<^sub>s\<^sub>t \"] unfolding subst_apply_stateful_strand_def by auto qed (simp_all add: ik\<^sub>s\<^sub>s\<^sub>t_def subst_apply_stateful_strand_def) thus ?case using Cons.IH ik\<^sub>s\<^sub>s\<^sub>t_append by (metis append_Cons append_Nil image_Un map_append subst_apply_stateful_strand_def) qed simp lemma ik\<^sub>s\<^sub>s\<^sub>t_set_subset: "set A \ set B \ ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by blast lemma ik\<^sub>s\<^sub>s\<^sub>t_prefix_subset: "prefix A B \ ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" (is "?P A B \ ?P' A B") "prefix A (C@D) \ \prefix A C \ ik\<^sub>s\<^sub>s\<^sub>t C \ ik\<^sub>s\<^sub>s\<^sub>t A" (is "?Q \ ?Q' \ ?Q''") proof - show "?P A B \ ?P' A B" for A B by (metis set_mono_prefix ik\<^sub>s\<^sub>s\<^sub>t_set_subset) thus "?Q \ ?Q' \ ?Q''" by (metis prefixI prefix_same_cases) qed lemma ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty: assumes "\a \ set A. \is_Receive a" shows "ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = {}" using assms in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ A] by fastforce lemma ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_eq: assumes "\s. a = receive\s\" shows "ik\<^sub>s\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty[of "[a]" \] ik\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding is_Receive_def by auto lemma db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t: "set (db'\<^sub>s\<^sub>s\<^sub>t A I D) = dbupd\<^sub>s\<^sub>s\<^sub>t A I (set D)" (is "?A = ?B") proof show "?A \ ?B" proof fix t s show "(t,s) \ ?A \ (t,s) \ ?B" by (induct rule: db'\<^sub>s\<^sub>s\<^sub>t.induct) auto qed show "?B \ ?A" proof fix t s show "(t,s) \ ?B \ (t,s) \ ?A" by (induct arbitrary: D rule: dbupd\<^sub>s\<^sub>s\<^sub>t.induct) auto qed qed lemma db\<^sub>s\<^sub>s\<^sub>t_no_upd: assumes "\a \ set A. \is_Insert a \ \is_Delete a" shows "db'\<^sub>s\<^sub>s\<^sub>t A I D = D" using assms proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_no_upd_append: assumes "\b \ set B. \is_Insert b \ \is_Delete b" shows "db'\<^sub>s\<^sub>s\<^sub>t A = db'\<^sub>s\<^sub>s\<^sub>t (A@B)" using assms proof (induction A) case Nil thus ?case by (simp add: db\<^sub>s\<^sub>s\<^sub>t_no_upd) next case (Cons a A) thus ?case by (cases a) simp_all qed lemma db\<^sub>s\<^sub>s\<^sub>t_append: "db'\<^sub>s\<^sub>s\<^sub>t (A@B) I D = db'\<^sub>s\<^sub>s\<^sub>t B I (db'\<^sub>s\<^sub>s\<^sub>t A I D)" proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) auto qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_in_cases: assumes "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" shows "(t,s) \ set D \ (\t' s'. insert\t',s'\ \ set A \ t = t' \ I \ s = s' \ I)" using assms proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_in_cases': assumes "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" and "(t,s) \ set D" shows "\B C t' s'. A = B@insert\t',s'\#C \ t = t' \ I \ s = s' \ I \ (\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I)" using assms(1) proof (induction A rule: List.rev_induct) case (snoc a A) note * = snoc db\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] thus ?case proof (cases a) case (Insert t' s') thus ?thesis using * by (cases "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)") force+ next case (Delete t' s') hence **: "t \ t' \ I \ s \ s' \ I" using * by simp have "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" using * Delete by force then obtain B C u v where B: "A = B@insert\u,v\#C" "t = u \ I" "s = v \ I" "\t' s'. delete\t',s'\ \ set C \ t \ t' \ I \ s \ s' \ I" - using snoc.IH by moura + using snoc.IH by atomize_elim auto have "A@[a] = B@insert\u,v\#(C@[a])" "\t' s'. delete\t',s'\ \ set (C@[a]) \ t \ t' \ I \ s \ s' \ I" using B(1,4) Delete ** by auto thus ?thesis using B(2,3) by blast qed force+ qed (simp add: assms(2)) lemma db\<^sub>s\<^sub>s\<^sub>t_filter: "db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t (filter is_Update A) I D" by (induct A I D rule: db'\<^sub>s\<^sub>s\<^sub>t.induct) simp_all lemma db\<^sub>s\<^sub>s\<^sub>t_subst_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t A. I x = J x" shows "db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t A J D" using assms proof (induction A arbitrary: D) case (Cons a A) hence "\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. I x = J x" "\D. db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t A J D" by auto thus ?case by (cases a) (simp_all add: term_subst_eq[of _ I J]) qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd: assumes "\a \ set A. \is_Insert a \ \is_Delete a" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D" using assms proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_no_deletes: assumes "list_all (\a. \is_Delete a) A" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D \ {(t \ I, s \ I) | t s. insert\t,s\ \ set A}" (is "?Q A D") using assms proof (induction A arbitrary: D) case (Cons a A) hence IH: "?Q A D" for D by auto have "\is_Delete a" using Cons.prems by simp thus ?case using IH by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_append: "dbupd\<^sub>s\<^sub>s\<^sub>t (A@B) I D = dbupd\<^sub>s\<^sub>s\<^sub>t B I (dbupd\<^sub>s\<^sub>s\<^sub>t A I D)" proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_filter: "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = dbupd\<^sub>s\<^sub>s\<^sub>t (filter is_Update A) I D" by (induct A I D rule: dbupd\<^sub>s\<^sub>s\<^sub>t.induct) simp_all lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases: assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" shows "(t,s) \ D \ (\t' s'. insert\t',s'\ \ set A \ t = t' \ I \ s = s' \ I)" (is ?P) and "\u v B. suffix (delete\u,v\#B) A \ (t,s) = (u,v) \\<^sub>p I \ (\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B)" (is ?Q) proof - show ?P using assms proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp show ?Q using assms proof (induction A arbitrary: D rule: List.rev_induct) case (snoc a A) note 0 = snoc.IH snoc.prems note 1 = suffix_snoc[of _ A a] have 2: "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I D" when "\is_Update a" using that dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] by (cases a) auto have 3: "suffix (delete\u,v\#B) A \ suffix (delete\u,v\#B@[a]) (A@[a])" when "\is_Update a" for u v B using that by simp have 4: "\C. B = C@[a] \ suffix (delete\u,v\#C) A" when a: "\is_Delete a" "suffix (delete\u,v\#B) (A@[a])" for u v B proof - have a': "a \ delete\u,v\" using a(1) by force obtain C where C: "delete\u,v\#B = C@[a]" "suffix C A" using 1 a(2) by blast show ?thesis using a' C by (cases C) auto qed note 5 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I] show ?case proof (cases "is_Update a") case True then obtain u v where "a = insert\u,v\ \ a = delete\u,v\" by (cases a) auto thus ?thesis proof assume a: "a = insert\u,v\" hence a': "\is_Delete a" by simp have 6: "insert\u,v\ \ set B" when B: "suffix (delete\u',v'\#B) (A@[a])" for u' v' B using 4[OF a' B] unfolding a by fastforce have 7: "(t,s) = (u,v) \\<^sub>p I \ (t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using snoc.prems 5 a by auto show ?thesis proof (cases "(t,s) = (u,v) \\<^sub>p I") case True have "insert\u,v\ \ set B" when B: "suffix (delete\u',v'\#B) (A@[a])" for u' v' B using 4[OF a' B] unfolding a by fastforce thus ?thesis using True by blast next case False hence 8: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using 7 by blast have "\u'' v''. (t,s) = (u'',v'') \\<^sub>p I \ insert\u'',v''\ \ set B" when B: "suffix (delete\u',v'\#B) (A @ [a])" "(t,s) = (u',v') \\<^sub>p I" for u' v' B proof - obtain C where C: "B = C@[a]" "suffix (delete\u',v'\#C) A" using 4[OF a' B(1)] by blast thus ?thesis using snoc.IH[OF 8] B(2) unfolding a by fastforce qed thus ?thesis by blast qed next assume a: "a = delete\u,v\" hence "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D - {((u,v) \\<^sub>p I)}" using snoc.prems 5 by auto hence 6: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" "(t,s) \ (u,v) \\<^sub>p I" by (blast,blast) have "(\C. B = C@[a] \ suffix (delete\u',v'\#C) A) \ (B = [] \ u' = u \ v' = v)" when B: "suffix (delete\u',v'\#B) (A@[a])" for B u' v' proof - obtain C where C: "delete\u',v'\#B = C@[a]" "suffix C A" using B 1 by blast show ?thesis proof (cases "B = []") case True thus ?thesis using C unfolding a by simp next case False then obtain b B' where B': "B = B'@[b]" by (meson rev_exhaust) show ?thesis using C unfolding a B' by auto qed qed hence "\C. B = C@[a] \ suffix (delete\u',v'\#C) A" when "suffix (delete\u',v'\#B) (A@[a])" "(t,s) = (u',v') \\<^sub>p I" for B u' v' using that 6 by blast thus ?thesis using snoc.IH[OF 6(1)] unfolding a by fastforce qed next case False have "\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B" when B: "suffix (delete\u,v\#B) (A@[a])" "(t,s) = (u,v) \\<^sub>p I" for u v B proof - obtain C where C: "B = C@[a]" "suffix (delete\u,v\#C) A" using 4[OF _ B(1)] False by blast show ?thesis using B(2) snoc.IH[OF snoc.prems[unfolded 2[OF False]]] C by fastforce qed thus ?thesis by blast qed qed simp qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_iff: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ ((\u v B. suffix (delete\u,v\#B) A \ (t,s) = (u,v) \\<^sub>p I \ (\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B)) \ ((t,s) \ D \ (\u v. (t,s) = (u,v) \\<^sub>p I \ insert\u,v\ \ set A)))" (is "?P A D \ ?Q1 A \ ?Q2 A D") proof show "?P A D \ ?Q1 A \ ?Q2 A D" using dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases by fast show "?Q1 A \ ?Q2 A D \ ?P A D" proof (induction A arbitrary: D) case (Cons a A) have Q1: "?Q1 A" using Cons.prems suffix_Cons[of _ a A] by blast show ?case proof (cases "is_Update a") case False thus ?thesis using Q1 Cons.IH Cons.prems by (cases a) auto next case True then obtain t' s' where "a = insert\t',s'\ \ a = delete\t',s'\" by (cases a) auto thus ?thesis proof assume a: "a = insert\t',s'\" hence "?Q2 A (insert ((t',s') \\<^sub>p I) D)" using Cons.prems by auto thus ?thesis using Q1 Cons.IH unfolding a by auto next assume a: "a = delete\t',s'\" hence "?Q2 A (D - {(t',s') \\<^sub>p I})" using Cons.prems by auto thus ?thesis using Q1 Cons.IH unfolding a by auto qed qed qed simp qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases': fixes A::"('a,'b) stateful_strand" assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" and "(t,s) \ D" shows "\B C t' s'. A = B@insert\t',s'\#C \ t = t' \ I \ s = s' \ I \ (\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I)" using assms(1) proof (induction A rule: List.rev_induct) case (snoc a A) note 0 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] have 1: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" when "\is_Update a" using that snoc.prems 0 by (cases a) auto show ?case proof (cases "is_Update a") case False obtain B C t' s' where B: "A = B@insert\t',s'\#C" "t = t' \ I" "s = s' \ I" "\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I" using snoc.IH[OF 1[OF False]] by blast have "A@[a] = B@insert\t',s'\#(C@[a])" "\t'' s''. delete\t'',s''\ \ set (C@[a]) \ t \ t'' \ I \ s \ s'' \ I" using False B(1,4) by auto thus ?thesis using B(2,3) by blast next case True then obtain t' s' where "a = insert\t',s'\ \ a = delete\t',s'\" by (cases a) auto thus ?thesis proof assume a: "a = insert\t',s'\" hence "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) I D = insert ((t',s') \\<^sub>p I) (dbupd\<^sub>s\<^sub>s\<^sub>t A I D)" using 0 by simp hence "(t,s) = (t',s') \\<^sub>p I \ (t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using snoc.prems by blast thus ?thesis proof assume 2: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" show ?thesis using snoc.IH[OF 2] unfolding a by force qed (force simp add: a) next assume a: "a = delete\t',s'\" hence 2: "t \ t' \ I \ s \ s' \ I" using 0 snoc.prems by simp have "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using 0 snoc.prems a by force then obtain B C u v where B: "A = B@insert\u,v\#C" "t = u \ I" "s = v \ I" "\t' s'. delete\t',s'\ \ set C \ t \ t' \ I \ s \ s' \ I" - using snoc.IH by moura + using snoc.IH by atomize_elim auto have "A@[a] = B@insert\u,v\#(C@[a])" "\t' s'. delete\t',s'\ \ set (C@[a]) \ t \ t' \ I \ s \ s' \ I" using B(1,4) a 2 by auto thus ?thesis using B(2,3) by blast qed qed qed (simp add: assms(2)) lemma dbupd\<^sub>s\<^sub>s\<^sub>t_mono: assumes "D \ E" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ dbupd\<^sub>s\<^sub>s\<^sub>t A I E" using assms proof (induction A arbitrary: D E) case (Cons a A) thus ?case proof (cases a) case (Insert t s) have "insert ((t,s) \\<^sub>p I) D \ insert ((t,s) \\<^sub>p I) E" using Cons.prems by fast thus ?thesis using Cons.IH unfolding Insert by simp next case (Delete t s) have "D - {(t,s) \\<^sub>p I} \ E - {(t,s) \\<^sub>p I}" using Cons.prems by fast thus ?thesis using Cons.IH unfolding Delete by simp qed auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_db_narrow: assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I (D \ E)" and "(t,s) \ D" shows "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I E" using assms proof (induction A arbitrary: D E) case (Cons a A) thus ?case proof (cases a) case (Delete t' s') thus ?thesis using Cons.prems Cons.IH[of "D - {(t',s') \\<^sub>p I}" "E - {(t',s') \\<^sub>p I}"] by (simp add: Un_Diff) qed auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff: assumes f: "f \ k" and A: "\t s. insert\t,s\ \ set A \ (\g ts. s = Fun g ts)" shows "(t,Fun f ts) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ (t,Fun f ts) \ dbupd\<^sub>s\<^sub>s\<^sub>t (filter (\a. \s ss. a = insert\s,Fun k ss\) A) I D" (is "?P A D \ ?P (?f A) D") proof show "?P A D \ ?P (?f A) D" using A proof (induction A arbitrary: D) case (Cons a A) have IH: "?P A D \ ?P (?f A) D" for D using Cons.prems(2) Cons.IH by simp show ?case proof (cases "is_Update a") case True then obtain u s where "a = insert\u,s\ \ a = delete\u,s\" by (cases a) auto thus ?thesis proof assume a: "a = insert\u,s\" obtain g ss where s: "s = Fun g ss" using a Cons.prems(2) by fastforce have 0: "?P A (insert ((u, s) \\<^sub>p I) D)" using a Cons.prems(1) by fastforce show ?thesis proof (cases "g = k") case g: True have "?f (a#A) = ?f A" unfolding a s g by force moreover have "(t,Fun f ts) \ (u, Fun g ss) \\<^sub>p I" using f unfolding g by auto ultimately show ?thesis using IH[OF 0] dbupd\<^sub>s\<^sub>s\<^sub>t_db_narrow[of t "Fun f ts" "?f A" I "{(u, s) \\<^sub>p I}" D] unfolding a s g by force next case g: False have "?f (a#A) = a#?f A" using g unfolding a s by force thus ?thesis using Cons.prems Cons.IH g unfolding a s by force qed next assume a: "a = delete\u,s\" hence "?f (a#A) = a#?f A" by auto thus ?thesis using Cons.prems Cons.IH unfolding a by fastforce qed next case a: False hence "?P A D" using Cons.prems(1) by (cases a) auto hence "?P (?f A) D" using Cons.IH Cons.prems(2) a by fastforce thus ?thesis using a by (cases a) auto qed qed simp have "dbupd\<^sub>s\<^sub>s\<^sub>t (?f A) I D \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" proof (induction A arbitrary: D) case (Cons a A) show ?case proof (cases a) case (Insert t s) have "?f (a#A) = a#?f A \ ?f (a#A) = ?f A" unfolding Insert by force hence "dbupd\<^sub>s\<^sub>s\<^sub>t (?f (a#A)) I D \ dbupd\<^sub>s\<^sub>s\<^sub>t (?f A) I (insert ((t,s) \\<^sub>p I) D)" using dbupd\<^sub>s\<^sub>s\<^sub>t_mono[of D "insert ((t, s) \\<^sub>p I) D"] unfolding Insert by auto thus ?thesis using Cons.IH unfolding Insert by fastforce qed (use Cons.prems Cons.IH in auto) qed simp thus "?P (?f A) D \ ?P A D" by blast qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap: fixes t s defines "fvs \ \A D. fv\<^sub>s\<^sub>s\<^sub>t A \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` D)" assumes "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" (is "?in \ A D") and "\x \ fvs A D. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ D. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set A \ delete\u,v\ \ set A \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" (is "?A \ \ D") and "\x \ fvs A D. \c. \ x = Fun c []" (is "?B \") and "\x \ fvs A D. \c. \ x = Fun c []" (is "?B \") and "\x \ fvs A D. \y \ fvs A D. \ x = \ y \ \ x = \ y" (is "?C \ \ A D") shows "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" (is "?in \ A D") using assms(2-) proof (induction A arbitrary: D rule: List.rev_induct) case Nil then obtain u v where u: "(u,v) \ D" "t \ \ = u \ \" "s \ \ = v \ \" by auto let ?X = "fv t \ fv u" let ?Y = "fv s \ fv v" have 0: "fv u \ fvs [] D" "fv v \ fvs [] D" "fv t \ fvs [] D" "fv s \ fvs [] D" using u(1) unfolding fvs_def by (blast, blast, blast, blast) have 1: "\x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ u))" "\x \ ?Y. \ x = \ x \ (\(\ x \ s) \ \(\ x \ v))" using Nil.prems(2) u(1) unfolding fvs_def by (blast,blast) have 2: "\x \ ?X. \c. \ x = Fun c []" "\x \ ?X. \c. \ x = Fun c []" "\x \ ?Y. \c. \ x = Fun c []" "\x \ ?Y. \c. \ x = Fun c []" using Nil.prems(3,4) 0 by (blast,blast,blast,blast) have 3: "\x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" "\x \ ?Y. \y \ ?Y. \ x = \ y \ \ x = \ y" using Nil.prems(5) 0 by (blast,blast) have "t \ \ = u \ \" "s \ \ = v \ \" using subst_const_swap_eq'[OF u(2) 1(1) 2(1,2) 3(1)] subst_const_swap_eq'[OF u(3) 1(2) 2(3,4) 3(2)] by argo+ thus ?case using u(1) by force next case (snoc a A) have 0: "fvs A D \ fvs (A@[a]) D" "set A \ set (A@[a])" unfolding fvs_def by auto note 1 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] have IH: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.IH[of D] snoc.prems(2-) 0 by blast let ?q0 = "\t s \ \. \x \ fv t \ fv s. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s))" let ?q1 = "\t s \. \x \ fv t \ fv s. \c. \ x = Fun c []" let ?q2 = "\t s \ \. \x \ fv t \ fv s. \y \ fv t \ fv s. \ x = \ y \ \ x = \ y" show ?case proof (cases "is_Update a") case False hence "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) = dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) = dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using 1 by (cases a; auto)+ thus ?thesis using IH snoc.prems(1) by blast next case True then obtain u v where u: "a = insert\u,v\ \ a = delete\u,v\" by (cases a) auto have uv_in: "insert\u,v\ \ set (A@[a]) \ delete\u,v\ \ set (A@[a])" using u by force hence fv_uv: "fv u \ fvs (A@[a]) D" "fv v \ fvs (A@[a]) D" unfolding fvs_def by (force,force) have fv_ts: "fv t \ fvs (A@[a]) D" "fv s \ fvs (A@[a]) D" unfolding fvs_def by (blast,blast) have q0: "?q0 t u \ \" "?q0 s v \ \" "?q0 t u \ \" "?q0 s v \ \" proof - show "?q0 t u \ \" "?q0 s v \ \" using snoc.prems(2) 0 fv_ts fv_uv uv_in by (blast,blast) show "?q0 t u \ \" proof fix x assume "x \ fv t \ fv u" hence "x \ fvs (A@[a]) D" using fv_ts(1) fv_uv(1) by blast thus "\ x = \ x \ (\(\ x \ t) \ \(\ x \ u))" using snoc.prems(2) uv_in by auto qed show "?q0 s v \ \" proof fix x assume "x \ fv s \ fv v" hence "x \ fvs (A@[a]) D" using fv_ts(2) fv_uv(2) by blast thus "\ x = \ x \ (\(\ x \ s) \ \(\ x \ v))" using snoc.prems(2) uv_in by auto qed qed have q1: "?q1 t u \" "?q1 t u \" "?q1 s v \" "?q1 s v \" using snoc.prems(3,4) 0 fv_ts fv_uv by (blast,blast,blast,blast) have q2: "?q2 t u \ \" "?q2 s v \ \" "?q2 t u \ \" "?q2 s v \ \" using snoc.prems(5) 0 fv_ts fv_uv by (blast,blast,blast,blast) from u show ?thesis proof assume a: "a = insert\u,v\" show ?thesis proof (cases "(t \ \, s \ \) = (u,v) \\<^sub>p \") case True hence "(t \ \, s \ \) = (u,v) \\<^sub>p \" using subst_const_swap_eq'[OF _ q0(1) q1(1,2) q2(1)] subst_const_swap_eq'[OF _ q0(2) q1(3,4) q2(2)] by fast thus ?thesis using 1 unfolding a by simp next case False hence "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.prems(1) 1 unfolding a by force hence "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using IH by blast thus ?thesis using 1 unfolding a by simp qed next assume a: "a = delete\u,v\" have "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using snoc.prems(1) dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding a by fastforce hence 2: "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using subst_const_swap_eq'[OF _ q0(3) q1(2,1) q2(3)] subst_const_swap_eq'[OF _ q0(4) q1(4,3) q2(4)] by fast have "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.prems(1) 1 unfolding a by fastforce hence 3: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using IH by blast show ?thesis using 2 3 dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding a by auto qed qed qed lemma subst_sst_cons: "a#A \\<^sub>s\<^sub>s\<^sub>t \ = (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#(A \\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_snoc: "A@[a] \\<^sub>s\<^sub>s\<^sub>t \ = (A \\<^sub>s\<^sub>s\<^sub>t \)@[a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_append[simp]: "A@B \\<^sub>s\<^sub>s\<^sub>t \ = (A \\<^sub>s\<^sub>s\<^sub>t \)@(B \\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_list_all: "list_all is_Send S \ list_all is_Send (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Receive S \ list_all is_Receive (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Equality S \ list_all is_Equality (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Insert S \ list_all is_Insert (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Delete S \ list_all is_Delete (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_InSet S \ list_all is_InSet (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_NegChecks S \ list_all is_NegChecks (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Assignment S \ list_all is_Assignment (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Check S \ list_all is_Check (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Update S \ list_all is_Update (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Check_or_Assignment S \ list_all is_Check_or_Assignment (S \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) note * = list_all_def subst_apply_stateful_strand_def { case 1 thus ?case using Cons.IH(1) unfolding * by (cases x) auto } { case 2 thus ?case using Cons.IH(2) unfolding * by (cases x) auto } { case 3 thus ?case using Cons.IH(3) unfolding * by (cases x) auto } { case 4 thus ?case using Cons.IH(4) unfolding * by (cases x) auto } { case 5 thus ?case using Cons.IH(5) unfolding * by (cases x) auto } { case 6 thus ?case using Cons.IH(6) unfolding * by (cases x) auto } { case 7 thus ?case using Cons.IH(7) unfolding * by (cases x) auto } { case 8 thus ?case using Cons.IH(8) unfolding * by (cases x) fastforce+ } { case 9 thus ?case using Cons.IH(9) unfolding * by (cases x) auto } { case 10 thus ?case using Cons.IH(10) unfolding * by (cases x) auto } { case 11 thus ?case using Cons.IH(11) unfolding * by (cases x) auto } qed simp_all lemma subst_sstp_id_subst: "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p Var = a" by (cases a) auto lemma subst_sst_id_subst: "A \\<^sub>s\<^sub>s\<^sub>t Var = A" by (induct A) (simp, metis subst_sstp_id_subst subst_sst_cons) lemma sst_vars_append_subset: "fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t (A@B)" "bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t (A@B)" "fv\<^sub>s\<^sub>s\<^sub>t B \ fv\<^sub>s\<^sub>s\<^sub>t (A@B)" "bvars\<^sub>s\<^sub>s\<^sub>t B \ bvars\<^sub>s\<^sub>s\<^sub>t (A@B)" by auto lemma sst_vars_disj_cons[simp]: "fv\<^sub>s\<^sub>s\<^sub>t (a#A) \ bvars\<^sub>s\<^sub>s\<^sub>t (a#A) = {} \ fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by auto lemma fv\<^sub>s\<^sub>s\<^sub>t_cons_subset[simp]: "fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t (a#A)" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) - set X" by simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ set X" (is ?A) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [(t,s)] \\: []\) = fv t \ fv s \ set X" (is ?B) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [] \\: [(t,s)]\) = fv t \ fv s \ set X" (is ?C) proof show ?A ?B ?C by auto qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ set X" (is ?A) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [(t,s)] \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ rm_vars (set X) \) \ fv (s \ rm_vars (set X) \) \ set X" (is ?B) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [] \\: [(t,s)]\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ rm_vars (set X) \) \ fv (s \ rm_vars (set X) \) \ set X" (is ?C) proof show ?A ?B ?C by auto qed simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset: "bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t (a#A)" by auto lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" by (cases a) auto lemma bvars\<^sub>s\<^sub>s\<^sub>t_subst: "bvars\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>s\<^sub>t A" using bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[of _ \] by (induct A) (simp_all add: subst_apply_stateful_strand_def) lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_set_cases[simp]: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\)) = set X" by simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks: "\is_NegChecks a \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = []" by (cases a) simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks: "bvars\<^sub>s\<^sub>s\<^sub>t A = bvars\<^sub>s\<^sub>s\<^sub>t (filter is_NegChecks A)" proof (induction A) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t_append[simp]: "vars\<^sub>s\<^sub>s\<^sub>t (A@B) = vars\<^sub>s\<^sub>s\<^sub>t A \ vars\<^sub>s\<^sub>s\<^sub>t B" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma vars\<^sub>s\<^sub>s\<^sub>t_Nil[simp]: "vars\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma vars\<^sub>s\<^sub>s\<^sub>t_Cons: "vars\<^sub>s\<^sub>s\<^sub>t (a#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ vars\<^sub>s\<^sub>s\<^sub>t A" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma fv\<^sub>s\<^sub>s\<^sub>t_Cons: "fv\<^sub>s\<^sub>s\<^sub>t (a#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ fv\<^sub>s\<^sub>s\<^sub>t A" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by simp lemma bvars\<^sub>s\<^sub>s\<^sub>t_Cons: "bvars\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ bvars\<^sub>s\<^sub>s\<^sub>t A" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by auto lemma vars\<^sub>s\<^sub>s\<^sub>t_Cons'[simp]: "vars\<^sub>s\<^sub>s\<^sub>t (send\ts\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (receive\ts\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\a: t \ s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\a: t \ s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (insert\t,s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\a: t \ s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\a: t \ s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\X\\\: F \\: G\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) \ vars\<^sub>s\<^sub>s\<^sub>t A" by (simp_all add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_if_no_bvars: assumes a: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = []" shows "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" proof (cases a) case (NegChecks X F G) hence "set X = {}" using a by fastforce thus ?thesis using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ \] unfolding NegChecks by simp qed (auto simp add: subst_list_set_fv subst_apply_fv_unfold) lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars: assumes A: "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>s\<^sub>t A)" using assms proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_if_no_bvars[of a \] fv\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] subst_sst_cons[of a A \] by simp qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: fixes x::"('a,'b) stateful_strand_step" shows "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" proof (cases x) case (NegChecks X F G) thus ?thesis by (induct F) force+ qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t: fixes S::"('a,'b) stateful_strand" shows "vars\<^sub>s\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of x] by (auto simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = set X \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (simp_all add: sup_commute sup_left_commute vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p) lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = X" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\[]\\\: F \\: G\)) = {}" by simp_all lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set X" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\[]\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t != s\) = fv t \ fv s" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t not in s\) = fv t \ fv s" by simp_all lemma fv\<^sub>s\<^sub>s\<^sub>t_append[simp]: "fv\<^sub>s\<^sub>s\<^sub>t (A@B) = fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t B" by simp lemma bvars\<^sub>s\<^sub>s\<^sub>t_append[simp]: "bvars\<^sub>s\<^sub>s\<^sub>t (A@B) = bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t B" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t_mono: "set A \ set B \ fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t B" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" shows "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" using assms var_is_subterm proof (cases a) case (NegChecks X F F') hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X" using assms by simp thus ?thesis using NegChecks var_is_subterm by fastforce qed force+ lemma fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>s\<^sub>t A \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p by (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t A") auto qed simp lemma var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using assms vars_iff_subtermeq proof (cases a) case (NegChecks X F F') hence "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F')" using assms by simp thus ?thesis using NegChecks vars_iff_subtermeq by force qed force+ lemma var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A) \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) show ?case proof (cases "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A)") case True thus ?thesis using Cons.IH by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) next case False thus ?thesis using Cons.prems var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p by (fastforce simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed qed simp lemma var_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ trms\<^sub>s\<^sub>s\<^sub>t A \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" by (meson var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t UN_I term.order_refl) lemma ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset: "ik\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>s\<^sub>t A" by (force simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A) \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" using var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast lemma var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: assumes "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t A" proof - obtain ts where ts: "Receive ts \ set A" "Var x \\<^sub>s\<^sub>e\<^sub>t set ts" - using assms unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by moura + using assms unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by atomize_elim auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t A" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by force thus ?thesis using ts(2) subterm_is_var by fastforce qed lemma fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t A" using var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t assms var_is_subterm by fastforce lemma fv_trms\<^sub>s\<^sub>s\<^sub>t_subset: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S) \ vars\<^sub>s\<^sub>s\<^sub>t S" "fv\<^sub>s\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" proof (induction S) case (Cons x S) have *: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (x#S)) = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" "fv\<^sub>s\<^sub>s\<^sub>t (x#S) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>s\<^sub>t S" "vars\<^sub>s\<^sub>s\<^sub>t (x#S) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ vars\<^sub>s\<^sub>s\<^sub>t S" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def vars\<^sub>s\<^sub>s\<^sub>t_def by auto { case 1 show ?case using Cons.IH(1) proof (cases x) case (NegChecks X F G) hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ set X" by (simp, meson vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(7)) hence "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of G] by auto thus ?thesis using Cons.IH(1) *(1,3) by blast qed auto } { case 2 show ?case using Cons.IH(2) proof (cases x) case (NegChecks X F G) hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) - set X" by auto hence "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of G] by auto thus ?thesis using Cons.IH(2) *(1,2) by blast qed auto } qed simp_all lemma fv_ik_subset_fv_sst'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>s\<^sub>t S" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma fv_ik_subset_vars_sst'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ vars\<^sub>s\<^sub>s\<^sub>t S" using fv_ik_subset_fv_sst' fv_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast lemma ik\<^sub>s\<^sub>s\<^sub>t_var_is_fv: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A) \ x \ fv\<^sub>s\<^sub>s\<^sub>t A" by (meson fv_ik_subset_fv_sst'[of A] fv_subset_subterms subsetCE term.set_intros(3)) lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases': assumes x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" using x vars_term_subst[of _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(1,2,3,4,5,6) vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(1,2)[of _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(3,6)[of _ _ _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(4,5)[of _ _ \] proof (cases s) case (NegChecks X F G) let ?\' = "rm_vars (set X) \" have "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\') \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\') \ x \ set X" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] x NegChecks by simp hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (?\' ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (?\' ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) \ x \ set X" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ ?\'] by blast hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) \ x \ set X" using rm_vars_fv\<^sub>s\<^sub>e\<^sub>t_subst by fast thus ?thesis using NegChecks vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(7)[of X F G] by auto qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t_subst_cases: assumes "x \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t S \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>s\<^sub>s\<^sub>t S)" using assms proof (induction S) case (Cons s S) thus ?case proof (cases "x \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False note * = subst_sst_cons[of s S \] vars\<^sub>s\<^sub>s\<^sub>t_Cons[of "s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "S \\<^sub>s\<^sub>s\<^sub>t \"] vars\<^sub>s\<^sub>s\<^sub>t_Cons[of s S] have **: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems False * by simp show ?thesis using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases'[OF **] * by auto qed (auto simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed simp lemma subset_subst_pairs_diff_exists: fixes \::"('a,'b) subst" and D D'::"('a,'b) dbstate" shows "\Di. Di \ D \ Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - D'" by (metis (no_types, lifting) Diff_subset subset_image_iff) lemma subset_subst_pairs_diff_exists': fixes \::"('a,'b) subst" and D::"('a,'b) dbstate" assumes "finite D" shows "\Di. Di \ D \ Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \} \ d \\<^sub>p \ \ (D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction D rule: finite_induct) case (insert d' D) - then obtain Di where IH: "Di \ D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by moura + then obtain Di where IH: "Di \ D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by atomize_elim auto show ?case proof (cases "d' \\<^sub>p \ = d \\<^sub>p \") case True hence "insert d' Di \ insert d' D" "insert d' Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (insert d' D - insert d' Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using IH by auto thus ?thesis by metis next case False hence "Di \ insert d' D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (insert d' D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using IH by auto thus ?thesis by metis qed qed simp lemma stateful_strand_step_subst_inI[intro]: "send\ts\ \ set A \ send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "receive\ts\ \ set A \ receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\c: t \ s\ \ set A \ \c: (t \ \) \ (s \ \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "insert\t, s\ \ set A \ insert\t \ \, s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "delete\t, s\ \ set A \ delete\t \ \, s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\c: t \ s\ \ set A \ \c: (t \ \) \ (s \ \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\X\\\: F \\: G\ \ set A \ \X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\: (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\t != s\ \ set A \ \t \ \ != s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\t not in s\ \ set A \ \t \ \ not in s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction A) case (Cons a A) note * = subst_sst_cons[of a A \] { case 1 thus ?case using Cons.IH(1) * by (cases a) auto } { case 2 thus ?case using Cons.IH(2) * by (cases a) auto } { case 3 thus ?case using Cons.IH(3) * by (cases a) auto } { case 4 thus ?case using Cons.IH(4) * by (cases a) auto } { case 5 thus ?case using Cons.IH(5) * by (cases a) auto } { case 6 thus ?case using Cons.IH(6) * by (cases a) auto } { case 7 thus ?case using Cons.IH(7) * by (cases a) auto } { case 8 thus ?case using Cons.IH(8) * by (cases a) auto } { case 9 thus ?case using Cons.IH(9) * by (cases a) auto } qed simp_all lemma stateful_strand_step_cases_subst: "is_Send a = is_Send (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Receive a = is_Receive (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Equality a = is_Equality (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Insert a = is_Insert (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Delete a = is_Delete (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_InSet a = is_InSet (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_NegChecks a = is_NegChecks (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Assignment a = is_Assignment (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Check a = is_Check (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Update a = is_Update (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Check_or_Assignment a = is_Check_or_Assignment (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" by (cases a; simp_all)+ lemma stateful_strand_step_substD: "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts\ \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = send\ts'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts\ \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = receive\ts'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \c: t \ s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \c: t' \ s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = insert\t,s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = insert\t',s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = delete\t,s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = delete\t',s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \c: t \ s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \c: t' \ s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \X\\\: F \\: G\ \ \F' G'. F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ a = \X\\\: F' \\: G'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t != s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \t' != s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t not in s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \t' not in s'\" by (cases a; auto simp add: subst_apply_pairs_def; fail)+ lemma stateful_strand_step_mem_substD: "send\ts\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ send\ts'\ \ set S" "receive\ts\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ receive\ts'\ \ set S" "\c: t \ s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \c: t' \ s'\ \ set S" "insert\t,s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ insert\t',s'\ \ set S" "delete\t,s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ delete\t',s'\ \ set S" "\c: t \ s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \c: t' \ s'\ \ set S" "\X\\\: F \\: G\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \F' G'. F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ \X\\\: F' \\: G'\ \ set S" "\t != s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \t' != s'\ \ set S" "\t not in s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \t' not in s'\ \ set S" proof (induction S) case (Cons a S) have *: "x \ set (S \\<^sub>s\<^sub>s\<^sub>t \)" when "x \ set (a#S \\<^sub>s\<^sub>s\<^sub>t \)" "x \ a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" for x using that by (simp add: subst_apply_stateful_strand_def) { case 1 thus ?case using Cons.IH(1)[OF *] by (cases a) auto } { case 2 thus ?case using Cons.IH(2)[OF *] by (cases a) auto } { case 3 thus ?case using Cons.IH(3)[OF *] by (cases a) auto } { case 4 thus ?case using Cons.IH(4)[OF *] by (cases a) auto } { case 5 thus ?case using Cons.IH(5)[OF *] by (cases a) auto } { case 6 thus ?case using Cons.IH(6)[OF *] by (cases a) auto } { case 7 thus ?case using Cons.IH(7)[OF *] by (cases a) auto } { case 8 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "\t != s\ = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using NegChecks stateful_strand_step_substD(8)[of a \ t s] by force qed (use 8 Cons.IH(8)[OF *] in auto) qed (use 8 Cons.IH(8)[OF *] in simp_all) } { case 9 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "\t not in s\ = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using NegChecks stateful_strand_step_substD(9)[of a \ t s] by force qed (use 9 Cons.IH(9)[OF *] in auto) qed (use 9 Cons.IH(9)[OF *] in simp_all) } qed simp_all lemma stateful_strand_step_fv_subset_cases: "send\ts\ \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t S" "receive\ts\ \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t S" "\c: t \ s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\c: t \ s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set X \ fv\<^sub>s\<^sub>s\<^sub>t S" "\t != s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\t not in s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" proof (induction S) case (Cons a S) { case 1 thus ?case using Cons.IH(1) by auto } { case 2 thus ?case using Cons.IH(2) by auto } { case 3 thus ?case using Cons.IH(3) by auto } { case 4 thus ?case using Cons.IH(4) by auto } { case 5 thus ?case using Cons.IH(5) by auto } { case 6 thus ?case using Cons.IH(6) by auto } { case 7 thus ?case using Cons.IH(7) by fastforce } { case 8 thus ?case using Cons.IH(8) by fastforce } { case 9 thus ?case using Cons.IH(9) by fastforce } qed simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "trms\<^sub>s\<^sub>s\<^sub>t [] = {}" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by simp lemma trms\<^sub>s\<^sub>s\<^sub>t_mono: "set M \ set N \ trms\<^sub>s\<^sub>s\<^sub>t M \ trms\<^sub>s\<^sub>s\<^sub>t N" by auto lemma trms\<^sub>s\<^sub>s\<^sub>t_memI[intro?]: "send\ts\ \ set S \ t \ set ts \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "receive\ts\ \ set S \ t \ set ts \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\ac: t \ s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\ac: t \ s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce+ lemma trms\<^sub>s\<^sub>s\<^sub>t_in: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t S" shows "\a \ set S. t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using assms unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by simp lemma trms\<^sub>s\<^sub>s\<^sub>t_cons: "trms\<^sub>s\<^sub>s\<^sub>t (a#A) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ trms\<^sub>s\<^sub>s\<^sub>t A" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force lemma trms\<^sub>s\<^sub>s\<^sub>t_append[simp]: "trms\<^sub>s\<^sub>s\<^sub>t (A@B) = trms\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>s\<^sub>t B" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" shows "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (NegChecks X F G) hence "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using NegChecks image_Un by simp_all thus ?thesis by (metis trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst) qed simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst': assumes "\is_NegChecks a" shows "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" using assms by (cases a) simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "\s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \ rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" proof (cases "is_NegChecks b") case True - then obtain X F G where *: "b = NegChecks X F G" by (cases b) moura+ + then obtain X F G where *: "b = NegChecks X F G" by (cases b) auto thus ?thesis using assms trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ "rm_vars (set X) \"] by auto next case False hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>s\<^sub>e\<^sub>t rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst' bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks by fastforce thus ?thesis using assms by fast qed lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''': fixes t::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>e\<^sub>t \" shows "\s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \ rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \ \\<^sub>s \" proof - - obtain s where s: "s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "t = s \ \" using assms by moura + obtain s where s: "s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "t = s \ \" using assms by atomize_elim auto show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF s(1)] s(2) by auto qed lemma trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons a S) hence IH: "trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" by auto show ?case using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *] IH by (auto simp add: subst_apply_stateful_strand_def) qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_subst_cons: "trms\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \)" using subst_sst_cons[of a A \] trms\<^sub>s\<^sub>s\<^sub>t_cons[of a A] trms\<^sub>s\<^sub>s\<^sub>t_append by simp lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using assms proof (cases a) case (NegChecks X F G) hence *: "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by simp have "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using NegChecks image_Un by simp hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using * assms by auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using wf_trms_subst_rm_vars[of \ "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "set X"] wf_trms_subst_rm_vars[of \ "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "set X"] by fast+ thus ?thesis using * trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ "rm_vars (set X) \"] by auto qed auto lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset: "t \ trms\<^sub>s\<^sub>s\<^sub>t A \ fv t \ vars\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t S" "subst_domain \ \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons s S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>s\<^sub>t S") case True hence "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH Cons.prems by auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto next case False hence *: "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" "subst_domain \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" using Cons.prems by auto hence "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases s) case (NegChecks X F G) hence **: "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using *(1) by auto have ***: "rm_vars (set X) \ = \" using *(2) NegChecks rm_vars_apply' by auto have "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using ** *** trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset[of t _ \] by auto thus ?thesis using *(2) using NegChecks vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] by blast qed auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset': assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" "fv t \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" "fv (t \ \) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons s S) show ?case proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)") case True hence "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH Cons.prems by auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto next case False hence 0: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" "fv t \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" "fv (t \ \) \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" using Cons.prems by auto note 1 = UN_Un UN_insert fv\<^sub>s\<^sub>e\<^sub>t.simps subst_apply_fv_subset subst_apply_fv_unfold subst_apply_term_empty sup_bot.comm_neutral fv_subterms_set fv_subset[OF 0(1)] note 2 = subst_apply_fv_union have "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases s) case (Send ts) have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using fv_subset[OF 0(1)] unfolding Send fv_subterms_set by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (metis subst_apply_fv_subset subst_apply_fv_unfold_set) thus ?thesis using Send by simp next case (Receive ts) have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using fv_subset[OF 0(1)] unfolding Receive fv_subterms_set by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (metis subst_apply_fv_subset subst_apply_fv_unfold_set) thus ?thesis using Receive by simp next case (NegChecks X F G) hence 3: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G)" using 0(1) by auto have "t \ rm_vars (set X) \ = t \ \" using 0(2) NegChecks rm_vars_ident[of t] by auto hence "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using 3 trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset'[of t _ "rm_vars (set X) \"] by fastforce thus ?thesis using 0(2,3) NegChecks fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] by auto qed (metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(3) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(3), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(4) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(4), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(5) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(5), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(6) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(6)) thus ?thesis using subst_sst_cons[of s S \] unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_funs_term_cases: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "f \ funs_term t" shows "(\u \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. f \ funs_term u) \ (\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. f \ funs_term (\ x))" using assms proof (cases s) case (NegChecks X F G) hence "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using assms(1) by auto hence "(\u\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term u) \ (\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term (rm_vars (set X) \ x)) \ (\u\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term u) \ (\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_funs_term_cases[OF _ assms(2), of _ "rm_vars (set X) \"] by meson hence "(\u \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term u) \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x))" by blast thus ?thesis proof assume "\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x)" then obtain x where x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "f \ funs_term (rm_vars (set X) \ x)" by auto hence "x \ set X" "rm_vars (set X) \ x = \ x" by auto thus ?thesis using x by (auto simp add: assms NegChecks) qed (auto simp add: assms NegChecks) qed (use assms funs_term_subst[of _ \] in auto) lemma trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "f \ funs_term t" shows "(\u \ trms\<^sub>s\<^sub>s\<^sub>t S. f \ funs_term u) \ (\x \ fv\<^sub>s\<^sub>s\<^sub>t S. f \ funs_term (\ x))" using assms(1) proof (induction S) case (Cons s S) thus ?case proof (cases "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False hence "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(1) subst_sst_cons[of s S \] trms\<^sub>s\<^sub>s\<^sub>t_cons by auto thus ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_funs_term_cases[OF _ assms(2)] by fastforce qed auto qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t T" and "bvars\<^sub>s\<^sub>s\<^sub>t T \ subst_domain \ = {}" shows "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" using trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms(2)] subterms_subst_subset'[of \ "trms\<^sub>s\<^sub>s\<^sub>t T"] fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t[OF assms(1)] by (metis (no_types, lifting) image_iff subset_iff eval_term.simps(1)) lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ bvars\<^sub>s\<^sub>s\<^sub>t S" "fv (\ x) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (\ x) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons a S) note 1 = fv_subst_subset[of _ _ \] note 2 = subst_apply_fv_union subst_apply_fv_unfold[of _ \] fv_subset image_eqI note 3 = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases note 4 = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps from Cons show ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t S") case False hence 5: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" " fv (\ x) \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "x \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" using Cons.prems by auto hence "fv (\ x) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases a) case (Send ts) thus ?thesis using 1 5(1) 3(1) 4(1) by auto next case (Receive ts) thus ?thesis using 1 5(1) 3(2) 4(2) by auto next case (NegChecks X F G) let ?\ = "rm_vars (set X) \" have *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using NegChecks 5(1) by auto have **: "fv (\ x) \ set X = {}" using NegChecks 5(2) by simp have ***: "\ x = ?\ x" using NegChecks 5(3) by auto have "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset[of x _ ?\] * *** by auto thus ?thesis using NegChecks ** by auto qed (metis (full_types) 2 5(1) 3(3) 4(3), metis (full_types) 2 5(1) 3(4) 4(4), metis (full_types) 2 5(1) 3(5) 4(5), metis (full_types) 2 5(1) 3(6) 4(6)) thus ?thesis by (auto simp add: subst_sst_cons[of a S \]) qed (auto simp add: subst_sst_cons[of a S \]) qed simp lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \))" using assms proof (induction A) case (Cons a A) hence IH: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \))" and *: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \)" by auto have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" by (rule wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *]) thus ?case using IH trms\<^sub>s\<^sub>s\<^sub>t_subst_cons[of a A \] by blast qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\y \ fv\<^sub>s\<^sub>s\<^sub>t S. x \ fv (\ y)" using assms proof (induction S) case (Cons s S) hence "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ \y \ fv\<^sub>s\<^sub>s\<^sub>t S. x \ fv (\ y)" using bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of S s] by blast thus ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False hence *: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(1) subst_sst_cons[of s S \] by fastforce have "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. x \ fv (\ y)" proof (cases s) case (NegChecks X F G) hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" and **: "x \ set X" using * by simp_all then obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv ((rm_vars (set X) \) y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] by blast have "y \ set X" proof assume y_in: "y \ set X" hence "(rm_vars (set X) \) y = Var y" by auto hence "x = y" using y(2) by simp thus False using ** y_in by metis qed thus ?thesis using NegChecks y by auto qed (use * fv_subst_obtain_var in force)+ thus ?thesis by auto qed auto qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_subset_range_vars_if_subset_domain: assumes "fv\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \" shows "fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ range_vars \" using assms fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[of _ S \] subst_dom_vars_in_subst[of _ \] subst_fv_imgI[of \] by (metis (no_types) in_mono subsetI) lemma fv\<^sub>s\<^sub>s\<^sub>t_in_fv_trms\<^sub>s\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>s\<^sub>t S \ x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" proof (induction S) case (Cons s S) thus ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t S") case False hence *: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" using Cons.prems by simp hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" proof (cases s) case (NegChecks X F G) hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using * by simp_all thus ?thesis using * fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_in_fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of x] NegChecks by auto qed auto thus ?thesis by simp qed simp qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)" proof - note 0 = fv_ground_subst_compose have 1: "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" using assms(2,3) by (blast,blast) note 2 = 0[OF assms, of _ \] show ?thesis proof (cases a) case (NegChecks X F G) have 3: "range_vars \ \ set X = {}" "range_vars (rm_vars (set X) \) = {}" "range_vars \ \ set X = {}" "range_vars (rm_vars (set X) \) = {}" using assms(2,3) rm_vars_img_fv_subset[of "set X"] by auto have 4: "subst_domain (rm_vars (set X) \) = subst_domain (rm_vars (set X) \)" using assms(1) rm_vars_dom[of "set X"] by blast have 5: "fv (t \ rm_vars (set X) (\ \\<^sub>s \)) = fv (t \ rm_vars (set X) (\ \\<^sub>s \))" for t using 2[of t] rm_vars_comp[OF 3(1), of t \] rm_vars_comp[OF 3(3), of t \] 0[OF 4 3(2,4), of t "rm_vars (set X) \"] by argo have 6: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \)) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \))" for H proof - have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r (h \\<^sub>p rm_vars (set X) (\ \\<^sub>s \)) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r (h \\<^sub>p rm_vars (set X) (\ \\<^sub>s \))" for h proof - obtain s t where h: "h = (s,t)" by (metis surj_pair) show ?thesis using 5[of s] 5[of t] unfolding h by fast qed thus ?thesis unfolding subst_apply_pairs_def by auto qed show ?thesis using 5 6 unfolding NegChecks by simp qed (use 2 in auto) qed lemma fv\<^sub>s\<^sub>s\<^sub>t_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) = fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" by (induct S) (auto simp add: fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_ground_subst_compose[OF assms] fv\<^sub>s\<^sub>s\<^sub>t_Cons subst_sst_cons) lemma stateful_strand_step_subst_comp: assumes "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) = {}" shows "x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = (x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" proof (cases x) case (NegChecks X F G) hence *: "range_vars \ \ set X = {}" using assms by simp have "H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \) = (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" for H using pairs_subst_comp rm_vars_comp[OF *] by (induct H) (auto simp add: subst_apply_pairs_def) thus ?thesis using NegChecks by simp qed simp_all lemma stateful_strand_subst_comp: assumes "range_vars \ \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \" using assms proof (induction S) case (Cons s S) hence IH: "S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \" using Cons by auto have "s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using Cons.prems stateful_strand_step_subst_comp[of \ s \] unfolding range_vars_alt_def by auto thus ?case using IH by (simp add: subst_apply_stateful_strand_def) qed simp lemma subst_apply_bvars_disj_NegChecks: assumes "set X \ subst_domain \ = {}" shows "NegChecks X F G \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = NegChecks X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - have "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto thus ?thesis by simp qed lemma subst_apply_NegChecks_no_bvars[simp]: "\[]\\\: F \\: F'\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\: (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)\" "\[]\\\: [] \\: F'\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: [] \\: (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)\" "\[]\\\: F \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\: []\" "\[]\\\: [] \\: [(t,s)]\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: [] \\: ([(t \ \,s \ \)])\" "\[]\\\: [(t,s)] \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: ([(t \ \,s \ \)]) \\: []\" by simp_all lemma setops\<^sub>s\<^sub>s\<^sub>t_mono: "set M \ set N \ setops\<^sub>s\<^sub>s\<^sub>t M \ setops\<^sub>s\<^sub>s\<^sub>t N" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "setops\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_cons[simp]: "setops\<^sub>s\<^sub>s\<^sub>t (a#A) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ setops\<^sub>s\<^sub>s\<^sub>t A" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_cons_subset[simp]: "setops\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>s\<^sub>s\<^sub>t (a#A)" using setops\<^sub>s\<^sub>s\<^sub>t_cons[of a A] by blast lemma setops\<^sub>s\<^sub>s\<^sub>t_append: "setops\<^sub>s\<^sub>s\<^sub>t (A@B) = setops\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>s\<^sub>s\<^sub>t B" proof (induction A) case (Cons a A) thus ?case by (cases a) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_member_iff: "(t,s) \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ (x = Insert t s \ x = Delete t s \ (\ac. x = InSet ac t s) \ (\X F F'. x = NegChecks X F F' \ (t,s) \ set F'))" by (cases x) auto lemma setops\<^sub>s\<^sub>s\<^sub>t_member_iff: "(t,s) \ setops\<^sub>s\<^sub>s\<^sub>t A \ (Insert t s \ set A \ Delete t s \ set A \ (\ac. InSet ac t s \ set A) \ (\X F F'. NegChecks X F F' \ set A \ (t,s) \ set F'))" (is "?P \ ?Q") proof (induction A) case (Cons a A) thus ?case proof (cases "(t, s) \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case True thus ?thesis using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_member_iff[of t s a] by auto qed auto qed simp lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" shows "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (NegChecks X F G) hence "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = set G \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using NegChecks image_Un by simp_all thus ?thesis by (simp add: subst_apply_pairs_def) qed simp_all lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst': assumes "\is_NegChecks a" shows "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms by (cases a) auto lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'': fixes t::"('a,'b) term \ ('a,'b) term" and \::"('a,'b) subst" assumes t: "t \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "\s \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \\<^sub>p rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" proof (cases "is_NegChecks b") case True - then obtain X F G where b: "b = NegChecks X F G" by (cases b) moura+ + then obtain X F G where b: "b = NegChecks X F G" by (cases b) auto hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b = set G" "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \)" by simp_all thus ?thesis using t subst_apply_pairs_pset_subst[of G] by blast next case False hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>p\<^sub>s\<^sub>e\<^sub>t rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst' bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks by fastforce thus ?thesis using t by blast qed lemma setops\<^sub>s\<^sub>s\<^sub>t_subst: assumes "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons a S) have "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" using Cons.prems by auto hence IH: "setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using Cons.IH by auto show ?case using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *] IH unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: subst_apply_stateful_strand_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_subst': fixes p::"('a,'b) term \ ('a,'b) term" and \::"('a,'b) subst" assumes "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\s \ setops\<^sub>s\<^sub>s\<^sub>t S. \X. set X \ bvars\<^sub>s\<^sub>s\<^sub>t S \ p = s \\<^sub>p rm_vars (set X) \" using assms proof (induction S) case (Cons a S) note 0 = setops\<^sub>s\<^sub>s\<^sub>t_cons[of a S] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of a S] note 1 = setops\<^sub>s\<^sub>s\<^sub>t_cons[of "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "S \\<^sub>s\<^sub>s\<^sub>t \"] subst_sst_cons[of a S \] have "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ p \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems 1 by auto thus ?case proof assume *: "p \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" show ?thesis using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF *] 0 by blast next assume *: "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" show ?thesis using Cons.IH[OF *] 0 by blast qed qed simp subsection \Stateful Constraint Semantics\ context intruder_model begin definition negchecks_model where "negchecks_model (\::('a,'b) subst) (D::('a,'b) dbstate) X F G \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ \ \\<^sub>s \ \ s \ \ \\<^sub>s \) \ (\(t,s) \ set G. (t,s) \\<^sub>p \ \\<^sub>s \ \ D))" fun strand_sem_stateful:: "('fun,'var) terms \ ('fun,'var) dbstate \ ('fun,'var) stateful_strand \ ('fun,'var) subst \ bool" ("\_; _; _\\<^sub>s") where "\M; D; []\\<^sub>s = (\\. True)" | "\M; D; Send ts#S\\<^sub>s = (\\. (\t \ set ts. M \ t \ \) \ \M; D; S\\<^sub>s \)" | "\M; D; Receive ts#S\\<^sub>s = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S\\<^sub>s \)" | "\M; D; Equality _ t t'#S\\<^sub>s = (\\. t \ \ = t' \ \ \ \M; D; S\\<^sub>s \)" | "\M; D; Insert t s#S\\<^sub>s = (\\. \M; insert ((t,s) \\<^sub>p \) D; S\\<^sub>s \)" | "\M; D; Delete t s#S\\<^sub>s = (\\. \M; D - {(t,s) \\<^sub>p \}; S\\<^sub>s \)" | "\M; D; InSet _ t s#S\\<^sub>s = (\\. (t,s) \\<^sub>p \ \ D \ \M; D; S\\<^sub>s \)" | "\M; D; NegChecks X F F'#S\\<^sub>s = (\\. negchecks_model \ D X F F' \ \M; D; S\\<^sub>s \)" lemmas strand_sem_stateful_induct = strand_sem_stateful.induct[case_names Nil ConsSnd ConsRcv ConsEq ConsIns ConsDel ConsIn ConsNegChecks] abbreviation constr_sem_stateful (infix "\\<^sub>s" 91) where "\ \\<^sub>s A \ \{}; {}; A\\<^sub>s \" lemma stateful_strand_sem_NegChecks_no_bvars: "\M; D; [\t not in s\]\\<^sub>s \ \ (t \ \, s \ \) \ D" "\M; D; [\t != s\]\\<^sub>s \ \ t \ \ \ s \ \" by (simp_all add: negchecks_model_def empty_dom_iff_empty_subst) lemma strand_sem_ik_mono_stateful: "\M; D; A\\<^sub>s \ \ \M \ M'; D; A\\<^sub>s \" proof (induction A arbitrary: M M' D rule: strand_sem_stateful.induct) case (2 M D ts S) hence "\t \ set ts. M \ t \ \" "\M \ M'; D; S\\<^sub>s \" by auto thus ?case using ideduct_mono[of M _ "M \ M'"] strand_sem_stateful.simps(2)[of "M \ M'" D ts S] by fastforce next case (3 M D ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M) \ M'; D; S\\<^sub>s \" by simp hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ M'); D; S\\<^sub>s \" by (metis Un_assoc) thus ?case by simp qed simp_all lemma strand_sem_append_stateful: "\M; D; A@B\\<^sub>s \ \ \M; D; A\\<^sub>s \ \ \M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \); dbupd\<^sub>s\<^sub>s\<^sub>t A \ D; B\\<^sub>s \" (is "?P \ ?Q \ ?R") proof - have 1: "?P \ ?Q" by (induct A rule: strand_sem_stateful.induct) auto have 2: "?P \ ?R" proof (induction A arbitrary: M D B) case (Cons a A) thus ?case proof (cases a) case (Receive ts) have "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)) = M \ (ik\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t A \ D = dbupd\<^sub>s\<^sub>s\<^sub>t (a#A) \ D" using Receive by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Cons Receive by (metis (no_types, lifting) Un_assoc append_Cons strand_sem_stateful.simps(3)) qed (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) have 3: "?Q \ ?R \ ?P" proof (induction A arbitrary: M D) case (Cons a A) thus ?case proof (cases a) case (Receive ts) have "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)) = M \ (ik\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t A \ D = dbupd\<^sub>s\<^sub>s\<^sub>t (a#A) \ D" using Receive by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Cons Receive by (metis (no_types, lifting) Un_assoc append_Cons strand_sem_stateful.simps(3)) qed (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis by (metis 1 2 3) qed lemma negchecks_model_db_subset: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "D' \ D" and "negchecks_model \ D X F F'" shows "negchecks_model \ D' X F F'" proof - have "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D'" when "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D" for \::"('a,'b) subst" using that assms(1) by blast thus ?thesis using assms(2) unfolding negchecks_model_def by meson qed lemma negchecks_model_db_supset: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "D' \ D" and "\f \ set F'. \\. subst_domain \ = set X \ ground (subst_range \) \ f \\<^sub>p (\ \\<^sub>s \) \ D - D'" and "negchecks_model \ D' X F F'" shows "negchecks_model \ D X F F'" proof - have "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D" when "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D'" "subst_domain \ = set X \ ground (subst_range \)" for \::"('a,'b) subst" using that assms(1,2) by blast thus ?thesis using assms(3) unfolding negchecks_model_def by meson qed lemma negchecks_model_subst: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" shows "negchecks_model (\ \\<^sub>s \) D X F F' \ negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - have 0: "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" when \: "subst_domain \ = set X" "ground (subst_range \)" for \ by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using * unfolding case_prod_unfold by (induct F) auto hence "(fst f \ \) \ \ \\<^sub>s \ \ (snd f \ \) \ \ \\<^sub>s \" using 0[OF \] by simp moreover have "(fst f \ \, snd f \ \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" using f(1) by fastforce } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" obtain f where f: "f \ set F'" "f \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" using * by (induct F') auto hence "f \\<^sub>p \ \\<^sub>p \ \\<^sub>s \ \ D" using 0[OF \] by (metis subst_pair_compose) moreover have "f \\<^sub>p \ \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). (s,t) \\<^sub>p \ \\<^sub>s \ \ D" using f(1) by (metis (no_types, lifting) case_prodI2) } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" obtain f where f: "f \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) auto then obtain g where g: "g \ set F" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "fst g \ \ \\<^sub>s (\ \\<^sub>s \) \ snd g \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) g 0[OF \] by (simp add: prod.case_eq_if) hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using g by fastforce } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). (s,t) \\<^sub>p (\ \\<^sub>s \) \ D" obtain f where f: "f \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "f \\<^sub>p \ \\<^sub>s \ \ D" using * by (induct F') auto then obtain g where g: "g \ set F'" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "g \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" using f(2) g 0[OF \] by (simp add: prod.case_eq_if) hence "\(s,t) \ set F'. (s,t) \\<^sub>p (\ \\<^sub>s (\ \\<^sub>s \)) \ D" using g by (metis (no_types, lifting) case_prodI2) } ultimately show ?thesis using assms unfolding negchecks_model_def by meson qed lemma strand_sem_subst_stateful: fixes \::"('fun,'var) subst" assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "\M; D; S\\<^sub>s (\ \\<^sub>s \) \ \M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" proof note [simp] = subst_sst_cons[of _ _ \] subst_subst_compose[of _ \ \] have "(subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" when \: "(subst_domain \ \ range_vars \) \ set X = {}" and \: "subst_domain \ = set X" "ground (subst_range \)" for X and \::"('fun,'var) subst" using \ \ unfolding range_vars_alt_def by auto hence 0: "\ \\<^sub>s \ = \ \\<^sub>s \" when \: "(subst_domain \ \ range_vars \) \ set X = {}" and \: "subst_domain \ = set X" "ground (subst_range \)" for \ X by (metis \ \ subst_comp_eq_if_disjoint_vars) show "\M; D; S\\<^sub>s (\ \\<^sub>s \) \ \M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" using assms proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence "\t \ set ts. M \ t \ \ \ \" and IH: "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by auto hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp thus ?case using IH by simp next case (ConsRcv M D ts S) hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by simp hence "\(set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by (metis list.set_map subst_comp_all) thus ?case by simp next case (ConsNegChecks M D X F F' S) hence *: "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" and **: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def negchecks_model_def by (force, auto) have "negchecks_model (\ \\<^sub>s \) D X F F'" using ConsNegChecks by auto hence "negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using 0[OF **] negchecks_model_subst[OF **] by blast moreover have "rm_vars (set X) \ = \" using ConsNegChecks.prems(2) by force ultimately show ?case using * by auto qed simp_all show "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \ \ \M; D; S\\<^sub>s (\ \\<^sub>s \)" using assms proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" and IH: "\M; D; S\\<^sub>s (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \ t \ \ \ \" by simp thus ?case using IH by simp next case (ConsRcv M D ts S) hence "\(set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S\\<^sub>s (\ \\<^sub>s \)" by simp hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; D; S\\<^sub>s (\ \\<^sub>s \)" by (metis list.set_map subst_comp_all) thus ?case by simp next case (ConsNegChecks M D X F F' S) have \: "rm_vars (set X) \ = \" using ConsNegChecks.prems(2) by force hence *: "\M; D; S\\<^sub>s (\ \\<^sub>s \)" and **: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsNegChecks unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def negchecks_model_def by auto have "negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsNegChecks.prems(1) \ by (auto simp add: subst_compose_assoc negchecks_model_def) hence "negchecks_model (\ \\<^sub>s \) D X F F'" using 0[OF **] negchecks_model_subst[OF **] by blast thus ?case using * by auto qed simp_all qed lemma strand_sem_receive_prepend_stateful: assumes "\M; D; S\\<^sub>s \" and "list_all is_Receive S'" shows "\M; D; S@S'\\<^sub>s \" using assms(2) proof (induction S' rule: List.rev_induct) case (snoc x S') hence "\t. x = receive\t\" "\M; D; S@S'\\<^sub>s \" unfolding list_all_iff is_Receive_def by auto thus ?case using strand_sem_append_stateful[of M D "S@S'" "[x]" \] by auto qed (simp add: assms(1)) lemma negchecks_model_model_swap: fixes I J::"('a,'b) subst" assumes "\x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) - set X. I x = J x" and "negchecks_model I D X F G" shows "negchecks_model J D X F G" proof - have 0: "\x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G). (\ \\<^sub>s I) x = (\ \\<^sub>s J) x" when "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" using that assms(1) by (metis DiffI ground_subst_range_empty_fv subst_comp_notin_dom_eq subst_ground_ident_compose(1)) have 1: "s \ (\ \\<^sub>s J) \ t \ (\ \\<^sub>s J)" when "s \ (\ \\<^sub>s I) \ t \ (\ \\<^sub>s I)" "(s,t) \ set F" "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" and s t using that(1,2) 0[OF that(3,4)] term_subst_eq_conv[of _ "\ \\<^sub>s I" "\ \\<^sub>s J"] UnCI fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4,5)[OF that(2)] by metis have 2: "(s,t) \\<^sub>p (\ \\<^sub>s J) \ D" when "(s,t) \\<^sub>p (\ \\<^sub>s I) \ D" "(s,t) \ set G" "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" and s t using that(1,2) 0[OF that(3,4)] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4,5)[of s t G] term_subst_eq_conv[of s "\ \\<^sub>s I" "\ \\<^sub>s J"] term_subst_eq_conv[of t "\ \\<^sub>s I" "\ \\<^sub>s J"] by simp have 3: "(\(s,t) \ set F. s \ \ \\<^sub>s J \ t \ \ \\<^sub>s J) \ (\(s,t) \ set G. (s, t) \\<^sub>p \ \\<^sub>s J \ D)" when "subst_domain \ = set X" "ground (subst_range \)" "(\(s,t) \ set F. s \ \ \\<^sub>s I \ t \ \ \\<^sub>s I) \ (\(s,t) \ set G. (s, t) \\<^sub>p \ \\<^sub>s I \ D)" for \::"('a,'b) subst" using 1[OF _ _ that(1,2)] 2[OF _ _ that(1,2)] that(3) by blast thus ?thesis using assms(2) unfolding negchecks_model_def by simp qed lemma strand_sem_model_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t S. I x = J x" and "\M; D; S\\<^sub>s I" shows "\M; D; S\\<^sub>s J" using assms(2,1) proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence *: "\M; D; S\\<^sub>s J" "\t \ set ts. M \ t \ I" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). I x = J x" by (fastforce, fastforce, fastforce) have "\t \ set ts. M \ t \ J" using *(2,3) term_subst_eq_conv[of _ I J] by (metis fv_subset subsetD) thus ?case using *(1) by auto next case (ConsRcv M D ts S) hence *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t I) \ M; D; S\\<^sub>s J" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). I x = J x" by (fastforce, fastforce) have "set ts \\<^sub>s\<^sub>e\<^sub>t I = set ts \\<^sub>s\<^sub>e\<^sub>t J" using *(2) term_subst_eq_conv[of _ I J] by (meson fv_subset image_cong subsetD) thus ?case using *(1) by simp next case (ConsEq M D ac t t' S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of t' I J] by force next case (ConsIns M D t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsDel M D t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsIn M D uv t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsNegChecks M D X F F' S) hence "\M; D; S\\<^sub>s J" "negchecks_model I D X F F'" "\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X. I x = J x" by (fastforce, fastforce, fastforce) thus ?case using negchecks_model_model_swap[of F F' X I J D] by fastforce qed simp lemma strand_sem_receive_send_append: assumes A: "\M; D; A\\<^sub>s I" shows "\M; D; A@[receive\[t]\,send\[t]\]\\<^sub>s I" proof - have "M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[receive\[t]\]) \\<^sub>s\<^sub>e\<^sub>t I) \ t \ I" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of t "A@[receive\[t]\]"] by auto thus ?thesis using A strand_sem_append_stateful[of M D A "[receive\[t]\]" I] strand_sem_append_stateful[of M D "A@[receive\[t]\]" "[send\[t]\]" I] by force qed lemma strand_sem_stateful_if_no_send_or_check: assumes A: "list_all (\a. \is_Send a \ \is_Check_or_Assignment a) A" shows "\M; D; A\\<^sub>s I" using A proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" and a: "\is_Send a" "\is_Check_or_Assignment a" by auto from a have "\M; D; [a]\\<^sub>s I" for M D by (cases a) auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] by blast qed simp lemma strand_sem_stateful_if_sends_deduct: assumes "list_all is_Send A" and "\ts. send\ts\ \ set A \ (\t \ set ts. M \ t \ I)" shows "\M; D; A\\<^sub>s I" using assms proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" by auto obtain ts where a: "a = send\ts\" "\t \ set ts. M \ t \ I" using snoc.prems by (cases a) auto have "\M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I); D; [a]\\<^sub>s I" for D using ideduct_mono[of M _ "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)"] a by auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] by blast qed simp lemma strand_sem_stateful_if_checks: assumes "list_all is_Check_or_Assignment A" and "\ac t s. \ac: t \ s\ \ set A \ t \ I = s \ I" and "\ac t s. \ac: t \ s\ \ set A \ (t \ I, s \ I) \ D" and "\X F G. \X\\\: F \\: G\ \ set A \ negchecks_model I D X F G" shows "\M; D; A\\<^sub>s I" using assms proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" and a: "is_Check_or_Assignment a" by auto have 0: "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D" using snoc.prems(1) dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of A I D] unfolding list_all_iff by auto have "\M; D; [a]\\<^sub>s I" for M using a snoc.prems(2,3,4) by (cases a) auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] unfolding 0 by blast qed simp lemma strand_sem_stateful_sends_deduct: assumes A: "\M; D; A\\<^sub>s I" and ts: "send\ts\ \ set A" and t: "t \ set ts" shows "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ t \ I" using A ts proof (induction A arbitrary: M D rule: List.rev_induct) case (snoc a A) have 0: "\M; D; A\\<^sub>s I" using strand_sem_append_stateful snoc.prems(1) by fast have 1: "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t I)" by auto have 2: "M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[send\ts\]) \\<^sub>s\<^sub>e\<^sub>t I) = M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ A] in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "A@[send\ts\]"] by fastforce show ?case proof (cases "send\ts\ \ set A") case True show ?thesis by (rule ideduct_mono[OF snoc.IH[OF 0 True] 1]) next case False hence a: "a = send\ts\" using snoc.prems(2) by force show ?thesis using strand_sem_append_stateful[of M D A "[a]" I] snoc.prems(1) t unfolding a 2 by auto qed qed simp end subsection \Well-Formedness Lemmata\ lemma wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[simp]: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" by (induction S) (auto simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def split: stateful_strand_step.split poscheckvariant.split) lemma wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_append: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (S@S') = wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S'" by (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_union[simp]: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (S@T) = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t T" by (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_singleton: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t [s] = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" by (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_fv_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ S] unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by force lemma wf\<^sub>s\<^sub>s\<^sub>t_prefix[dest]: "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S') \ wf'\<^sub>s\<^sub>s\<^sub>t V S" by (induct S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) auto lemma wf\<^sub>s\<^sub>s\<^sub>t_vars_mono: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ W) S" proof (induction S arbitrary: V) case (Cons x S) thus ?case proof (cases x) case (Send ts) have "wf'\<^sub>s\<^sub>s\<^sub>t (V \ W \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" using Cons.prems(1) Cons.IH Send by (metis sup_assoc sup_commute wf'\<^sub>s\<^sub>s\<^sub>t.simps(3)) thus ?thesis by (simp add: Send) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ W) S" "fv t' \ V \ W" using Equality Cons.prems(1) Cons.IH by auto thus ?thesis using Equality Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using Equality Cons by auto qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ W) S" using InSet Cons.prems(1) Cons.IH by auto thus ?thesis using InSet Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using InSet Cons by auto qed qed auto qed simp lemma wf\<^sub>s\<^sub>s\<^sub>tI[intro]: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) = V" using Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using Send by simp next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t' \ V" using Equality Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Equality Assign by simp next case Check thus ?thesis using Equality Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t \ fv t' \ V" using InSet Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono InSet Assign by (simp add: Un_assoc) next case Check thus ?thesis using InSet Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed (simp_all add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>tI'[intro]: assumes "\((\x. case x of Receive ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign _ t' \ fv t' | Insert t t' \ fv t \ fv t' | _ \ {}) ` set S) \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V S" using assms proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Equality a t t') thus ?thesis using Cons by (cases a) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) next case (InSet a t t') thus ?thesis using Cons by (cases a) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Un_assoc) qed (simp_all add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_append_exec: "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S') \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" proof (induction S arbitrary: V) case (Cons x S V) thus ?case proof (cases x) case (Send ts) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Cons.prems Cons.IH by simp thus ?thesis using Send unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Assign unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case Check hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Check unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using InSet Cons.prems Cons.IH by auto thus ?thesis using InSet Assign unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case Check hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using InSet Cons.prems Cons.IH by auto thus ?thesis using InSet Check unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) qed qed (auto simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_append: "wf'\<^sub>s\<^sub>s\<^sub>t X S \ wf'\<^sub>s\<^sub>s\<^sub>t Y T \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (S@T)" proof (induction X S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case 1 thus ?case by (metis wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Un_commute append_Nil) next case 3 thus ?case by (metis append_Cons Un_commute Un_assoc wf'\<^sub>s\<^sub>s\<^sub>t.simps(3)) next case (4 V t t' S) hence *: "fv t' \ V" and "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ Y) (S @ T)" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ Y \ fv t) (S @ T)" by (metis Un_commute Un_assoc) thus ?case using * by auto next case (8 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ Y) (S @ T)" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ Y \ fv t \ fv t') (S @ T)" by (metis Un_commute Un_assoc) thus ?case by auto qed auto lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" proof (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case (2 V ts S) hence *: "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" "wf'\<^sub>s\<^sub>s\<^sub>t V S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V" using "2.prems"(2) unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce thus ?case using "2.IH" * by simp next case (3 V ts S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using "3.prems"(2) unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?case using "3.IH" * by simp next case (4 V t t' S) hence *: "fv t' \ V" "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t := t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (\t := t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv t)" using "4.prems"(2) by blast thus ?case using "4.IH" * by simp next case (6 V t t' S) hence *: "fv t \ fv t' \ V" "wf'\<^sub>s\<^sub>s\<^sub>t V S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (insert\t,t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V" using "6.prems"(2) by blast thus ?case using "6.IH" * by simp next case (8 V t t' S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t') S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (select\t,t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (select\t,t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv t \ fv t')" using "8.prems"(2) by blast thus ?case using "8.IH" * by simp qed (simp_all add: wf\<^sub>s\<^sub>s\<^sub>tI wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix': assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "\((\x. case x of Receive ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign _ t' \ fv t' | Insert t t' \ fv t \ fv t' | _ \ {}) ` set S') \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" using assms by (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>tI' wf\<^sub>s\<^sub>s\<^sub>t_vars_mono wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_subst_apply: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction S arbitrary: V rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case (2 V ts S) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "2.IH" subst_apply_fv_subset by (simp, force) thus ?case by (simp add: subst_apply_stateful_strand_def) next case (3 V ts S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)))) (S \\<^sub>s\<^sub>s\<^sub>t \)" using "3.IH" by metis hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" using subst_apply_fv_unfold_set[of \ ts] by fastforce thus ?case by (simp add: subst_apply_stateful_strand_def) next case (4 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S" "fv t' \ V" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>s\<^sub>t \)" and *: "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "4.IH" subst_apply_fv_subset by force+ hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_union) thus ?case using * by (simp add: subst_apply_stateful_strand_def) next case (6 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t \ fv t' \ V" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "6.IH" subst_apply_fv_subset by force+ thus ?case by (simp add: sup_assoc subst_apply_stateful_strand_def) next case (8 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t') S" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t \ fv t'))) (S \\<^sub>s\<^sub>s\<^sub>t \)" using "8.IH" subst_apply_fv_subset by force hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_union) thus ?case by (simp add: subst_apply_stateful_strand_def) qed (auto simp add: subst_apply_stateful_strand_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_induct[consumes 1, case_names Nil ConsSnd ConsRcv ConsEq ConsEq2 ConsIn ConsIns ConsDel ConsNegChecks]: fixes S::"('a,'b) stateful_strand" assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" "P []" "\ts S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Send ts])" "\ts S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Receive ts])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Equality Assign t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Equality Check t t'])" "\ac t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[InSet ac t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv t \ fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Insert t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Delete t t'])" "\X F G S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[NegChecks X F G])" shows "P S" using assms proof (induction S rule: List.rev_induct) case (snoc x S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t V S" "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) [x]" by (metis wf\<^sub>s\<^sub>s\<^sub>t_prefix, metis wf\<^sub>s\<^sub>s\<^sub>t_append_exec) have IH: "P S" using snoc.IH[OF *(1)] snoc.prems by auto note ** = snoc.prems(3-)[OF *(1) IH] *(2) show ?case proof (cases x) case (Send ts) thus ?thesis using **(1) by blast next case (Receive ts) thus ?thesis using **(2,9) by simp next case (Equality ac t t') thus ?thesis using **(3,4,9) by (cases ac) auto next case (Insert t t') thus ?thesis using **(6,9) by force next case (Delete t t') thus ?thesis using **(7) by presburger next case (NegChecks X F G) thus ?thesis using **(8) by presburger next case (InSet ac t t') thus ?thesis using **(5) by (cases ac) auto qed qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_strand_first_Send_var_split: assumes "wf'\<^sub>s\<^sub>s\<^sub>t {} S" "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. \(\w \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ \ \ w) \ ( (\ts. S = S\<^sub>p\<^sub>r\<^sub>e@send\ts\#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (\s u. S = S\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ s \ \ \ \(t \ \ \\<^sub>s\<^sub>e\<^sub>t \ ` fv u)) \ (\s u. S = S\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#S\<^sub>s\<^sub>u\<^sub>f \ (t \ \ \ s \ \ \ t \ \ \ u \ \)))" (is "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. ?P S\<^sub>p\<^sub>r\<^sub>e \ ?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f") using assms proof (induction S rule: wf\<^sub>s\<^sub>s\<^sub>t_induct) case (ConsSnd ts' S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ w") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsSnd.IH by moura + using ConsSnd.IH by atomize_elim auto thus ?thesis by fastforce next case False then obtain v where v: "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts')" "t \ \ \ \ v" using ConsSnd.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain t' where t': "t' \ set ts'" "v \ fv t'" by auto have "t \ \ \ t' \ \" using v(2) t'(2) subst_mono[of "Var v" t' \] vars_iff_subtermeq[of v] term.order_trans by auto hence "t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts' \\<^sub>s\<^sub>e\<^sub>t \" using v(1) t'(1) by blast thus ?thesis using False v by blast qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsRcv.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsRcv.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsRcv.IH by moura + using ConsRcv.IH by atomize_elim auto thus ?case by fastforce next case (ConsEq s s' S) have *: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsEq.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsEq.IH by moura + using ConsEq.IH by atomize_elim auto thus ?thesis by fastforce next case False then obtain v where "v \ fv s" "t \ \ \ \ v" and **: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsEq.prems * unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto hence "t \ \ \ s \ \" using vars_iff_subtermeq[of v s] subst_mono[of "Var v" s \] term.order_trans by auto thus ?thesis using False ** by fastforce qed next case (ConsEq2 s s' S) have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (S@[Equality Check s s']) = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsEq2.prems by metis - then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq2.IH by moura + then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq2.IH by atomize_elim auto thus ?case by fastforce next case (ConsNegChecks X F G S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp - then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsNegChecks.IH by moura + then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsNegChecks.IH by atomize_elim auto thus ?case by fastforce next case (ConsIn ac s s' S) show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsIn.IH by moura + using ConsIn.IH by atomize_elim auto thus ?thesis by fastforce next case False hence ac: "ac = assign" using ConsIn.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by (cases ac) auto obtain v where "v \ fv s \ fv s'" "t \ \ \ \ v" using ConsIn.prems False unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def ac by auto hence *: "t \ \ \ s \ \ \ t \ \ \ s' \ \" using vars_iff_subtermeq[of v s] vars_iff_subtermeq[of v s'] subst_mono[of "Var v" s \] subst_mono[of "Var v" s' \] term.order_trans by auto show ?thesis using * False unfolding ac by fast qed next case (ConsIns s s' S) have "fv s \ fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsIns.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsIns.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce - then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIns.IH by moura + then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIns.IH by atomize_elim auto thus ?case by fastforce next case (ConsDel s s' S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp - then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsDel.IH by moura + then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsDel.IH by atomize_elim auto thus ?case by fastforce qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_vars_mono': "wf'\<^sub>s\<^sub>s\<^sub>t V S \ V \ W \ wf'\<^sub>s\<^sub>s\<^sub>t W S" by (metis Diff_partition wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_receives_only_eq: assumes "list_all is_Receive S" shows "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>s\<^sub>t S" using assms proof (induction S) case (Cons a A) obtain ts where a: "a = receive\ts\" using Cons.prems by (cases a) auto have IH: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A = fv\<^sub>s\<^sub>s\<^sub>t A" using Cons.prems Cons.IH by simp show ?case using IH unfolding a wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_receives_only_empty: assumes "list_all is_Receive S" shows "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S = {}" using assms proof (induction S) case (Cons a A) obtain ts where a: "a = receive\ts\" using Cons.prems by (cases a) auto have IH: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t A = {}" using Cons.prems Cons.IH by simp show ?case using IH unfolding a wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by simp qed (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_sends_only: assumes "list_all is_Send S" shows "wf'\<^sub>s\<^sub>s\<^sub>t V S" using assms proof (induction S arbitrary: V) case (Cons s S) thus ?case by (cases s) auto qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_sends_only_prepend: assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "list_all is_Send S'" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S'@S)" using wf\<^sub>s\<^sub>s\<^sub>t_append[OF wf\<^sub>s\<^sub>s\<^sub>t_sends_only[OF assms(2), of "{}"] assms(1)] by simp lemma wf\<^sub>s\<^sub>s\<^sub>t_receives_only_fv_subset: assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "list_all is_Receive S" shows "fv\<^sub>s\<^sub>s\<^sub>t S \ V" using assms proof (induction rule: wf\<^sub>s\<^sub>s\<^sub>t_induct) case (ConsRcv ts S) thus ?case using wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_receives_only_empty[of S] by auto qed auto lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix'': assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" using assms by (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>tI' wf\<^sub>s\<^sub>s\<^sub>t_vars_mono wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Typing.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Typing.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Typing.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Typing.thy @@ -1,2010 +1,2010 @@ (* Title: Stateful_Typing.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Extending the Typing Result to Stateful Constraints\ theory Stateful_Typing imports Typing_Result Stateful_Strands begin text \\label{sec:Stateful-Typing}Locale setup\ locale stateful_typed_model = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" + fixes Pair::"'fun" assumes Pair_arity: "arity Pair = 2" and Ana_subst': "\f T \ K M. Ana (Fun f T) = (K,M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" begin lemma Ana_invar_subst'[simp]: "Ana_invar_subst \" using Ana_subst' unfolding Ana_invar_subst_def by force definition pair where "pair d \ case d of (t,t') \ Fun Pair [t,t']" fun tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s:: "(('fun,'var) term \ ('fun,'var) term) list \ ('fun,'var) dbstatelist \ (('fun,'var) term \ ('fun,'var) term) list list" where "tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s [] D = [[]]" | "tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) D = concat (map (\d. map ((#) (pair (s,t), pair d)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)) D)" text \ A translation/reduction \tr\ from stateful constraints to (lists of) "non-stateful" constraints. The output represents a finite disjunction of constraints whose models constitute exactly the models of the input constraint. The typing result for "non-stateful" constraints is later lifted to the stateful setting through this reduction procedure. \ fun tr::"('fun,'var) stateful_strand \ ('fun,'var) dbstatelist \ ('fun,'var) strand list" where "tr [] D = [[]]" | "tr (send\ts\#A) D = map ((#) (send\ts\\<^sub>s\<^sub>t)) (tr A D)" | "tr (receive\ts\#A) D = map ((#) (receive\ts\\<^sub>s\<^sub>t)) (tr A D)" | "tr (\ac: t \ t'\#A) D = map ((#) (\ac: t \ t'\\<^sub>s\<^sub>t)) (tr A D)" | "tr (insert\t,s\#A) D = tr A (List.insert (t,s) D)" | "tr (delete\t,s\#A) D = concat (map (\Di. map (\B. (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])@B) (tr A [d\D. d \ set Di])) (subseqs D))" | "tr (\ac: t \ s\#A) D = concat (map (\B. map (\d. \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#B) D) (tr A D))" | "tr (\X\\\: F \\: F'\#A) D = map ((@) (map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))) (tr A D)" text \Type-flaw resistance of stateful constraint steps\ fun tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = ((\\. Unifier \ t t') \ \ t = \ t')" | "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks X F F') = ( (F' = [] \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F-set X. \a. \ (Var x) = TAtom a)) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') \ T = [] \ (\s \ set T. s \ Var ` set X)))" | "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = True" text \Type-flaw resistance of stateful constraints\ definition tfr\<^sub>s\<^sub>s\<^sub>t where "tfr\<^sub>s\<^sub>s\<^sub>t S \ tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S) \ list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p S" subsection \Minor Lemmata\ lemma pair_in_pair_image_iff: "pair (s,t) \ pair ` P \ (s,t) \ P" unfolding pair_def by fast lemma subst_apply_pairs_pair_image_subst: "pair ` set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = pair ` set F \\<^sub>s\<^sub>e\<^sub>t \" unfolding subst_apply_pairs_def pair_def by (induct F) auto lemma Ana_subst_subterms_cases: fixes \::"('fun,'var) subst" assumes t: "t \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" and s: "s \ set (snd (Ana t))" shows "(\u \ subterms\<^sub>s\<^sub>e\<^sub>t M. t = u \ \ \ s \ set (snd (Ana u)) \\<^sub>s\<^sub>e\<^sub>t \) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. t \ \ x)" proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \") case True - then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t M" "t = u \ \" by moura + then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t M" "t = u \ \" by blast show ?thesis proof (cases u) case (Var x) hence "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using fv_subset_subterms[OF u(1)] by simp thus ?thesis using u(2) Var by fastforce next case (Fun f T) hence "set (snd (Ana t)) = set (snd (Ana u)) \\<^sub>s\<^sub>e\<^sub>t \" using Ana_subst'[of f T _ _ \] u(2) by (cases "Ana u") auto thus ?thesis using s u by blast qed qed (use s t subterms\<^sub>s\<^sub>e\<^sub>t_subst in blast) lemma Ana_snd_subst_nth_inv: fixes \::"('fun,'var) subst" and f ts defines "R \ snd (Ana (Fun f ts \ \))" assumes r: "r = R ! i" "i < length R" shows "r = snd (Ana (Fun f ts)) ! i \ \" proof - obtain K R where "Ana (Fun f ts) = (K,R)" by (metis surj_pair) thus ?thesis using Ana_subst'[of f ts K R \] r unfolding R_def by auto qed lemma Ana_snd_subst_inv: fixes \::"('fun,'var) subst" assumes r: "r \ set (snd (Ana (Fun f ts \ \)))" shows "\t \ set (snd (Ana (Fun f ts))). r = t \ \" proof - obtain K R where "Ana (Fun f ts) = (K,R)" by (metis surj_pair) thus ?thesis using Ana_subst'[of f ts K R \] r by auto qed lemma fun_pair_eq[dest]: "pair d = pair d' \ d = d'" proof - - obtain t s t' s' where "d = (t,s)" "d' = (t',s')" by moura + obtain t s t' s' where "d = (t,s)" "d' = (t',s')" by atomize_elim auto thus "pair d = pair d' \ d = d'" unfolding pair_def by simp qed lemma fun_pair_subst: "pair d \ \ = pair (d \\<^sub>p \)" using surj_pair[of d] unfolding pair_def by force lemma fun_pair_subst_set: "pair ` M \\<^sub>s\<^sub>e\<^sub>t \ = pair ` (M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" proof show "pair ` M \\<^sub>s\<^sub>e\<^sub>t \ \ pair ` (M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using fun_pair_subst[of _ \] by fastforce show "pair ` (M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) \ pair ` M \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume t: "t \ pair ` (M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" then obtain p where p: "p \ M" "t = pair (p \\<^sub>p \)" by blast thus "t \ pair ` M \\<^sub>s\<^sub>e\<^sub>t \" using fun_pair_subst[of p \] by force qed qed lemma fun_pair_eq_subst: "pair d \ \ = pair d' \ \ \ d \\<^sub>p \ = d' \\<^sub>p \" by (metis fun_pair_subst fun_pair_eq[of "d \\<^sub>p \" "d' \\<^sub>p \"]) lemma setops\<^sub>s\<^sub>s\<^sub>t_pair_image_cons[simp]: "pair ` setops\<^sub>s\<^sub>s\<^sub>t (x#S) = pair ` setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (send\ts\#S) = pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (receive\ts\#S) = pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\ac: t \ t'\#S) = pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (insert\t,s\#S) = {pair (t,s)} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#S) = {pair (t,s)} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\ac: t \ s\#S) = {pair (t,s)} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\X\\\: F \\: G\#S) = pair ` set G \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by auto lemma setops\<^sub>s\<^sub>s\<^sub>t_pair_image_subst_cons[simp]: "pair ` setops\<^sub>s\<^sub>s\<^sub>t (x#S \\<^sub>s\<^sub>s\<^sub>t \) = pair ` setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (send\ts\#S \\<^sub>s\<^sub>s\<^sub>t \) = pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (receive\ts\#S \\<^sub>s\<^sub>s\<^sub>t \) = pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\ac: t \ t'\#S \\<^sub>s\<^sub>s\<^sub>t \) = pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (insert\t,s\#S \\<^sub>s\<^sub>s\<^sub>t \) = {pair (t,s) \ \} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#S \\<^sub>s\<^sub>s\<^sub>t \) = {pair (t,s) \ \} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\ac: t \ s\#S \\<^sub>s\<^sub>s\<^sub>t \) = {pair (t,s) \ \} \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\X\\\: F \\: G\#S \\<^sub>s\<^sub>s\<^sub>t \) = pair ` set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using subst_sst_cons[of _ S \] unfolding setops\<^sub>s\<^sub>s\<^sub>t_def pair_def by auto lemma setops\<^sub>s\<^sub>s\<^sub>t_are_pairs: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ \s s'. t = pair (s,s')" proof (induction A) case (Cons a A) thus ?case by (cases a) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma fun_pair_wf\<^sub>t\<^sub>r\<^sub>m: "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m t' \ wf\<^sub>t\<^sub>r\<^sub>m (pair (t,t'))" using Pair_arity unfolding wf\<^sub>t\<^sub>r\<^sub>m_def pair_def by auto lemma wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_pairs: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (pair ` set F)" using fun_pair_wf\<^sub>t\<^sub>r\<^sub>m by blast lemma wf_fun_pair_ineqs_map: assumes "wf\<^sub>s\<^sub>t X A" shows "wf\<^sub>s\<^sub>t X (map (\d. \Y\\\: [(pair (t, s), pair d)]\\<^sub>s\<^sub>t) D@A)" using assms by (induct D) auto lemma wf_fun_pair_negchecks_map: assumes "wf\<^sub>s\<^sub>t X A" shows "wf\<^sub>s\<^sub>t X (map (\G. \Y\\\: (F@G)\\<^sub>s\<^sub>t) M@A)" using assms by (induct M) auto lemma wf_fun_pair_eqs_ineqs_map: fixes A::"('fun,'var) strand" assumes "wf\<^sub>s\<^sub>t X A" "Di \ set (subseqs D)" "\(t,t') \ set D. fv t \ fv t' \ X" shows "wf\<^sub>s\<^sub>t X ((map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])@A)" proof - let ?c1 = "map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di" let ?c2 = "map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]" have 1: "wf\<^sub>s\<^sub>t X (?c2@A)" using wf_fun_pair_ineqs_map[OF assms(1)] by simp have 2: "\(t,t') \ set Di. fv t \ fv t' \ X" using assms(2,3) by (meson contra_subsetD subseqs_set_subset(1)) have "wf\<^sub>s\<^sub>t X (?c1@B)" when "wf\<^sub>s\<^sub>t X B" for B::"('fun,'var) strand" using 2 that by (induct Di) auto thus ?thesis using 1 by simp qed lemma trms\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and t: "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\s \. s \ trms\<^sub>s\<^sub>s\<^sub>t S \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" using t proof (induction S) case (Cons s S) thus ?case proof (cases "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False hence "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems trms\<^sub>s\<^sub>s\<^sub>t_subst_cons[of s S \] by auto then obtain u where u: "u \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" "t = u \ rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)) \" using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'' by blast thus ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_subst_cons[of s S \] wt_subst_rm_vars[OF \(1), of "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)"] wf_trms_subst_rm_vars'[OF \(2), of "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)"] by fastforce qed auto qed simp lemma setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\s \. s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" using t proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Insert t' s) hence "t = pair (t',s) \ \ \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.prems subst_sst_cons[of _ S \] unfolding pair_def by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Insert Cons.IH \ by (cases "t = pair (t', s) \ \") (fastforce, auto) next case (Delete t' s) hence "t = pair (t',s) \ \ \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.prems subst_sst_cons[of _ S \] unfolding pair_def by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Delete Cons.IH \ by (cases "t = pair (t', s) \ \") (fastforce, auto) next case (InSet ac t' s) hence "t = pair (t',s) \ \ \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.prems subst_sst_cons[of _ S \] unfolding pair_def by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using InSet Cons.IH \ by (cases "t = pair (t', s) \ \") (fastforce, auto) next case (NegChecks X F F') hence "t \ pair ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.prems subst_sst_cons[of _ S \] unfolding pair_def by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis proof assume "t \ pair ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" then obtain s where s: "t = s \ rm_vars (set X) \" "s \ pair ` set F'" using subst_apply_pairs_pair_image_subst[of F' "rm_vars (set X) \"] by auto thus ?thesis using NegChecks setops\<^sub>s\<^sub>s\<^sub>t_pair_image_cons(8)[of X F F' S] wt_subst_rm_vars[OF \(1), of "set X"] wf_trms_subst_rm_vars'[OF \(2), of "set X"] by fast qed (use Cons.IH in auto) qed (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def subst_sst_cons[of _ S \]) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" proof - show "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) hence 0: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" by auto thus ?case proof (cases a) case (NegChecks X F F') hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F')" using 0 by simp thus ?thesis using NegChecks wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_pairs[of F'] 0 by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def dest: fun_pair_wf\<^sub>t\<^sub>r\<^sub>m) qed (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" by fast qed lemma SMP_MP_split: assumes "t \ SMP M" and M: "\m \ M. is_Fun m" shows "(\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t \ M \\<^sub>s\<^sub>e\<^sub>t \) \ t \ SMP ((subterms\<^sub>s\<^sub>e\<^sub>t M \ \((set \ fst \ Ana) ` M)) - M)" (is "?P t \ ?Q t") using assms(1) proof (induction t rule: SMP.induct) case (MP t) have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "M \\<^sub>s\<^sub>e\<^sub>t Var = M" by simp_all thus ?case using MP by metis next case (Subterm t t') show ?case using Subterm.IH proof assume "?P t" - then obtain s \ where s: "s \ M" "t = s \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by moura + then obtain s \ where s: "s \ M" "t = s \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by auto then obtain f T where fT: "s = Fun f T" using M by fast have "(\s'. s' \ s \ t' = s' \ \) \ (\x \ fv s. t' \ \ x)" using subterm_subst_unfold[OF Subterm.hyps(2)[unfolded s(2)]] by blast thus ?thesis proof assume "\s'. s' \ s \ t' = s' \ \" - then obtain s' where s': "s' \ s" "t' = s' \ \" by moura + then obtain s' where s': "s' \ s" "t' = s' \ \" by fast show ?thesis proof (cases "s' \ M") case True thus ?thesis using s' \ by blast next case False hence "s' \ (subterms\<^sub>s\<^sub>e\<^sub>t M \ \((set \ fst \ Ana) ` M)) - M" using s'(1) s(1) by force thus ?thesis using SMP.Substitution[OF SMP.MP[of s'] \] s' by presburger qed next assume "\x \ fv s. t' \ \ x" - then obtain x where x: "x \ fv s" "t' \ \ x" by moura + then obtain x where x: "x \ fv s" "t' \ \ x" by fast have "Var x \ M" using M by blast hence "Var x \ (subterms\<^sub>s\<^sub>e\<^sub>t M \ \((set \ fst \ Ana) ` M)) - M" using s(1) var_is_subterm[OF x(1)] by blast hence "\ x \ SMP ((subterms\<^sub>s\<^sub>e\<^sub>t M \ \((set \ fst \ Ana) ` M)) - M)" using SMP.Substitution[OF SMP.MP[of "Var x"] \] by auto thus ?thesis using SMP.Subterm x(2) by presburger qed qed (metis SMP.Subterm[OF _ Subterm.hyps(2)]) next case (Substitution t \) show ?case using Substitution.IH proof assume "?P t" - then obtain \ where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ M \\<^sub>s\<^sub>e\<^sub>t \" by moura + then obtain \ where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ M \\<^sub>s\<^sub>e\<^sub>t \" by fast hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \)" using wt_subst_compose[of \, OF _ Substitution.hyps(2)] wf_trm_subst_compose[of \ _ \, OF _ wf_trm_subst_rangeD[OF Substitution.hyps(3)]] wf_trm_subst_range_iff by (argo, blast, auto) thus ?thesis by blast next assume "?Q t" thus ?thesis using SMP.Substitution[OF _ Substitution.hyps(2,3)] by meson qed next case (Ana t K T k) show ?case using Ana.IH proof assume "?P t" - then obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ M \\<^sub>s\<^sub>e\<^sub>t \" by moura + then obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ M \\<^sub>s\<^sub>e\<^sub>t \" by fast then obtain s where s: "s \ M" "t = s \ \" by auto then obtain f S where fT: "s = Fun f S" using M by (cases s) auto obtain K' T' where s_Ana: "Ana s = (K', T')" by (metis surj_pair) hence "set K = set K' \\<^sub>s\<^sub>e\<^sub>t \" "set T = set T' \\<^sub>s\<^sub>e\<^sub>t \" using Ana_subst'[of f S K' T'] fT Ana.hyps(2) s(2) by auto then obtain k' where k': "k' \ set K'" "k = k' \ \" using Ana.hyps(3) by fast show ?thesis proof (cases "k' \ M") case True thus ?thesis using k' \(1,2) by blast next case False hence "k' \ (subterms\<^sub>s\<^sub>e\<^sub>t M \ \((set \ fst \ Ana) ` M)) - M" using k'(1) s_Ana s(1) by force thus ?thesis using SMP.Substitution[OF SMP.MP[of k'] \(1,2)] k'(2) by presburger qed next assume "?Q t" thus ?thesis using SMP.Ana[OF _ Ana.hyps(2,3)] by meson qed qed lemma setops_subterm_trms: assumes t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" and s: "s \ t" shows "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" proof - obtain u u' where u: "pair (u,u') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "t = pair (u,u')" using t setops\<^sub>s\<^sub>s\<^sub>t_are_pairs[of _ S] by blast hence "s \ u \ s \ u'" using s unfolding pair_def by auto thus ?thesis using u setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of u u' S] unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force qed lemma setops_subterms_cases: assumes t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (pair ` setops\<^sub>s\<^sub>s\<^sub>t S)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S) \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" proof - obtain s s' where s: "pair (s,s') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "t \ pair (s,s')" using t setops\<^sub>s\<^sub>s\<^sub>t_are_pairs[of _ S] by blast hence "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \ t \ s \ t \ s'" unfolding pair_def by auto thus ?thesis using s setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of s s' S] unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force qed lemma setops_SMP_cases: assumes "t \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t S)" and "\p. Ana (pair p) = ([], [])" shows "(\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \) \ t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S)" proof - have 0: "\((set \ fst \ Ana) ` pair ` setops\<^sub>s\<^sub>s\<^sub>t S) = {}" proof (induction S) case (Cons x S) thus ?case using assms(2) by (cases x) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) have 1: "\m \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S. is_Fun m" proof (induction S) case (Cons x S) thus ?case unfolding pair_def by (cases x) (auto simp add: assms(2) setops\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) have 2: "subterms\<^sub>s\<^sub>e\<^sub>t (pair ` setops\<^sub>s\<^sub>s\<^sub>t S) \ \((set \ fst \ Ana) ` (pair ` setops\<^sub>s\<^sub>s\<^sub>t S)) - pair ` setops\<^sub>s\<^sub>s\<^sub>t S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" using 0 setops_subterms_cases by fast show ?thesis using SMP_MP_split[OF assms(1) 1] SMP_mono[OF 2] SMP_subterms_eq[of "trms\<^sub>s\<^sub>s\<^sub>t S"] by blast qed lemma constraint_model_priv_const_in_constr_prefix: assumes A: "wf\<^sub>s\<^sub>s\<^sub>t A" and I: "I \\<^sub>s A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" and c: "\public c" "arity c = 0" "Fun c [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t A" using const_subterms_subst_cases[OF c(3)] proof assume "Fun c [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t A" thus ?thesis using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by blast next assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A). x \ subst_domain I \ Fun c [] \ I x" then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A)" "Fun c [] \ I x" by blast have 0: "wf\<^sub>t\<^sub>r\<^sub>m (I x)" "wf'\<^sub>s\<^sub>s\<^sub>t {} A" "Fun c [] \ I = Fun c []" using I A by simp_all have 1: "x \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A" using x(1) in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ A] unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by force hence 2: "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A. Fun c [] \ I \ I v" using 0(3) x(2) by force obtain A\<^sub>p\<^sub>r\<^sub>e A\<^sub>s\<^sub>u\<^sub>f where A': "\(\w \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e. Fun c [] \ I w)" "(\ts. A = A\<^sub>p\<^sub>r\<^sub>e@send\ts\#A\<^sub>s\<^sub>u\<^sub>f \ Fun c [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t I) \ (\s u. A = A\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#A\<^sub>s\<^sub>u\<^sub>f \ Fun c [] \ s \ I \ \(Fun c [] \\<^sub>s\<^sub>e\<^sub>t (I ` fv u))) \ (\s u. A = A\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#A\<^sub>s\<^sub>u\<^sub>f \ (Fun c [] \ s \ I \ Fun c [] \ u \ I))" (is "?X \ ?Y \ ?Z") using wf\<^sub>s\<^sub>s\<^sub>t_strand_first_Send_var_split[OF 0(2) 2] by force show ?thesis using A'(2) proof (elim disjE) assume ?X then obtain ts where ts: "A = A\<^sub>p\<^sub>r\<^sub>e@send\ts\#A\<^sub>s\<^sub>u\<^sub>f" "Fun c [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t I" by blast hence "I \\<^sub>s (A\<^sub>p\<^sub>r\<^sub>e@[send\ts\])" using I(1) strand_sem_append_stateful[of "{}" "{}" "A\<^sub>p\<^sub>r\<^sub>e@[send\ts\]" A\<^sub>s\<^sub>u\<^sub>f I] by auto hence "(ik\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e) \\<^sub>s\<^sub>e\<^sub>t I \ t \ I" when "t \ set ts" for t using that strand_sem_append_stateful[of "{}" "{}" A\<^sub>p\<^sub>r\<^sub>e "[send\ts\]" I] strand_sem_stateful.simps(2)[of _ _ ts "[]"] unfolding list_all_iff by force hence "Fun c [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t I" using ts(2) c(1) private_fun_deduct_in_ik by fast hence "Fun c [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" using A'(1) const_subterms_subst_cases[of c I "ik\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e"] ik\<^sub>s\<^sub>s\<^sub>t_fv_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of A\<^sub>p\<^sub>r\<^sub>e] by fast thus ?thesis using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "A\<^sub>p\<^sub>r\<^sub>e"] unfolding ts(1) by fastforce next assume ?Y then obtain s u where su: "A = A\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#A\<^sub>s\<^sub>u\<^sub>f" "Fun c [] \ s \ I" "\(Fun c [] \\<^sub>s\<^sub>e\<^sub>t I ` fv u)" by fast hence "s \ I = u \ I" using I(1) strand_sem_append_stateful[of "{}" "{}" "A\<^sub>p\<^sub>r\<^sub>e@[\assign: s \ u\]" A\<^sub>s\<^sub>u\<^sub>f I] strand_sem_append_stateful[of _ _ A\<^sub>p\<^sub>r\<^sub>e "[\assign: s \ u\]" I] by auto hence "Fun c [] \ u" using su(2,3) const_subterm_subst_var_obtain[of c u I] by auto thus ?thesis unfolding su(1) by auto next assume ?Z then obtain s u where su: "A = A\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#A\<^sub>s\<^sub>u\<^sub>f" "Fun c [] \ s \ I \ Fun c [] \ u \ I" by fast hence "(s,u) \\<^sub>p I \ dbupd\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e I {}" using I(1) strand_sem_append_stateful[of "{}" "{}" "A\<^sub>p\<^sub>r\<^sub>e@[\assign: s \ u\]" A\<^sub>s\<^sub>u\<^sub>f I] strand_sem_append_stateful[of _ _ A\<^sub>p\<^sub>r\<^sub>e "[\assign: s \ u\]" I] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain s' u' where su': "insert\s',u'\ \ set A\<^sub>p\<^sub>r\<^sub>e" "s \ I = s' \ I" "u \ I = u' \ I" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "s \ I" "u \ I" A\<^sub>p\<^sub>r\<^sub>e I "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of A\<^sub>p\<^sub>r\<^sub>e I "[]"] by fastforce have "fv s' \ fv u' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" using su'(1) unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by force hence "Fun c [] \ s' \ Fun c [] \ u'" using su(2) A'(1) su'(2,3) const_subterm_subst_cases[of c s' I] const_subterm_subst_cases[of c u' I] by auto thus ?thesis using su'(1) unfolding su(1) by force qed qed lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_empty_case: assumes "tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D = []" shows "D = []" "F \ []" proof - show "F \ []" using assms by (auto intro: ccontr) have "tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F (a#A) \ []" for a A by (induct F "a#A" rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) fastforce+ thus "D = []" using assms by (cases D) simp_all qed lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_elem_length_eq: assumes "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" shows "length G = length F" using assms by (induct F D arbitrary: G rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) auto lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_index: assumes "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "i < length F" shows "\d \ set D. G ! i = (pair (F ! i), pair d)" using assms proof (induction F D arbitrary: i G rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D) obtain d G' where G: "d \ set D" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "G = (pair (s,t), pair d)#G'" - using "2.prems"(1) by moura + using "2.prems"(1) by atomize_elim auto show ?case using "2.IH"[OF G(1,2)] "2.prems"(2) G(1,3) by (cases i) auto qed simp lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_cons: assumes "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "d \ set D" shows "(pair (s,t), pair d)#G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) D)" using assms by auto lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_has_pair_lists: assumes "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "g \ set G" shows "\f \ set F. \d \ set D. g = (pair f, pair d)" using assms proof (induction F D arbitrary: G rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D) obtain d G' where G: "d \ set D" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "G = (pair (s,t), pair d)#G'" - using "2.prems"(1) by moura + using "2.prems"(1) by atomize_elim auto show ?case using "2.IH"[OF G(1,2)] "2.prems"(2) G(1,3) by (cases "g \ set G'") auto qed simp lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_pair_lists: assumes "f \ set F" "d \ set D" shows "\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D). (pair f, pair d) \ set G" (is "?P F D f d") proof - have "\f \ set F. \d \ set D. ?P F D f d" proof (induction F D rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D) hence IH: "\f \ set F. \d \ set D. ?P F D f d" by metis moreover have "\d \ set D. ?P ((s,t)#F) D (s,t) d" proof fix d assume d: "d \ set D" then obtain G where G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_empty_case(1) by force hence "(pair (s, t), pair d)#G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) D)" using d by auto thus "?P ((s,t)#F) D (s,t) d" using d G by auto qed ultimately show ?case by fastforce qed simp thus ?thesis by (metis assms) qed lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_db_append_subset: "set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D) \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F (D@E))" (is ?A) "set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F E) \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F (D@E))" (is ?B) proof - show ?A proof (induction F D rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D) show ?case proof fix G assume "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) D)" then obtain d G' where G': "d \ set D" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "G = (pair (s,t), pair d)#G'" - by moura + by atomize_elim auto have "d \ set (D@E)" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F (D@E))" using "2.IH"[OF G'(1)] G'(1,2) by auto thus "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) (D@E))" using G'(3) by auto qed qed simp show ?B proof (induction F E rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F E) show ?case proof fix G assume "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) E)" then obtain d G' where G': "d \ set E" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F E)" "G = (pair (s,t), pair d)#G'" - by moura + by atomize_elim auto have "d \ set (D@E)" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F (D@E))" using "2.IH"[OF G'(1)] G'(1,2) by auto thus "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) (D@E))" using G'(3) by auto qed qed simp qed lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ pair ` set F \ pair ` set D" proof (induction F D arbitrary: G rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D G) obtain d G' where G: "d \ set D" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "G = (pair (s,t), pair d)#G'" - using "2.prems"(1) by moura + using "2.prems"(1) by atomize_elim auto show ?case using "2.IH"[OF G(1,2)] G(1,3) by auto qed simp lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset': "\(trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)) \ pair ` set F \ pair ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset by blast lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s D" proof (induction F D arbitrary: G rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F D G) obtain d G' where G: "d \ set D" "G' \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)" "G = (pair (s,t), pair d)#G'" - using "2.prems"(1) by moura + using "2.prems"(1) by atomize_elim auto show ?case using "2.IH"[OF G(1,2)] G(1,3) unfolding pair_def by auto qed simp lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset': "\(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F D)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset[of _ F D] by blast lemma tr_trms_subset: "A' \ set (tr A D) \ trms\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" proof (induction A D arbitrary: A' rule: tr.induct) case 1 thus ?case by simp next case (2 t A D) - then obtain A'' where A'': "A' = send\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by moura + then obtain A'' where A'': "A' = send\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by atomize_elim auto hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" by (metis "2.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (3 t A D) - then obtain A'' where A'': "A' = receive\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by moura + then obtain A'' where A'': "A' = receive\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by atomize_elim auto hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" by (metis "3.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (4 ac t t' A D) - then obtain A'' where A'': "A' = \ac: t \ t'\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by moura + then obtain A'' where A'': "A' = \ac: t \ t'\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" by atomize_elim auto hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" by (metis "4.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (5 t s A D) hence "A' \ set (tr A (List.insert (t,s) D))" by simp hence "trms\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set (List.insert (t, s) D)" by (metis "5.IH") thus ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (6 t s A D) from 6 obtain Di A'' B C where A'': "Di \ set (subseqs D)" "A'' \ set (tr A [d\D. d \ set Di])" "A' = (B@C)@A''" "B = map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di" "C = map (\d. Inequality [] [(pair (t,s) , pair d)]) [d\D. d \ set Di]" - by moura + by atomize_elim auto hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set [d\D. d \ set Di]" by (metis "6.IH") hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) \ pair ` set D" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "trms\<^sub>s\<^sub>t (B@C) \ insert (pair (t,s)) (pair ` set D)" using A''(4,5) subseqs_set_subset[OF A''(1)] by auto moreover have "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (Delete t s#A)" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(3) trms\<^sub>s\<^sub>t_append[of "B@C" A'] by auto next case (7 ac t s A D) from 7 obtain d A'' where A'': "d \ set D" "A'' \ set (tr A D)" "A' = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#A''" - by moura + by atomize_elim auto hence "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" by (metis "7.IH") moreover have "trms\<^sub>s\<^sub>t A' = {pair (t,s), pair d} \ trms\<^sub>s\<^sub>t A''" using A''(1,3) by auto ultimately show ?case using A''(1) by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (8 X F F' A D) from 8 obtain A'' where A'': "A'' \ set (tr A D)" "A' = (map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))@A''" - by moura + by atomize_elim auto define B where "B \ \(trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))" have "trms\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" by (metis A''(1) "8.IH") hence "trms\<^sub>s\<^sub>t A' \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D" using A'' B_def by auto moreover have "B \ pair ` set F' \ pair ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset'[of F' D] B_def by simp moreover have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (\X\\\: F \\: F'\#A) = pair ` set F' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case by auto qed lemma tr_vars_subset: assumes "A' \ set (tr A D)" shows "fv\<^sub>s\<^sub>t A' \ fv\<^sub>s\<^sub>s\<^sub>t A \ (\(t,t') \ set D. fv t \ fv t')" (is ?P) and "bvars\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>s\<^sub>t A" (is ?Q) proof - show ?P using assms proof (induction A arbitrary: A' D rule: strand_sem_stateful_induct) case (ConsIn A' D ac t s A) then obtain A'' d where *: "d \ set D" "A' = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto hence "fv\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t A \ (\(t,t')\set D. fv t \ fv t')" by (metis ConsIn.IH) thus ?case using * unfolding pair_def by auto next case (ConsDel A' D t s A) define Dfv where "Dfv \ \D::('fun,'var) dbstatelist. (\(t,t')\set D. fv t \ fv t')" define fltD where "fltD \ \Di. filter (\d. d \ set Di) D" define constr where "constr \ \Di. (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) (fltD Di))" from ConsDel obtain A'' Di where *: "Di \ set (subseqs D)" "A' = (constr Di)@A''" "A'' \ set (tr A (fltD Di))" - unfolding constr_def fltD_def by moura + unfolding constr_def fltD_def by atomize_elim auto hence "fv\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t A \ Dfv (fltD Di)" unfolding Dfv_def constr_def fltD_def by (metis ConsDel.IH) moreover have "Dfv (fltD Di) \ Dfv D" unfolding Dfv_def constr_def fltD_def by auto moreover have "Dfv Di \ Dfv D" using subseqs_set_subset(1)[OF *(1)] unfolding Dfv_def constr_def fltD_def by fast moreover have "fv\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ (Dfv Di \ Dfv (fltD Di))" unfolding Dfv_def constr_def fltD_def pair_def by auto moreover have "fv\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) = fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t A" by auto moreover have "fv\<^sub>s\<^sub>t A' = fv\<^sub>s\<^sub>t (constr Di) \ fv\<^sub>s\<^sub>t A''" using * by force ultimately have "fv\<^sub>s\<^sub>t A' \ fv\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) \ Dfv D" by auto thus ?case unfolding Dfv_def fltD_def constr_def by simp next case (ConsNegChecks A' D X F F' A) then obtain A'' where A'': "A'' \ set (tr A D)" "A' = (map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))@A''" - by moura + by atomize_elim auto define B where "B \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))" have 1: "fv\<^sub>s\<^sub>t (map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D)) \ (B \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" unfolding B_def by auto have 2: "B \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset'[of F' D] unfolding B_def by simp have "fv\<^sub>s\<^sub>t A' \ ((fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s D \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X) \ fv\<^sub>s\<^sub>t A''" using 1 2 A''(2) by fastforce thus ?case using ConsNegChecks.IH[OF A''(1)] by auto qed fastforce+ show ?Q using assms by (induct A arbitrary: A' D rule: strand_sem_stateful_induct) fastforce+ qed lemma tr_vars_disj: assumes "A' \ set (tr A D)" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" shows "fv\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>t A' = {}" using assms tr_vars_subset by fast lemma tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_alt_def: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p S = ((\ac t t'. Equality ac t t' \ set S \ (\\. Unifier \ t t') \ \ t = \ t') \ (\X F F'. NegChecks X F F' \ set S \ ( (F' = [] \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F-set X. \a. \ (Var x) = TAtom a)) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') \ T = [] \ (\s \ set T. s \ Var ` set X)))))" (is "?P S = ?Q S") proof show "?P S \ ?Q S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp show "?Q S \ ?P S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp qed lemma tfr\<^sub>s\<^sub>s\<^sub>t_Nil[simp]: "tfr\<^sub>s\<^sub>s\<^sub>t []" by (simp add: tfr\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def) lemma tfr\<^sub>s\<^sub>s\<^sub>t_append: "tfr\<^sub>s\<^sub>s\<^sub>t (A@B) \ tfr\<^sub>s\<^sub>s\<^sub>t A" proof - assume assms: "tfr\<^sub>s\<^sub>s\<^sub>t (A@B)" let ?M = "trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A" let ?N = "trms\<^sub>s\<^sub>s\<^sub>t (A@B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (A@B)" let ?P = "\t t'. \x \ fv t \ fv t'. \a. \ (Var x) = Var a" let ?Q = "\X t t'. X = [] \ (\x \ (fv t \ fv t')-set X. \a. \ (Var x) = Var a)" have *: "SMP ?M - Var`\ \ SMP ?N - Var`\" "?M \ ?N" using SMP_mono[of ?M ?N] setops\<^sub>s\<^sub>s\<^sub>t_append[of A B] by auto { fix s t assume **: "tfr\<^sub>s\<^sub>e\<^sub>t ?N" "s \ SMP ?M - Var`\" "t \ SMP ?M - Var`\" "(\\. Unifier \ s t)" hence "s \ SMP ?N - Var`\" "t \ SMP ?N - Var`\" using * by auto hence "\ s = \ t" using **(1,4) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast } moreover have "\t \ ?N. wf\<^sub>t\<^sub>r\<^sub>m t \ \t \ ?M. wf\<^sub>t\<^sub>r\<^sub>m t" using * by blast ultimately have "tfr\<^sub>s\<^sub>e\<^sub>t ?N \ tfr\<^sub>s\<^sub>e\<^sub>t ?M" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast hence "tfr\<^sub>s\<^sub>e\<^sub>t ?M" using assms unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by metis thus "tfr\<^sub>s\<^sub>s\<^sub>t A" using assms unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by simp qed lemma tfr\<^sub>s\<^sub>s\<^sub>t_append': "tfr\<^sub>s\<^sub>s\<^sub>t (A@B) \ tfr\<^sub>s\<^sub>s\<^sub>t B" proof - assume assms: "tfr\<^sub>s\<^sub>s\<^sub>t (A@B)" let ?M = "trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?N = "trms\<^sub>s\<^sub>s\<^sub>t (A@B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (A@B)" let ?P = "\t t'. \x \ fv t \ fv t'. \a. \ (Var x) = Var a" let ?Q = "\X t t'. X = [] \ (\x \ (fv t \ fv t')-set X. \a. \ (Var x) = Var a)" have *: "SMP ?M - Var`\ \ SMP ?N - Var`\" "?M \ ?N" using SMP_mono[of ?M ?N] setops\<^sub>s\<^sub>s\<^sub>t_append[of A B] by auto { fix s t assume **: "tfr\<^sub>s\<^sub>e\<^sub>t ?N" "s \ SMP ?M - Var`\" "t \ SMP ?M - Var`\" "(\\. Unifier \ s t)" hence "s \ SMP ?N - Var`\" "t \ SMP ?N - Var`\" using * by auto hence "\ s = \ t" using **(1,4) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast } moreover have "\t \ ?N. wf\<^sub>t\<^sub>r\<^sub>m t \ \t \ ?M. wf\<^sub>t\<^sub>r\<^sub>m t" using * by blast ultimately have "tfr\<^sub>s\<^sub>e\<^sub>t ?N \ tfr\<^sub>s\<^sub>e\<^sub>t ?M" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast hence "tfr\<^sub>s\<^sub>e\<^sub>t ?M" using assms unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by metis thus "tfr\<^sub>s\<^sub>s\<^sub>t B" using assms unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by simp qed lemma tfr\<^sub>s\<^sub>s\<^sub>t_cons: "tfr\<^sub>s\<^sub>s\<^sub>t (a#A) \ tfr\<^sub>s\<^sub>s\<^sub>t A" using tfr\<^sub>s\<^sub>s\<^sub>t_append'[of "[a]" A] by simp lemma tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes s: "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) \ range_vars \ = {}" shows "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases s) case (Equality a t t') thus ?thesis proof (cases "\\. Unifier \ (t \ \) (t' \ \)") case True hence "\\. Unifier \ t t'" by (metis subst_subst_compose[of _ \]) moreover have "\ t = \ (t \ \)" "\ t' = \ (t' \ \)" by (metis wt_subst_trm''[OF assms(2)])+ ultimately have "\ (t \ \) = \ (t' \ \)" using s Equality by simp thus ?thesis using Equality True by simp qed simp next case (NegChecks X F G) let ?P = "\F G. G = [] \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F-set X. \a. \ (Var x) = TAtom a)" let ?Q = "\F G. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set G) \ T = [] \ (\s \ set T. s \ Var ` set X)" let ?\ = "rm_vars (set X) \" have "?P F G \ ?Q F G" using NegChecks assms(1) by simp hence "?P (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ ?Q (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" proof assume *: "?P F G" have "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\ = []" using * by simp moreover have "\a. \ (Var x) = TAtom a" when x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) - set X" for x proof - obtain t t' where t: "(t,t') \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" "x \ fv t \ fv t' - set X" using x(1) by auto then obtain u u' where u: "(u,u') \ set F" "u \ ?\ = t" "u' \ ?\ = t'" unfolding subst_apply_pairs_def by auto obtain y where y: "y \ fv u \ fv u' - set X" "x \ fv (?\ y)" using t(2) u(2,3) rm_vars_fv_obtain by fast hence a: "\a. \ (Var y) = TAtom a" using u * by auto have a': "\ (Var y) = \ (?\ y)" using wt_subst_trm''[OF wt_subst_rm_vars[OF \(1), of "set X"], of "Var y"] by simp have "(\z. ?\ y = Var z) \ (\c. ?\ y = Fun c [])" proof (cases "?\ y \ subst_range \") case True thus ?thesis using a a' \(2) const_type_inv_wf by (cases "?\ y") fastforce+ qed fastforce hence "?\ y = Var x" using y(2) by fastforce hence "\ (Var x) = \ (Var y)" using a' by simp thus ?thesis using a by presburger qed ultimately show ?thesis by simp next assume *: "?Q F G" have **: "set X \ range_vars ?\ = {}" using \(3) NegChecks rm_vars_img_fv_subset[of "set X" \] by auto have "?Q (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" using ineq_subterm_inj_cond_subst[OF ** *] trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of F "rm_vars (set X) \"] subst_apply_pairs_pair_image_subst[of G "rm_vars (set X) \"] by (metis (no_types, lifting) image_Un) thus ?thesis by simp qed thus ?thesis using NegChecks by simp qed simp_all lemma tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply: assumes S: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "bvars\<^sub>s\<^sub>s\<^sub>t S \ range_vars \ = {}" shows "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>s\<^sub>t \)" proof - have "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) \ range_vars \ = {}" when "s \ set S" for s using that \(3) unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def range_vars_alt_def by fastforce thus ?thesis using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF _ \(1,2)] S unfolding list_all_iff by (auto simp add: subst_apply_stateful_strand_def) qed lemma tfr_setops_if_tfr_trms: assumes "Pair \ \(funs_term ` SMP (trms\<^sub>s\<^sub>s\<^sub>t S))" and "\p. Ana (pair p) = ([], [])" and "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S. \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S. (\\. Unifier \ s t) \ \ s = \ t" and "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S. \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S. (\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)) \ (\\. Unifier \ s t)" and tfr: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S)" proof - have 0: "t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var \ t \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t S) - range Var" when "t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S) - range Var" for t using that SMP_union by blast have 1: "s \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" when st: "s \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t S) - range Var" "t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" "\\. Unifier \ s t" for s t proof - have "(\\. s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \) \ s \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" using st setops_SMP_cases[of s S] assms(2) by blast moreover { fix \ assume \: "s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" then obtain s' where s': "s' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "s = s' \ \" by blast then obtain u u' where u: "s' = Fun Pair [u,u']" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs[of s'] unfolding pair_def by fast hence *: "s = Fun Pair [u \ \, u' \ \]" using \ s' by simp obtain f T where fT: "t = Fun f T" using st(2) by (cases t) auto hence "f \ Pair" using st(2) assms(1) by auto hence False using st(3) * fT s' u by fast } ultimately show ?thesis by meson qed have 2: "\ s = \ t" when "s \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" "t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" "\\. Unifier \ s t" for s t using that tfr unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast have 3: "\ s = \ t" when st: "s \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t S) - range Var" "t \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t S) - range Var" "\\. Unifier \ s t" for s t proof - let ?P = "\s \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" have "(\\. ?P s \) \ s \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" "(\\. ?P t \) \ t \ SMP (trms\<^sub>s\<^sub>s\<^sub>t S) - range Var" using setops_SMP_cases[of _ S] assms(2) st(1,2) by auto hence "(\\ \'. ?P s \ \ ?P t \') \ \ s = \ t" by (metis 1 2 st) moreover { fix \ \' assume *: "?P s \" "?P t \'" then obtain s' t' where **: "s' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S" "s = s' \ \" "t = t' \ \'" by blast hence "\\. Unifier \ s' t'" using st(3) assms(4) * by blast hence "\ s' = \ t'" using assms(3) ** by blast hence "\ s = \ t" using * **(3,4) wt_subst_trm''[of \ s'] wt_subst_trm''[of \' t'] by argo } ultimately show ?thesis by blast qed show ?thesis using 0 1 2 3 unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by metis qed end subsection \The Typing Result for Stateful Constraints\ subsubsection \Correctness of the Constraint Reduction\ context stateful_typed_model begin context begin private lemma tr_wf': assumes "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "\(t,t') \ set D. fv t \ fv t' \ X" and "wf'\<^sub>s\<^sub>s\<^sub>t X A" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "A' \ set (tr A D)" shows "wf\<^sub>s\<^sub>t X A'" proof - define P where "P = (\(D::('fun,'var) dbstatelist) (A::('fun,'var) stateful_strand). (\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}) \ fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {})" have "P D A" using assms(1,4) by (simp add: P_def) with assms(5,3,2) show ?thesis proof (induction A arbitrary: A' D X rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case 1 thus ?case by simp next case (2 X ts A A') then obtain A'' where A'': "A' = receive\ts\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A" "\(s,s') \ set D. fv s \ fv s' \ X" "P D A" using 2(1,2,3,4) apply (force, force) using 2(5) unfolding P_def by force show ?case using "2.IH"[OF A''(2) *] A''(1,3) by simp next case (3 X ts A A') then obtain A'' where A'': "A' = send\ts\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) A" "\(s,s') \ set D. fv s \ fv s' \ X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P D A" using 3(1,2,3,4) apply (force, force) using 3(5) unfolding P_def by force show ?case using "3.IH"[OF A''(2) *] A''(1) by simp next case (4 X t t' A A') then obtain A'' where A'': "A' = \assign: t \ t'\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" "fv t' \ X" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t) A" "\(s,s') \ set D. fv s \ fv s' \ X \ fv t" "P D A" using 4(1,2,3,4) apply (force, force) using 4(5) unfolding P_def by force show ?case using "4.IH"[OF A''(2) *] A''(1,3) by simp next case (5 X t t' A A') then obtain A'' where A'': "A' = \check: t \ t'\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A" "P D A" using 5(3) apply force using 5(5) unfolding P_def by force show ?case using "5.IH"[OF A''(2) *(1) 5(4) *(2)] A''(1) by simp next case (6 X t s A A') hence A': "A' \ set (tr A (List.insert (t,s) D))" "fv t \ X" "fv s \ X" by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A" "\(s,s') \ set (List.insert (t,s) D). fv s \ fv s' \ X" using 6 by auto have **: "P (List.insert (t,s) D) A" using 6(5) unfolding P_def by force show ?case using "6.IH"[OF A'(1) * **] A'(2,3) by simp next case (7 X t s A A') let ?constr = "\Di. (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])" from 7 obtain Di A'' where A'': "A' = ?constr Di@A''" "A'' \ set (tr A [d\D. d \ set Di])" "Di \ set (subseqs D)" - by moura + by atomize_elim force have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A" "\(t',s') \ set [d\D. d \ set Di]. fv t' \ fv s' \ X" using 7 by auto have **: "P [d\D. d \ set Di] A" using 7 unfolding P_def by force have ***: "\(t, t') \ set D. fv t \ fv t' \ X" using 7 by auto show ?case using "7.IH"[OF A''(2) * **] A''(1) wf_fun_pair_eqs_ineqs_map[OF _ A''(3) ***] by simp next case (8 X t s A A') then obtain d A'' where A'': "A' = \assign: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" "d \ set D" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t \ fv s) A" "\(t',s')\set D. fv t' \ fv s' \ X \ fv t \ fv s" "P D A" using 8(1,2,3,4) apply (force, force) using 8(5) unfolding P_def by force have **: "fv (pair d) \ X" using A''(3) "8.prems"(3) unfolding pair_def by fastforce have ***: "fv (pair (t,s)) = fv s \ fv t" unfolding pair_def by auto show ?case using "8.IH"[OF A''(2) *] A''(1) ** *** unfolding pair_def by (simp add: Un_assoc) next case (9 X t s A A') then obtain d A'' where A'': "A' = \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" "d \ set D" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A""P D A" using 9(3) apply force using 9(5) unfolding P_def by force have **: "fv (pair d) \ X" using A''(3) "9.prems"(3) unfolding pair_def by fastforce have ***: "fv (pair (t,s)) = fv s \ fv t" unfolding pair_def by auto show ?case using "9.IH"[OF A''(2) *(1) 9(4) *(2)] A''(1) ** *** by (simp add: Un_assoc) next case (10 X Y F F' A A') from 10 obtain A'' where A'': "A' = (map (\G. \Y\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))@A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X A" "\(t',s') \ set D. fv t' \ fv s' \ X" using 10 by auto have "bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t (\Y\\\: F \\: F'\#A)" "fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t (\Y\\\: F \\: F'\#A)" by auto hence **: "P D A" using 10 unfolding P_def by blast show ?case using "10.IH"[OF A''(2) * **] A''(1) wf_fun_pair_negchecks_map by simp qed qed private lemma tr_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "A' \ set (tr A [])" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t A')" using tr_trms_subset[OF assms(1)] setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2)[OF assms(2)] by auto lemma tr_wf: assumes "A' \ set (tr A [])" and "wf\<^sub>s\<^sub>s\<^sub>t A" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A)" shows "wf\<^sub>s\<^sub>t {} A'" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t A')" and "fv\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>t A' = {}" using tr_wf'[OF _ _ _ _ assms(1)] tr_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF assms(1,3)] tr_vars_disj[OF assms(1)] assms(2) by fastforce+ private lemma fun_pair_ineqs: assumes "d \\<^sub>p \ \\<^sub>p \ \ d' \\<^sub>p \" shows "pair d \ \ \ \ \ pair d' \ \" proof - have "d \\<^sub>p (\ \\<^sub>s \) \ d' \\<^sub>p \" using assms subst_pair_compose by metis hence "pair d \ (\ \\<^sub>s \) \ pair d' \ \" using fun_pair_eq_subst by metis thus ?thesis by simp qed private lemma tr_Delete_constr_iff_aux1: assumes "\d \ set Di. (t,s) \\<^sub>p \ = d \\<^sub>p \" and "\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \" shows "\M; (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])\\<^sub>d \" proof - from assms(2) have "\M; map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" proof (induction D) case (Cons d D) hence IH: "\M; map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D . d \ set Di]\\<^sub>d \" by auto thus ?case proof (cases "d \ set Di") case False hence "(t,s) \\<^sub>p \ \ d \\<^sub>p \" using Cons by simp hence "pair (t,s) \ \ \ pair d \ \" using fun_pair_eq_subst by metis moreover have "\t (\::('fun,'var) subst). subst_domain \ = {} \ t \ \ = t" by auto ultimately have "\\. subst_domain \ = {} \ pair (t,s) \ \ \ \ \ pair d \ \ \ \" by metis thus ?thesis using IH by (simp add: ineq_model_def) qed simp qed simp moreover { fix B assume "\M; B\\<^sub>d \" with assms(1) have "\M; (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@B\\<^sub>d \" unfolding pair_def by (induction Di) auto } ultimately show ?thesis by metis qed private lemma tr_Delete_constr_iff_aux2: assumes "ground M" and "\M; (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])\\<^sub>d \" shows "(\d \ set Di. (t,s) \\<^sub>p \ = d \\<^sub>p \) \ (\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \)" proof - let ?c1 = "map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di" let ?c2 = "map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]" have "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using assms(1) subst_all_ground_ident by metis moreover have "ik\<^sub>s\<^sub>t ?c1 = {}" by auto ultimately have *: "\M; map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di\\<^sub>d \" "\M; map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" using strand_sem_split(3,4)[of M ?c1 ?c2 \] assms(2) by auto from *(1) have 1: "\d \ set Di. (t,s) \\<^sub>p \ = d \\<^sub>p \" unfolding pair_def by (induct Di) auto from *(2) have 2: "\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \" proof (induction D arbitrary: Di) case (Cons d D) thus ?case proof (cases "d \ set Di") case False hence IH: "\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \" using Cons by force have "\t (\::('fun,'var) subst). subst_domain \ = {} \ ground (subst_range \) \ \ = Var" by auto moreover have "ineq_model \ [] [((pair (t,s)), (pair d))]" using False Cons.prems by simp ultimately have "pair (t,s) \ \ \ pair d \ \" by (simp add: ineq_model_def) thus ?thesis using IH unfolding pair_def by force qed simp qed simp show ?thesis by (metis 1 2) qed private lemma tr_Delete_constr_iff: fixes \::"('fun,'var) subst" assumes "ground M" shows "set Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {(t,s) \\<^sub>p \} \ (t,s) \\<^sub>p \ \ (set D - set Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ \M; (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])\\<^sub>d \" proof - let ?constr = "(map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di])" { assume "set Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {(t,s) \\<^sub>p \}" "(t,s) \\<^sub>p \ \ (set D - set Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" hence "\d \ set Di. (t,s) \\<^sub>p \ = d \\<^sub>p \" "\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \" by auto hence "\M; ?constr\\<^sub>d \" using tr_Delete_constr_iff_aux1 by simp } moreover { assume "\M; ?constr\\<^sub>d \" hence "\d \ set Di. (t,s) \\<^sub>p \ = d \\<^sub>p \" "\d \ set D - set Di. (t,s) \\<^sub>p \ \ d \\<^sub>p \" using assms tr_Delete_constr_iff_aux2 by auto hence "set Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {(t,s) \\<^sub>p \} \ (t,s) \\<^sub>p \ \ (set D - set Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by force } ultimately show ?thesis by metis qed private lemma tr_NotInSet_constr_iff: fixes \::"('fun,'var) subst" assumes "\(t,t') \ set D. (fv t \ fv t') \ set X = {}" shows "(\\. subst_domain \ = set X \ ground (subst_range \) \ (t,s) \\<^sub>p \ \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) \ \M; map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) D\\<^sub>d \" proof - { assume "\\. subst_domain \ = set X \ ground (subst_range \) \ (t,s) \\<^sub>p \ \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" with assms have "\M; map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) D\\<^sub>d \" proof (induction D) case (Cons d D) - obtain t' s' where d: "d = (t',s')" by moura + obtain t' s' where d: "d = (t',s')" by atomize_elim auto have "\M; map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) D\\<^sub>d \" "map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) (d#D) = \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t#map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) D" using Cons by auto moreover have "\\. subst_domain \ = set X \ ground (subst_range \) \ pair (t, s) \ \ \ \ \ pair d \ \" using fun_pair_ineqs[of \ _ "(t,s)" \ d] Cons.prems(2) by auto moreover have "(fv t' \ fv s') \ set X = {}" using Cons.prems(1) d by auto hence "\\. subst_domain \ = set X \ pair d \ \ = pair d" using d unfolding pair_def by auto ultimately show ?case by (simp add: ineq_model_def) qed simp } moreover { fix \::"('fun,'var) subst" assume "\M; map (\d. \X\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) D\\<^sub>d \" and \: "subst_domain \ = set X" "ground (subst_range \)" with assms have "(t,s) \\<^sub>p \ \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" proof (induction D) case (Cons d D) - obtain t' s' where d: "d = (t',s')" by moura + obtain t' s' where d: "d = (t',s')" by atomize_elim auto have "(t,s) \\<^sub>p \ \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" "pair (t,s) \ \ \ \ \ pair d \ \ \ \" using Cons d by (auto simp add: ineq_model_def simp del: subst_range.simps) moreover have "pair d \ \ = pair d" using Cons.prems(1) fun_pair_subst[of d \] d \(1) unfolding pair_def by auto ultimately show ?case unfolding pair_def by force qed simp } ultimately show ?thesis by metis qed lemma tr_NegChecks_constr_iff: "(\G\set L. ineq_model \ X (F@G)) \ \M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) L\\<^sub>d \" (is ?A) "negchecks_model \ D X F F' \ \M; D; [\X\\\: F \\: F'\]\\<^sub>s \" (is ?B) proof - show ?A by (induct L) auto show ?B by simp qed lemma tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv: fixes \::"('fun,'var) subst" assumes "\(t,t') \ set D. (fv t \ fv t') \ set X = {}" shows "negchecks_model \ (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) X F F' \ (\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D). ineq_model \ X (F@G))" proof - define P where "P \ \\::('fun,'var) subst. subst_domain \ = set X \ ground (subst_range \)" define Ineq where "Ineq \ \(\::('fun,'var) subst) F. \(s,t) \ set F. s \ \ \\<^sub>s \ \ t \ \ \\<^sub>s \" define Ineq' where "Ineq' \ \(\::('fun,'var) subst) F. \(s,t) \ set F. s \ \ \\<^sub>s \ \ t \ \" define Notin where "Notin \ \(\::('fun,'var) subst) D F'. \(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" have sublmm: "((s,t) \\<^sub>p \ \\<^sub>s \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) \ (list_all (\d. Ineq' \ [(pair (s,t),pair d)]) D)" for s t \ D unfolding pair_def by (induct D) (auto simp add: Ineq'_def) have "Notin \ D F' \ (\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D). Ineq' \ G)" (is "?A \ ?B") when "P \" for \ proof show "?A \ ?B" proof (induction F' D rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F' D) show ?case proof (cases "Notin \ D F'") case False hence "(s,t) \\<^sub>p \ \\<^sub>s \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using "2.prems" by (auto simp add: Notin_def) hence "pair (s,t) \ \ \\<^sub>s \ \ pair d \ \" when "d \ set D" for d using that sublmm Ball_set[of D "\d. Ineq' \ [(pair (s,t), pair d)]"] by (simp add: Ineq'_def) moreover have "\d \ set D. \G'. G = (pair (s,t), pair d)#G'" when "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F') D)" for G using that tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_index[OF that, of 0] by force ultimately show ?thesis by (simp add: Ineq'_def) qed (auto dest: "2.IH" simp add: Ineq'_def) qed (simp add: Notin_def) have "\?A \ \?B" proof (induction F' D rule: tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.induct) case (2 s t F' D) hence "\Notin \ D F'" "D \ []" unfolding Notin_def by auto then obtain G where G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D)" "\Ineq' \ G" using "2.IH" by (cases D) auto obtain d where d: "d \ set D" "pair (s,t) \ \ \\<^sub>s \ = pair d \ \" using "2.prems" unfolding pair_def by (auto simp add: Notin_def) thus ?case using G(2) tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_cons[OF G(1) d(1)] by (auto simp add: Ineq'_def) qed (simp add: Ineq'_def) thus "?B \ ?A" by metis qed hence *: "(\\. P \ \ Ineq \ F \ Notin \ D F') \ (\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D). \\. P \ \ Ineq \ F \ Ineq' \ G)" by auto have "t \ \ = t" when "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D)" "(s,t) \ set G" "P \" for \ s t G using assms that(3) tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_has_pair_lists[OF that(1,2)] unfolding pair_def by (fastforce simp add: P_def) hence **: "Ineq' \ G = Ineq \ G" when "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D)" "P \" for \ G using that unfolding Ineq_def Ineq'_def by force have ***: "negchecks_model \ (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) X F F' \ (\\. P \ \ Ineq \ F \ Notin \ D F')" unfolding P_def Ineq_def Notin_def negchecks_model_def by blast have "ineq_model \ X (F@G) \ (\\. P \ \ Ineq \ (F@G))" for G unfolding P_def Ineq_def ineq_model_def by blast hence ****: "ineq_model \ X (F@G) \ (\\. P \ \ Ineq \ F \ Ineq \ G)" for G unfolding Ineq_def by fastforce show ?thesis using * ** *** **** by simp qed lemma tr_sem_equiv': assumes "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "ground M" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \ \ (\A' \ set (tr A D). \M; A'\\<^sub>d \)" (is "?P \ ?Q") proof have \_grounds: "\t. fv (t \ \) = {}" by (rule interpretation_grounds[OF \]) have "\A' \ set (tr A D). \M; A'\\<^sub>d \" when ?P using that assms(1,2,3) proof (induction A arbitrary: D rule: strand_sem_stateful_induct) case (ConsRcv M D ts A) have "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground ((set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M)" using \ ConsRcv.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain A' where A': "A' \ set (tr A D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; A'\\<^sub>d \" by (metis ConsRcv.IH) thus ?case by auto next case (ConsSnd M D ts A) have "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "\t \ set ts. M \ t \ \" using \ ConsSnd.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain A' where A': "A' \ set (tr A D)" "\M; A'\\<^sub>d \" by (metis ConsSnd.IH) thus ?case using * by auto next case (ConsEq M D ac t t' A) have "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "t \ \ = t' \ \" using \ ConsEq.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain A' where A': "A' \ set (tr A D)" "\M; A'\\<^sub>d \" by (metis ConsEq.IH) thus ?case using * by auto next case (ConsIns M D t s A) have "\M; set (List.insert (t,s) D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "\(t,t') \ set (List.insert (t,s) D). (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" using ConsIns.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain A' where A': "A' \ set (tr A (List.insert (t,s) D))" "\M; A'\\<^sub>d \" by (metis ConsIns.IH) thus ?case by auto next case (ConsDel M D t s A) have *: "\M; (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {(t,s) \\<^sub>p \}; A\\<^sub>s \" "\(t,t')\set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" using ConsDel.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain Di where Di: "Di \ set D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {(t,s) \\<^sub>p \}" "(t,s) \\<^sub>p \ \ (set D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" - using subset_subst_pairs_diff_exists'[of "set D"] by moura + using subset_subst_pairs_diff_exists'[of "set D"] by atomize_elim auto hence **: "(set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {(t,s) \\<^sub>p \} = (set D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by blast obtain Di' where Di': "set Di' = Di" "Di' \ set (subseqs D)" - using subset_sublist_exists[OF Di(1)] by moura + using subset_sublist_exists[OF Di(1)] by atomize_elim auto hence ***: "(set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {(t,s) \\<^sub>p \} = (set [d\D. d \ set Di'] \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using Di ** by auto define constr where "constr \ map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di'@ map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di']" have ****: "\(t,t')\set [d\D. d \ set Di']. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" using *(2) Di(1) Di'(1) subseqs_set_subset[OF Di'(2)] by simp have "set D - Di = set [d\D. d \ set Di']" using Di Di' by auto hence *****: "\M; set [d\D. d \ set Di'] \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" using *(1) ** by metis obtain A' where A': "A' \ set (tr A [d\D. d \ set Di'])" "\M; A'\\<^sub>d \" - using ConsDel.IH[OF ***** **** *(3,4)] by moura + using ConsDel.IH[OF ***** **** *(3,4)] by atomize_elim auto hence constr_sat: "\M; constr\\<^sub>d \" using Di Di' *(1) *** tr_Delete_constr_iff[OF *(4), of \ Di' t s D] unfolding constr_def by auto have "constr@A' \ set (tr (Delete t s#A) D)" using A'(1) Di' unfolding constr_def by auto moreover have "ik\<^sub>s\<^sub>t constr = {}" unfolding constr_def by auto hence "\M \\<^sub>s\<^sub>e\<^sub>t \; constr\\<^sub>d \" "\M \ (ik\<^sub>s\<^sub>t constr \\<^sub>s\<^sub>e\<^sub>t \); A'\\<^sub>d \" using constr_sat A'(2) subst_all_ground_ident[OF *(4)] by simp_all ultimately show ?case using strand_sem_append(2)[of _ _ \] subst_all_ground_ident[OF *(4), of \] by metis next case (ConsIn M D ac t s A) have "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "(t,s) \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using \ ConsIn.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ then obtain A' where A': "A' \ set (tr A D)" "\M; A'\\<^sub>d \" by (metis ConsIn.IH) moreover obtain d where "d \ set D" "pair (t,s) \ \ = pair d \ \" using * unfolding pair_def by auto ultimately show ?case using * by auto next case (ConsNegChecks M D X F F' A) let ?ineqs = "(map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))" have 1: "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" "ground M" using ConsNegChecks by auto have 2: "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" using ConsNegChecks.prems(2,3) \ unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce+ have 3: "negchecks_model \ (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) X F F'" using ConsNegChecks.prems(1) by simp from 1 2 obtain A' where A': "A' \ set (tr A D)" "\M; A'\\<^sub>d \" by (metis ConsNegChecks.IH) have 4: "\(t,t') \ set D. (fv t \ fv t') \ set X = {}" using ConsNegChecks.prems(2) unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by auto have "\M; ?ineqs\\<^sub>d \" using 3 tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 4] tr_NegChecks_constr_iff by metis moreover have "ik\<^sub>s\<^sub>t ?ineqs = {}" by auto moreover have "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using 1(2) \ by (simp add: subst_all_ground_ident) ultimately show ?case using strand_sem_append(2)[of M ?ineqs \ A'] A' by force qed simp thus "?P \ ?Q" by metis have "(\A' \ set (tr A D). \M; A'\\<^sub>d \) \ ?P" using assms(1,2,3) proof (induction A arbitrary: D rule: strand_sem_stateful_induct) case (ConsRcv M D ts A) have "\A' \ set (tr A D). \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; A'\\<^sub>d \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground ((set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M)" using \ ConsRcv.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsRcv.IH) thus ?case by auto next case (ConsSnd M D ts A) have "\A' \ set (tr A D). \M; A'\\<^sub>d \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "\t \ set ts. M \ t \ \" using \ ConsSnd.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ hence "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsSnd.IH) thus ?case using * by auto next case (ConsEq M D ac t t' A) have "\A' \ set (tr A D). \M; A'\\<^sub>d \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "t \ \ = t' \ \" using \ ConsEq.prems unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ hence "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsEq.IH) thus ?case using * by auto next case (ConsIns M D t s A) hence "\A' \ set (tr A (List.insert (t,s) D)). \M; A'\\<^sub>d \" "\(t,t') \ set (List.insert (t,s) D). (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by auto+ hence "\M; set (List.insert (t,s) D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsIns.IH) thus ?case by auto next case (ConsDel M D t s A) define constr where "constr \ \Di. map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di@ map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]" let ?flt = "\Di. filter (\d. d \ set Di) D" have "\Di \ set (subseqs D). \B' \ set (tr A (?flt Di)). B = constr Di@B'" when "B \ set (tr (delete\t,s\#A) D)" for B using that unfolding constr_def by auto then obtain A' Di where A': "constr Di@A' \ set (tr (Delete t s#A) D)" "A' \ set (tr A (?flt Di))" "Di \ set (subseqs D)" "\M; constr Di@A'\\<^sub>d \" using ConsDel.prems(1) by blast have 1: "\(t,t')\set (?flt Di). (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" using ConsDel.prems(2) by auto have 2: "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" using ConsDel.prems(3) by force+ have "ik\<^sub>s\<^sub>t (constr Di) = {}" unfolding constr_def by auto hence 3: "\M; A'\\<^sub>d \" using subst_all_ground_ident[OF ConsDel.prems(4)] A'(4) strand_sem_split(4)[of M "constr Di" A' \] by simp have IH: "\M; set (?flt Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsDel.IH[OF _ 1 2 ConsDel.prems(4)] 3 A'(2)) have "\M; constr Di\\<^sub>d \" using subst_all_ground_ident[OF ConsDel.prems(4)] strand_sem_split(3) A'(4) by metis hence *: "set Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {(t,s) \\<^sub>p \}" "(t,s) \\<^sub>p \ \ (set D - set Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using tr_Delete_constr_iff[OF ConsDel.prems(4), of \ Di t s D] unfolding constr_def by auto have 4: "set (?flt Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {((t,s) \\<^sub>p \)}" proof show "set (?flt Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {((t,s) \\<^sub>p \)}" proof fix u u' assume u: "(u,u') \ set (?flt Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" then obtain v v' where v: "(v,v') \ set D - set Di" "(v,v') \\<^sub>p \ = (u,u')" by auto hence "(u,u') \ (t,s) \\<^sub>p \" using * by force thus "(u,u') \ (set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {((t,s) \\<^sub>p \)}" using u v * subseqs_set_subset[OF A'(3)] by auto qed show "(set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - {((t,s) \\<^sub>p \)} \ set (?flt Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using * subseqs_set_subset[OF A'(3)] by force qed show ?case using 4 IH by simp next case (ConsIn M D ac t s A) have "\A' \ set (tr A D). \M; A'\\<^sub>d \" "\(t,t') \ set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "ground M" and *: "(t,s) \\<^sub>p \ \ set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using ConsIn.prems(1,2,3,4) apply (fastforce, fastforce, fastforce, fastforce) using ConsIn.prems(1) tr.simps(7)[of ac t s A D] unfolding pair_def by fastforce hence "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsIn.IH) moreover obtain d where "d \ set D" "pair (t,s) \ \ = pair d \ \" using * unfolding pair_def by auto ultimately show ?case using * by auto next case (ConsNegChecks M D X F F' A) let ?ineqs = "(map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))" obtain B where B: "?ineqs@B \ set (tr (NegChecks X F F'#A) D)" "\M; ?ineqs@B\\<^sub>d \" "B \ set (tr A D)" - using ConsNegChecks.prems(1) by moura + using ConsNegChecks.prems(1) by atomize_elim auto moreover have "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using ConsNegChecks.prems(4) \ by (simp add: subst_all_ground_ident) moreover have "ik\<^sub>s\<^sub>t ?ineqs = {}" by auto ultimately have "\M; B\\<^sub>d \" using strand_sem_split(4)[of M ?ineqs B \] by simp moreover have "\(t,t')\set D. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" using ConsNegChecks.prems(2,3) unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by force+ ultimately have "\M; set D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; A\\<^sub>s \" by (metis ConsNegChecks.IH B(3) ConsNegChecks.prems(4)) moreover have "\(t, t')\set D. (fv t \ fv t') \ set X = {}" using ConsNegChecks.prems(2) unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by force ultimately show ?case using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv tr_NegChecks_constr_iff B(2) strand_sem_split(3)[of M ?ineqs B \] \M \\<^sub>s\<^sub>e\<^sub>t \ = M\ by simp qed simp thus "?Q \ ?P" by metis qed lemma tr_sem_equiv: assumes "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\ \\<^sub>s A \ (\A' \ set (tr A []). (\ \ \A'\))" using tr_sem_equiv'[OF _ assms(1) _ assms(2), of "[]" "{}"] unfolding constr_sem_d_def by auto end end subsubsection \Typing Result Locale Definition\ locale stateful_typing_result = stateful_typed_model arity public Ana \ Pair + typing_result arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" subsubsection \Type-Flaw Resistance Preservation of the Constraint Reduction\ context stateful_typing_result begin context begin private lemma tr_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "A' \ set (tr A D)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p A" and "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" (is "?P0 A D") and "\(t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" (is "?P1 A D") and "\t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D. \t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set D. (\\. Unifier \ t t') \ \ t = \ t'" (is "?P3 A D") shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p A'" proof - have sublmm: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p A" "?P0 A D" "?P1 A D" "?P3 A D" when p: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a#A)" "?P0 (a#A) D" "?P1 (a#A) D" "?P3 (a#A) D" for a A D using p(1) apply (simp add: tfr\<^sub>s\<^sub>s\<^sub>t_def) using p(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset apply fast using p(3) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset apply fast using p(4) setops\<^sub>s\<^sub>s\<^sub>t_cons_subset by fast show ?thesis using assms proof (induction A D arbitrary: A' rule: tr.induct) case 1 thus ?case by simp next case (2 t A D) note prems = "2.prems" note IH = "2.IH" from prems(1) obtain A'' where A'': "A' = send\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (3 t A D) note prems = "3.prems" note IH = "3.IH" from prems(1) obtain A'' where A'': "A' = receive\t\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (4 ac t t' A D) note prems = "4.prems" note IH = "4.IH" from prems(1) obtain A'' where A'': "A' = \ac: t \ t'\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" - by moura + by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson moreover have "(\\. Unifier \ t t') \ \ t = \ t'" using prems(2) by (simp add: tfr\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(1) by auto next case (5 t s A D) note prems = "5.prems" note IH = "5.IH" from prems(1) have A': "A' \ set (tr A (List.insert (t,s) D))" by simp have 1: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p A" using sublmm[OF prems(2,3,4,5)] by simp have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (Insert t s#A) \ pair`set D = pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair`set (List.insert (t,s) D)" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence 3: "?P3 A (List.insert (t,s) D)" using prems(5) by metis moreover have "?P1 A (List.insert (t, s) D)" using prems(3,4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of A] by auto ultimately have "list_all tfr\<^sub>s\<^sub>t\<^sub>p A'" using IH[OF A' sublmm(1,2)[OF prems(2,3,4,5)] _ 3] by metis thus ?case using A'(1) by auto next case (6 t s A D) note prems = "6.prems" note IH = "6.IH" define constr where constr: "constr \ (\Di. (map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di)@ (map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\D. d \ set Di]))" from prems(1) obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr A [d\D. d \ set Di])" "Di \ set (subseqs D)" unfolding constr by auto define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" have "set [d\D. d \ set Di] \ set D" "pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ pair ` set [d\D. d \ set Di] \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) \ pair ` set D" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence *: "?P3 A [d\D. d \ set Di]" using prems(5) by blast have **: "?P1 A [d\D. d \ set Di]" using prems(4,5) by auto have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(3,2) sublmm(1,2)[OF prems(2,3,4,5)] ** *] by metis have 2: "\ac: u \ u'\\<^sub>s\<^sub>t \ set A'' \ (\d \ set Di. u = pair (t,s) \ u' = pair d)" when "\ac: u \ u'\\<^sub>s\<^sub>t \ set A'" for ac u u' using that A''(1) unfolding constr by force have 3: "Inequality X U \ set A' \ Inequality X U \ set A'' \ (\d \ set [d\D. d \ set Di]. U = [(pair (t,s), pair d)] \ Q2 [(pair (t,s), pair d)] X)" for X U using A''(1) unfolding Q2_def constr by force have 4: "\d\set D. (\\. Unifier \ (pair (t,s)) (pair d)) \ \ (pair (t,s)) = \ (pair d)" using prems(5) by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) { fix ac u u' assume a: "\ac: u \ u'\\<^sub>s\<^sub>t \ set A'" "\\. Unifier \ u u'" hence "\ac: u \ u'\\<^sub>s\<^sub>t \ set A'' \ (\d \ set Di. u = pair (t,s) \ u' = pair d)" using 2 by metis hence "\ u = \ u'" using 1(1) 4 subseqs_set_subset[OF A''(3)] a(2) tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of A''] by blast } moreover { fix u U assume "\U\\\: u\\<^sub>s\<^sub>t \ set A'" hence "\U\\\: u\\<^sub>s\<^sub>t \ set A'' \ (\d \ set [d\D. d \ set Di]. u = [(pair (t,s), pair d)] \ Q2 u U)" using 3 by metis hence "Q1 u U \ Q2 u U" using 1 4 subseqs_set_subset[OF A''(3)] tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of A''] unfolding Q1_def Q2_def by blast } ultimately show ?case using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of A'] unfolding Q1_def Q2_def by blast next case (7 ac t s A D) note prems = "7.prems" note IH = "7.IH" from prems(1) obtain d A'' where A'': "A' = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#A''" "A'' \ set (tr A D)" "d \ set D" - by moura + by atomize_elim auto have "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis moreover have "(\\. Unifier \ (pair (t,s)) (pair d)) \ \ (pair (t,s)) = \ (pair d)" using prems(2,5) A''(3) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(1) by fastforce next case (8 X F F' A D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = (map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D))" define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(M::('fun,'var) terms) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M \ T = [] \ (\s \ set T. s \ Var ` set X))" have Q2_subset: "Q2 M' X" when "M' \ M" "Q2 M X" for X M M' using that unfolding Q2_def by auto have Q2_supset: "Q2 (M \ M') X" when "Q2 M X" "Q2 M' X" for X M M' using that unfolding Q2_def by auto from prems(1) obtain A'' where A'': "A' = constr@A''" "A'' \ set (tr A D)" - using constr_def by moura + using constr_def by atomize_elim auto have 0: "F' = [] \ constr = [\X\\\: F\\<^sub>s\<^sub>t]" unfolding constr_def by simp have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p A''" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "(F' = [] \ Q1 F X) \ Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" using prems(2) unfolding Q1_def Q2_def by simp have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p constr" when "F' = []" "Q1 F X" using that 0 2 tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of constr] unfolding Q1_def by auto { fix c assume "c \ set constr" hence "\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D). c = \X\\\: (F@G)\\<^sub>s\<^sub>t" unfolding constr_def by force } moreover { fix G assume G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' D)" and c: "\X\\\: (F@G)\\<^sub>s\<^sub>t \ set constr" and e: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" have d_Q2: "Q2 (pair ` set D) X" unfolding Q2_def proof (intro allI impI) fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (pair ` set D)" then obtain d where d: "d \ set D" "Fun f T \ subterms (pair d)" by auto hence "fv (pair d) \ set X = {}" using prems(4) unfolding pair_def by force thus "T = [] \ (\s \ set T. s \ Var ` set X)" by (metis fv_disj_Fun_subterm_param_cases d(2)) qed have "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F' \ pair ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset[OF G] by auto hence "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G)) X" using Q2_subset[OF _ Q2_supset[OF e d_Q2]] by metis hence "tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: (F@G)\\<^sub>s\<^sub>t)" by (metis Q2_def tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)) } ultimately have 4: "list_all tfr\<^sub>s\<^sub>t\<^sub>p constr" when "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" using that Ball_set by blast have 5: "list_all tfr\<^sub>s\<^sub>t\<^sub>p constr" using 2 3 4 by metis show ?case using 1 5 A''(1) by simp qed qed lemma tr_tfr: assumes "A' \ set (tr A [])" and "tfr\<^sub>s\<^sub>s\<^sub>t A" and "fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" shows "tfr\<^sub>s\<^sub>t A'" proof - have *: "trms\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A" using tr_trms_subset[OF assms(1)] by simp hence "SMP (trms\<^sub>s\<^sub>t A') \ SMP (trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" using SMP_mono by simp moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A)" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t A')" by (metis tfr_subset(2)[OF _ *]) have **: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p A" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast have "pair ` setops\<^sub>s\<^sub>s\<^sub>t A \ SMP (trms\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t A) - Var`\" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs unfolding pair_def by auto hence ***: "\t \ pair`setops\<^sub>s\<^sub>s\<^sub>t A. \t' \ pair`setops\<^sub>s\<^sub>s\<^sub>t A. (\\. Unifier \ t t') \ \ t = \ t'" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>e\<^sub>t_def by blast have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p A'" using tr_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF assms(1) ** assms(3)] *** unfolding pair_def by fastforce show ?thesis by (metis 1 2 tfr\<^sub>s\<^sub>t_def) qed end end subsubsection \Theorem: The Stateful Typing Result\ context stateful_typing_result begin theorem stateful_typing_result: assumes "wf\<^sub>s\<^sub>s\<^sub>t \" and "tfr\<^sub>s\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t \)" and "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ \\<^sub>s \" obtains \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "\\<^sub>\ \\<^sub>s \" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - obtain \' where \': "\' \ set (tr \ [])" "\ \ \\'\" using tr_sem_equiv[of \] assms(1,4,5) by auto have *: "wf\<^sub>s\<^sub>t {} \'" "fv\<^sub>s\<^sub>t \' \ bvars\<^sub>s\<^sub>t \' = {}" "tfr\<^sub>s\<^sub>t \'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t \')" using tr_wf[OF \'(1) assms(1,3)] tr_tfr[OF \'(1) assms(2)] assms(1) by metis+ obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\{}; \'\\<^sub>d \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_attack_if_tfr_attack_d * Ana_invar_subst' assms(4) \'(2) unfolding constr_sem_d_def - by moura + by atomize_elim auto thus ?thesis using that tr_sem_equiv[of \] assms(1,3) \'(1) unfolding constr_sem_d_def by auto qed end subsection \Proving Type-Flaw Resistance Automatically\ definition pair' where "pair' pair_fun d \ case d of (t,t') \ Fun pair_fun [t,t']" fun comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ pair_fun (\_: t \ t'\) = (mgu t t' \ None \ \ t = \ t')" | "comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ pair_fun (\X\\\: F \\: F'\) = ( (F' = [] \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. is_Var (\ (Var x)))) \ (\u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair' pair_fun ` set F'). is_Fun u \ (args u = [] \ (\s \ set (args u). s \ Var ` set X))))" | "comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ _ _ = True" definition comp_tfr\<^sub>s\<^sub>s\<^sub>t where "comp_tfr\<^sub>s\<^sub>s\<^sub>t arity Ana \ pair_fun M S \ list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ pair_fun) S \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (trms_list\<^sub>s\<^sub>s\<^sub>t S) \ has_all_wt_instances_of \ (trms\<^sub>s\<^sub>s\<^sub>t S \ pair' pair_fun ` setops\<^sub>s\<^sub>s\<^sub>t S) M \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" locale stateful_typed_model' = stateful_typed_model arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" + assumes \_Var_fst': "\\ n m. \ (Var (\,n)) = \ (Var (\,m))" and Ana_const': "\c T. arity c = 0 \ Ana (Fun c T) = ([], [])" begin sublocale typed_model' by (unfold_locales, rule \_Var_fst', metis Ana_const', metis Ana_subst') lemma pair_code: "pair d = pair' Pair d" by (simp add: pair_def pair'_def) end locale stateful_typing_result' = stateful_typed_model' arity public Ana \ Pair + stateful_typing_result arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" begin lemma tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair a" proof (cases a) case (Equality ac t t') thus ?thesis using mgu_always_unifies[of t _ t'] mgu_gives_MGU[of t t'] by auto next case (NegChecks X F F') thus ?thesis using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(2)[of X F F'] comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(2)[of \ Pair X F F'] Fun_range_case(2)[of "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F')"] unfolding is_Var_def pair_code[symmetric] by auto qed auto lemma tfr\<^sub>s\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>s\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>s\<^sub>t arity Ana \ Pair M S" shows "tfr\<^sub>s\<^sub>s\<^sub>t S" unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def proof have comp_tfr\<^sub>s\<^sub>e\<^sub>t_M: "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" using assms unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def by blast have SMP_repr_M: "finite_SMP_representation arity Ana \ M" using comp_tfr\<^sub>s\<^sub>e\<^sub>t_M unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def by blast have wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S)" and S_trms_instance_M: "has_all_wt_instances_of \ (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S) M" using assms setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2)[of S] trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[of S] finite_SMP_representationD[OF SMP_repr_M] unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def list_all_iff pair_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] finite_SMP_representation_def by (meson, argo, argo) show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S \ pair ` setops\<^sub>s\<^sub>s\<^sub>t S)" using tfr_subset(3)[OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF comp_tfr\<^sub>s\<^sub>e\<^sub>t_M] SMP_SMP_subset] SMP_I'[OF wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M S_trms_instance_M] by blast have "list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair) S" by (metis assms comp_tfr\<^sub>s\<^sub>s\<^sub>t_def) thus "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p S" by (induct S) (simp_all add: tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p) qed lemma tfr\<^sub>s\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>s\<^sub>t': fixes S defines "M \ SMP0 Ana \ (trms_list\<^sub>s\<^sub>s\<^sub>t S@map pair (setops_list\<^sub>s\<^sub>s\<^sub>t S))" assumes comp_tfr: "comp_tfr\<^sub>s\<^sub>s\<^sub>t arity Ana \ Pair (set M) S" shows "tfr\<^sub>s\<^sub>s\<^sub>t S" by (rule tfr\<^sub>s\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>s\<^sub>t[OF comp_tfr[unfolded M_def]]) end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy b/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy @@ -1,2896 +1,2896 @@ (* Title: Strands_and_Constraints.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Strands and Symbolic Intruder Constraints\ theory Strands_and_Constraints imports Messages More_Unification Intruder_Deduction begin subsection \Constraints, Strands and Related Definitions\ datatype poscheckvariant = Assign ("assign") | Check ("check") text \ A strand (or constraint) step is either a message transmission (either a message being sent \Send\ or being received \Receive\) or a check on messages (a positive check \Equality\---which can be either an "assignment" or just a check---or a negative check \Inequality\) \ datatype (funs\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>s\<^sub>t\<^sub>p: 'b) strand_step = Send "('a,'b) term list" ("send\_\\<^sub>s\<^sub>t" 80) | Receive "('a,'b) term list" ("receive\_\\<^sub>s\<^sub>t" 80) | Equality poscheckvariant "('a,'b) term" "('a,'b) term" ("\_: _ \ _\\<^sub>s\<^sub>t" [80,80]) | Inequality (bvars\<^sub>s\<^sub>t\<^sub>p: "'b list") "(('a,'b) term \ ('a,'b) term) list" ("\_\\\: _\\<^sub>s\<^sub>t" [80,80]) where "bvars\<^sub>s\<^sub>t\<^sub>p (Send _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Receive _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Equality _ _ _) = []" abbreviation "Send1 t \ Send [t]" abbreviation "Receive1 t \ Receive [t]" text \ A strand is a finite sequence of strand steps (constraints and strands share the same datatype) \ type_synonym ('a,'b) strand = "('a,'b) strand_step list" type_synonym ('a,'b) strands = "('a,'b) strand set" abbreviation "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ \(t,t') \ set F. {t,t'}" fun trms\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) terms" where "trms\<^sub>s\<^sub>t\<^sub>p (Send ts) = set ts" | "trms\<^sub>s\<^sub>t\<^sub>p (Receive ts) = set ts" | "trms\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" lemma vars\<^sub>s\<^sub>t\<^sub>p_unfold[simp]: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x) \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto text \The set of terms occurring in a strand\ definition trms\<^sub>s\<^sub>t where "trms\<^sub>s\<^sub>t S \ \(trms\<^sub>s\<^sub>t\<^sub>p ` set S)" fun trms_list\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) term list" where "trms_list\<^sub>s\<^sub>t\<^sub>p (Send ts) = ts" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Receive ts) = ts" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = concat (map (\(t,t'). [t,t']) F)" text \The set of terms occurring in a strand (list variant)\ definition trms_list\<^sub>s\<^sub>t where "trms_list\<^sub>s\<^sub>t S \ remdups (concat (map trms_list\<^sub>s\<^sub>t\<^sub>p S))" text \The set of variables occurring in a sent message\ definition fv\<^sub>s\<^sub>n\<^sub>d::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>n\<^sub>d x \ case x of Send t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | _ \ {}" text \The set of variables occurring in a received message\ definition fv\<^sub>r\<^sub>c\<^sub>v::"('a,'b) strand_step \ 'b set" where "fv\<^sub>r\<^sub>c\<^sub>v x \ case x of Receive t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | _ \ {}" text \The set of variables occurring in an equality constraint\ definition fv\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s \ fv t else {} | _ \ {}" text \The set of variables occurring at the left-hand side of an equality constraint\ definition fv_l\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_l\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s else {} | _ \ {}" text \The set of variables occurring at the right-hand side of an equality constraint\ definition fv_r\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_r\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv t else {} | _ \ {}" text \The free variables of inequality constraints\ definition fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q::"('a,'b) strand_step \ 'b set" where "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ case x of Inequality X F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X | _ \ {}" fun fv\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>t\<^sub>p (Send t) = fv\<^sub>s\<^sub>e\<^sub>t (set t)" | "fv\<^sub>s\<^sub>t\<^sub>p (Receive t) = fv\<^sub>s\<^sub>e\<^sub>t (set t)" | "fv\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\(t,t') \ set F. fv t \ fv t') - set X" text \The set of free variables of a strand\ definition fv\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "fv\<^sub>s\<^sub>t S \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" text \The set of bound variables of a strand\ definition bvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "bvars\<^sub>s\<^sub>t S \ \(set (map (set \ bvars\<^sub>s\<^sub>t\<^sub>p) S))" text \The set of all variables occurring in a strand\ definition vars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "vars\<^sub>s\<^sub>t S \ \(set (map vars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x \ case x of Inequality _ _ \ {} | Equality Check _ _ \ {} | _ \ vars\<^sub>s\<^sub>t\<^sub>p x" text \The variables of a strand whose occurrences might be restricted by well-formedness constraints\ definition wfrestrictedvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t S \ \(set (map wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfvarsoccs\<^sub>s\<^sub>t\<^sub>p where "wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x \ case x of Send t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | Equality Assign s t \ fv s | _ \ {}" text \The variables of a strand that occur in sent messages or in assignments\ definition wfvarsoccs\<^sub>s\<^sub>t where "wfvarsoccs\<^sub>s\<^sub>t S \ \(set (map wfvarsoccs\<^sub>s\<^sub>t\<^sub>p S))" text \The variables occurring at the right-hand side of assignment steps\ fun assignment_rhs\<^sub>s\<^sub>t where "assignment_rhs\<^sub>s\<^sub>t [] = {}" | "assignment_rhs\<^sub>s\<^sub>t (Equality Assign t t'#S) = insert t' (assignment_rhs\<^sub>s\<^sub>t S)" | "assignment_rhs\<^sub>s\<^sub>t (x#S) = assignment_rhs\<^sub>s\<^sub>t S" text \The set of function symbols occurring in a strand\ definition funs\<^sub>s\<^sub>t::"('a,'b) strand \ 'a set" where "funs\<^sub>s\<^sub>t S \ \(set (map funs\<^sub>s\<^sub>t\<^sub>p S))" fun subst_apply_strand_step::"('a,'b) strand_step \ ('a,'b) subst \ ('a,'b) strand_step" (infix "\\<^sub>s\<^sub>t\<^sub>p" 51) where "Send t \\<^sub>s\<^sub>t\<^sub>p \ = Send (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" | "Receive t \\<^sub>s\<^sub>t\<^sub>p \ = Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" | "Equality a t t' \\<^sub>s\<^sub>t\<^sub>p \ = Equality a (t \ \) (t' \ \)" | "Inequality X F \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" text \Substitution application for strands\ definition subst_apply_strand::"('a,'b) strand \ ('a,'b) subst \ ('a,'b) strand" (infix "\\<^sub>s\<^sub>t" 51) where "S \\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>s\<^sub>t\<^sub>p \) S" text \The semantics of inequality constraints\ definition "ineq_model (\::('a,'b) subst) X F \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,t') \ set F. t \ \ \\<^sub>s \ \ t' \ \ \\<^sub>s \))" fun simple\<^sub>s\<^sub>t\<^sub>p where "simple\<^sub>s\<^sub>t\<^sub>p (Receive t) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Send [Var v]) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\\. ineq_model \ X F)" | "simple\<^sub>s\<^sub>t\<^sub>p _ = False" text \Simple constraints\ definition simple where "simple S \ list_all simple\<^sub>s\<^sub>t\<^sub>p S" text \The intruder knowledge of a constraint\ fun ik\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) terms" where "ik\<^sub>s\<^sub>t [] = {}" | "ik\<^sub>s\<^sub>t (Receive t#S) = set t \ (ik\<^sub>s\<^sub>t S)" | "ik\<^sub>s\<^sub>t (_#S) = ik\<^sub>s\<^sub>t S" text \Strand well-formedness\ fun wf\<^sub>s\<^sub>t::"'b set \ ('a,'b) strand \ bool" where "wf\<^sub>s\<^sub>t V [] = True" | "wf\<^sub>s\<^sub>t V (Receive ts#S) = (fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wf\<^sub>s\<^sub>t V S)" | "wf\<^sub>s\<^sub>t V (Send ts#S) = wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" | "wf\<^sub>s\<^sub>t V (Equality Assign s t#S) = (fv t \ V \ wf\<^sub>s\<^sub>t (V \ fv s) S)" | "wf\<^sub>s\<^sub>t V (Equality Check s t#S) = wf\<^sub>s\<^sub>t V S" | "wf\<^sub>s\<^sub>t V (Inequality _ _#S) = wf\<^sub>s\<^sub>t V S" text \Well-formedness of constraint states\ definition wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r::"('a,'b) strand \ ('a,'b) subst \ bool" where "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>t {} S \ subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {})" declare trms\<^sub>s\<^sub>t_def[simp] declare fv\<^sub>s\<^sub>n\<^sub>d_def[simp] declare fv\<^sub>r\<^sub>c\<^sub>v_def[simp] declare fv\<^sub>e\<^sub>q_def[simp] declare fv_l\<^sub>e\<^sub>q_def[simp] declare fv_r\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>s\<^sub>t_def[simp] declare vars\<^sub>s\<^sub>t_def[simp] declare bvars\<^sub>s\<^sub>t_def[simp] declare wfrestrictedvars\<^sub>s\<^sub>t_def[simp] declare wfvarsoccs\<^sub>s\<^sub>t_def[simp] lemmas wf\<^sub>s\<^sub>t_induct = wf\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsEq2 ConsIneq] lemmas ik\<^sub>s\<^sub>t_induct = ik\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsIneq] lemmas assignment_rhs\<^sub>s\<^sub>t_induct = assignment_rhs\<^sub>s\<^sub>t.induct[case_names Nil ConsEq2 ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lexicographical measure on strands\ definition size\<^sub>s\<^sub>t::"('a,'b) strand \ nat" where "size\<^sub>s\<^sub>t S \ size_list (\x. Max (insert 0 (size ` trms\<^sub>s\<^sub>t\<^sub>p x))) S" definition measure\<^sub>s\<^sub>t::"((('a, 'b) strand \ ('a,'b) subst) \ ('a, 'b) strand \ ('a,'b) subst) set" where "measure\<^sub>s\<^sub>t \ measures [\(S,\). card (fv\<^sub>s\<^sub>t S), \(S,\). size\<^sub>s\<^sub>t S]" lemma measure\<^sub>s\<^sub>t_alt_def: "((s,x),(t,y)) \ measure\<^sub>s\<^sub>t = (card (fv\<^sub>s\<^sub>t s) < card (fv\<^sub>s\<^sub>t t) \ (card (fv\<^sub>s\<^sub>t s) = card (fv\<^sub>s\<^sub>t t) \ size\<^sub>s\<^sub>t s < size\<^sub>s\<^sub>t t))" by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) lemma measure\<^sub>s\<^sub>t_trans: "trans measure\<^sub>s\<^sub>t" by (simp add: trans_def measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) subsubsection \Some lemmata\ lemma trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t: "trms\<^sub>s\<^sub>t S = set (trms_list\<^sub>s\<^sub>t S)" unfolding trms\<^sub>s\<^sub>t_def trms_list\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma subst_apply_strand_step_def: "s \\<^sub>s\<^sub>t\<^sub>p \ = (case s of Send t \ Send (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) | Receive t \ Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) | Equality a t t' \ Equality a (t \ \) (t' \ \) | Inequality X F \ Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by (cases s) simp_all lemma subst_apply_strand_nil[simp]: "[] \\<^sub>s\<^sub>t \ = []" unfolding subst_apply_strand_def by simp lemma finite_funs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (funs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_funs\<^sub>s\<^sub>t[simp]: "finite (funs\<^sub>s\<^sub>t S)" unfolding funs\<^sub>s\<^sub>t_def by simp lemma finite_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[simp]: "finite (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s x)" by (induct x) auto lemma finite_trms\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (trms\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (vars\<^sub>s\<^sub>t\<^sub>p x)" by auto lemma finite_bvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (set (bvars\<^sub>s\<^sub>t\<^sub>p x))" by rule lemma finite_fv\<^sub>s\<^sub>n\<^sub>d[simp]: "finite (fv\<^sub>s\<^sub>n\<^sub>d x)" by (cases x) auto lemma finite_fv\<^sub>r\<^sub>c\<^sub>v[simp]: "finite (fv\<^sub>r\<^sub>c\<^sub>v x)" by (cases x) auto lemma finite_fv\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (fv\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t[simp]: "finite (vars\<^sub>s\<^sub>t S)" by simp lemma finite_bvars\<^sub>s\<^sub>t[simp]: "finite (bvars\<^sub>s\<^sub>t S)" by simp lemma finite_fv\<^sub>s\<^sub>t[simp]: "finite (fv\<^sub>s\<^sub>t S)" by simp lemma finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfrestrictedvars\<^sub>s\<^sub>t[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t S)" using finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfvarsoccs\<^sub>s\<^sub>t[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t S)" using finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_ik\<^sub>s\<^sub>t[simp]: "finite (ik\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) simp_all lemma finite_assignment_rhs\<^sub>s\<^sub>t[simp]: "finite (assignment_rhs\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) simp_all lemma ik\<^sub>s\<^sub>t_is_rcv_set: "ik\<^sub>s\<^sub>t A = {t | ts t. Receive ts \ set A \ t \ set ts}" by (induct A rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_snoc_no_receive_eq: assumes "\ts. a = receive\ts\\<^sub>s\<^sub>t" shows "ik\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms unfolding ik\<^sub>s\<^sub>t_is_rcv_set by (metis (no_types, lifting) Un_iff append_Nil2 set_ConsD set_append) lemma ik\<^sub>s\<^sub>tD[dest]: "t \ ik\<^sub>s\<^sub>t S \ \ts. t \ set ts \ Receive ts \ set S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD'[dest]: "t \ ik\<^sub>s\<^sub>t S \ t \ trms\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD''[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_subterm_exD: assumes "t \ ik\<^sub>s\<^sub>t S" shows "\x \ set S. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" using assms ik\<^sub>s\<^sub>tD by force lemma assignment_rhs\<^sub>s\<^sub>tD[dest]: "t \ assignment_rhs\<^sub>s\<^sub>t S \ \t'. Equality Assign t' t \ set S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs\<^sub>s\<^sub>tD'[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma bvars\<^sub>s\<^sub>t_split: "bvars\<^sub>s\<^sub>t (S@S') = bvars\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S'" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma bvars\<^sub>s\<^sub>t_singleton: "bvars\<^sub>s\<^sub>t [x] = set (bvars\<^sub>s\<^sub>t\<^sub>p x)" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma strand_fv_bvars_disjointD: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" shows "set X \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" using assms by (induct S) fastforce+ lemma strand_fv_bvars_disjoint_unfold: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" "Inequality Y G \ set S" shows "set Y \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) = {}" proof - have "set X \ bvars\<^sub>s\<^sub>t S" "set Y \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set Y \ fv\<^sub>s\<^sub>t S" using strand_fv_bvars_disjointD[OF assms(1)] assms(2,3) by auto thus ?thesis using assms(1) by fastforce qed lemma strand_subst_hom[iff]: "(S@S') \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)@(S' \\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>s\<^sub>t \ = (x \\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>s\<^sub>t \)" unfolding subst_apply_strand_def by auto lemma strand_subst_comp: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ \\<^sub>s \ = ((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) have *: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ (set (bvars\<^sub>s\<^sub>t\<^sub>p x)) = {}" using Cons bvars\<^sub>s\<^sub>t_split[of "[x]" S] append_Cons inf_sup_absorb by (metis (no_types, lifting) Int_iff Un_commute disjoint_iff_not_equal self_append_conv2, metis append_self_conv2 bvars\<^sub>s\<^sub>t_singleton inf_bot_right inf_left_commute) hence IH: "S \\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \" using Cons.IH by auto have "(x#S \\<^sub>s\<^sub>t \ \\<^sub>s \) = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#(S \\<^sub>s\<^sub>t \ \\<^sub>s \)" by (metis strand_subst_hom(2)) hence "... = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" by (metis IH) hence "... = ((x \\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>t\<^sub>p \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" using rm_vars_comp[OF *(2)] proof (induction x) case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) qed (simp_all add: subst_apply_strand_step_def) thus ?case using IH by auto qed (simp add: subst_apply_strand_def) lemma strand_substI[intro]: "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof - show "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "vars\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Send ts) thus ?case by (induct ts) auto next case (Receive ts) thus ?case by (induct ts) auto next case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) show "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "fv\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Send ts) thus ?case by (induct ts) auto next case (Receive ts) thus ?case by (induct ts) auto next case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) qed lemma strand_substI': "fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" by (metis inf_bot_right strand_substI(1), metis inf_bot_right strand_substI(2)) lemma strand_subst_set: "(set (S \\<^sub>s\<^sub>t \)) = ((\x. x \\<^sub>s\<^sub>t\<^sub>p \) ` (set S))" by (auto simp add: subst_apply_strand_def) lemma strand_map_inv_set_snd_rcv_subst: assumes "finite (M::('a,'b) terms)" shows "set ((map Send1 (inv set M)) \\<^sub>s\<^sub>t \) = Send1 ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?A) "set ((map Receive1 (inv set M)) \\<^sub>s\<^sub>t \) = Receive1 ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?B) proof - { fix f::"('a,'b) term \ ('a,'b) strand_step" assume f: "f = Send1 \ f = Receive1" from assms have "set ((map f (inv set M)) \\<^sub>s\<^sub>t \) = f ` (M \\<^sub>s\<^sub>e\<^sub>t \)" proof (induction rule: finite_induct) case empty thus ?case unfolding inv_def by auto next case (insert m M) have "set (map f (inv set (insert m M)) \\<^sub>s\<^sub>t \) = insert (f m \\<^sub>s\<^sub>t\<^sub>p \) (set (map f (inv set M) \\<^sub>s\<^sub>t \))" by (simp add: insert.hyps(1) inv_set_fset subst_apply_strand_def) thus ?case using f insert.IH by auto qed } thus "?A" "?B" by auto qed lemma strand_ground_subst_vars_subset: assumes "ground (subst_range \)" shows "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x" using ground_subst_fv_subset[OF assms] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" proof (induction F) case (Cons f F) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) = fv (t \ ?\) \ fv (t' \ ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by (auto simp add: subst_apply_pairs_def) thus ?case using ground_subst_fv_subset[OF ground_subset[OF rm_vars_img_subset assms, of "set X"]] Cons.IH by (metis (no_types, lifting) Un_mono) qed (simp add: subst_apply_pairs_def) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto thus ?case using Cons.IH by auto qed (simp add: subst_apply_strand_def) lemma ik_union_subset: "\(P ` ik\<^sub>s\<^sub>t S) \ (\x \ (set S). \(P ` trms\<^sub>s\<^sub>t\<^sub>p x))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty[simp]: "ik\<^sub>s\<^sub>t (map Send X) = {}" by (induct "map Send X" arbitrary: X rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty'[simp]: "ik\<^sub>s\<^sub>t [Send t] = {}" by simp lemma ik_append[iff]: "ik\<^sub>s\<^sub>t (S@S') = ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S'" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_cons: "ik\<^sub>s\<^sub>t (x#S) = ik\<^sub>s\<^sub>t [x] \ ik\<^sub>s\<^sub>t S" using ik_append[of "[x]" S] by simp lemma assignment_rhs_append[iff]: "assignment_rhs\<^sub>s\<^sub>t (S@S') = assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t S'" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma eqs_rcv_map_empty: "assignment_rhs\<^sub>s\<^sub>t (map Receive M) = {}" by auto lemma ik_rcv_map: assumes "ts \ set L" shows "set ts \ ik\<^sub>s\<^sub>t (map Receive L)" proof - { fix L L' have "set ts \ ik\<^sub>s\<^sub>t [Receive ts]" by auto hence "set ts \ ik\<^sub>s\<^sub>t (map Receive L@Receive ts#map Receive L')" using ik_append by auto hence "set ts \ ik\<^sub>s\<^sub>t (map Receive (L@ts#L'))" by auto } thus ?thesis using assms split_list_last by force qed lemma ik_subst: "ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by (induct rule: ik\<^sub>s\<^sub>t_induct) auto lemma ik_rcv_map': assumes "t \ ik\<^sub>s\<^sub>t (map Receive L)" shows "\ts \ set L. t \ set ts" using assms by force lemma ik_append_subset[simp]: "ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t (S@S')" "ik\<^sub>s\<^sub>t S' \ ik\<^sub>s\<^sub>t (S@S')" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs_append_subset[simp]: "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (S@S')" "assignment_rhs\<^sub>s\<^sub>t S' \ assignment_rhs\<^sub>s\<^sub>t (S@S')" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma trms\<^sub>s\<^sub>t_cons: "trms\<^sub>s\<^sub>t (x#S) = trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by simp lemma trm_strand_subst_cong: "t \ trms\<^sub>s\<^sub>t S \ t \ \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\X F. Inequality X F \ set S \ t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" (is "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S") "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S) \ (\X F. Inequality X F \ set S \ (\t' \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \))" (is "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S") proof - show "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t S") case True hence "?P t \ S" using Cons by simp thus ?thesis by (cases x) (metis (no_types, lifting) Un_iff list.set_intros(2) strand_subst_hom(2) trms\<^sub>s\<^sub>t_cons)+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp show "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)") case True hence "?Q t \ S" using Cons by simp thus ?thesis by (cases x) force+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \" by (induct F) (force simp add: subst_apply_pairs_def)+ thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp qed subsection \Lemmata: Free Variables of Strands\ lemma fv_trm_snd_rcv[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Send ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Receive ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by simp_all lemma in_strand_fv_subset: "x \ set S \ vars\<^sub>s\<^sub>t\<^sub>p x \ vars\<^sub>s\<^sub>t S" by fastforce lemma in_strand_fv_subset_snd: "Send ts \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" by fastforce lemma in_strand_fv_subset_rcv: "Receive ts \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by fastforce lemma fv\<^sub>s\<^sub>n\<^sub>dE: assumes "v \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" obtains ts where "send\ts\\<^sub>s\<^sub>t \ set S" "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" proof - have "\ts. send\ts\\<^sub>s\<^sub>t \ set S \ v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>s\<^sub>n\<^sub>d_def strand_step.collapse(1)) thus ?thesis by (metis that) qed lemma fv\<^sub>r\<^sub>c\<^sub>vE: assumes "v \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" obtains ts where "receive\ts\\<^sub>s\<^sub>t \ set S" "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" proof - have "\ts. receive\ts\\<^sub>s\<^sub>t \ set S \ v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>r\<^sub>c\<^sub>v_def strand_step.collapse(2)) thus ?thesis by (metis that) qed lemma vars\<^sub>s\<^sub>t\<^sub>pI[intro]: "x \ fv\<^sub>s\<^sub>t\<^sub>p s \ x \ vars\<^sub>s\<^sub>t\<^sub>p s" by (induct s rule: fv\<^sub>s\<^sub>t\<^sub>p.induct) auto lemma vars\<^sub>s\<^sub>tI[intro]: "x \ fv\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>t\<^sub>pI by fastforce lemma fv\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t[simp]: "fv\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>tI by force lemma vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t: "vars\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case proof (induction x) case (Inequality X F) thus ?case by (induct F) auto qed auto qed simp lemma fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p: "x \ fv\<^sub>s\<^sub>t\<^sub>p a \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p a)" using var_is_subterm by (cases a) force+ lemma fv\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>t A \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p by (cases "x \ fv\<^sub>s\<^sub>t A") auto qed simp lemma vars_st_snd_map: "vars\<^sub>s\<^sub>t (map Send tss) = fv\<^sub>s\<^sub>e\<^sub>t (Fun f ` set tss)" by auto lemma vars_st_rcv_map: "vars\<^sub>s\<^sub>t (map Receive tss) = fv\<^sub>s\<^sub>e\<^sub>t (Fun f ` set tss)" by auto lemma vars_snd_rcv_union: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_union: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_empty[simp]: "fv\<^sub>s\<^sub>n\<^sub>d x = {} \ fv\<^sub>r\<^sub>c\<^sub>v x = {}" by (cases x) simp_all lemma vars_snd_rcv_strand[iff]: "vars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S))) \ bvars\<^sub>s\<^sub>t S" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. vars\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ set (bvars\<^sub>s\<^sub>t\<^sub>p s) \ V" by (metis vars_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma fv_snd_rcv_strand[iff]: "fv\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)))" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. fv\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ V" by (metis fv_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma vars_snd_rcv_strand2[iff]: "wfrestrictedvars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S)))" by (induct S) (auto simp add: split: strand_step.split poscheckvariant.split) lemma fv_snd_rcv_strand_subset[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" proof - show "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand[of S] by auto show "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) show "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) qed lemma vars_snd_rcv_strand_subset2[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q assign) S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma wfrestrictedvars\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma subst_sends_strand_step_fv_to_img: "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars ?\" proof (induction F) case (Cons f F) thus ?case using subst_sends_fv_to_img[of _ ?\] by (auto simp add: subst_apply_pairs_def) qed (auto simp add: subst_apply_pairs_def) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" using rm_vars_img_subset[of "set X" \] fv_set_mono unfolding range_vars_alt_def by blast+ thus ?thesis using Inequality by (auto simp add: subst_apply_strand_step_def) qed (auto simp add: subst_apply_strand_step_def) lemma subst_sends_strand_fv_to_img: "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) have *: "fv\<^sub>s\<^sub>t (x#S \\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>t (x#S) \ range_vars \ = fv\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>t S \ range_vars \" by auto thus ?case using Cons.IH subst_sends_strand_step_fv_to_img[of x \] by auto qed simp lemma ineq_apply_subst: assumes "subst_domain \ \ set X = {}" shows "(Inequality X F) \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using rm_vars_apply'[OF assms] by (simp add: subst_apply_strand_step_def) lemma fv_strand_step_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (P x)) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" proof (cases x) case (Send ts) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by auto thus ?thesis using assms Send subst_apply_fv_unfold[of _ \] by fastforce next case (Receive ts) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by auto thus ?thesis using assms Receive subst_apply_fv_unfold[of _ \] by fastforce next case (Equality ac' t t') show ?thesis proof (cases "ac = ac'") case True hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] True by auto next case False hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = {}" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] False by auto qed next case (Inequality X F) hence 1: "set X \ (subst_domain \ \ range_vars \) = {}" "x \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "rm_vars (set X) \ = \" using assms ineq_apply_subst[of \ X F] rm_vars_apply'[of \ "set X"] unfolding range_vars_alt_def by force+ have 2: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using Inequality by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" using fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq[OF 1(1), of "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] by simp hence 3: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" by (metis fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst) have 4: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using 1(2) by auto show ?thesis using assms(1) Inequality subst_apply_fv_unfold[of _ \] 1(2) 2 3 4 unfolding fv\<^sub>e\<^sub>q_def fv\<^sub>r\<^sub>c\<^sub>v_def fv\<^sub>s\<^sub>n\<^sub>d_def by (metis (no_types) Sup_empty image_empty fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps fv\<^sub>s\<^sub>e\<^sub>t.simps fv\<^sub>s\<^sub>t\<^sub>p.simps(4) strand_step.simps(20)) qed lemma fv_strand_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S)))) = \(set (map P (S \\<^sub>s\<^sub>t \)))" using assms(2) proof (induction S) case (Cons x S) hence *: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence **: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" using fv_strand_step_subst[OF assms(1), of x \] by auto have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) \ (\(set (map P ((S \\<^sub>s\<^sub>t \)))))" using Cons unfolding range_vars_alt_def bvars\<^sub>s\<^sub>t_def by force hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = P (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S))))" using ** by simp thus ?case using Cons.IH[OF *(1)] unfolding bvars\<^sub>s\<^sub>t_def by simp qed simp lemma fv_strand_subst2: assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (wfrestrictedvars\<^sub>s\<^sub>t S)) = wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis (no_types, lifting) assms fv\<^sub>s\<^sub>e\<^sub>t.simps vars_snd_rcv_strand2 fv_strand_subst UN_Un image_Un) lemma fv_strand_subst': assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>t S)) = fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis assms fv_strand_subst fv\<^sub>s\<^sub>t_def) lemma fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_in_fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast lemma trms\<^sub>s\<^sub>t_append: "trms\<^sub>s\<^sub>t (A@B) = trms\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>t B" by auto lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (a \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s a \\<^sub>s\<^sub>e\<^sub>t \" by (auto simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset: "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (force simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" shows "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix x assume "x \ fv t" hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" using fv_subset[OF assms] fv_subterms_set[of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast hence "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset by fast } thus ?thesis by (meson fv_subst_obtain_var subset_iff) qed lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_funs_term_cases: assumes "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "f \ funs_term t" shows "(\u \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term u) \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term (\ x))" using assms(1) proof (induction F) case (Cons g F) obtain s u where g: "g = (s,u)" by (metis surj_pair) show ?case proof (cases "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False thus ?thesis using assms(2) Cons.prems g funs_term_subst[of _ \] by (auto simp add: subst_apply_pairs_def) qed (use Cons.IH in fastforce) qed simp lemma trm\<^sub>s\<^sub>t\<^sub>p_subst: assumes "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" shows "trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" proof - have "rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p a)) \ = \" using assms by force thus ?thesis using assms by (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def split: strand_step.splits) qed lemma trms\<^sub>s\<^sub>t_subst: assumes "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" shows "trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction A) case (Cons a A) have 1: "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" using Cons.prems by auto hence IH: "trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ = trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using Cons.IH by simp have "trms\<^sub>s\<^sub>t (a#A) = trms\<^sub>s\<^sub>t\<^sub>p a \ trms\<^sub>s\<^sub>t A" by auto hence 2: "trms\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis image_Un) have "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \)) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" by (auto simp add: subst_apply_strand_def) hence 3: "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using trm\<^sub>s\<^sub>t\<^sub>p_subst[OF 1(2)] by auto show ?case using IH 2 3 by metis qed (simp add: subst_apply_strand_def) lemma strand_map_set_subst: assumes \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = (\(set (map trms\<^sub>s\<^sub>t\<^sub>p S))) \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons x S) hence "bvars\<^sub>s\<^sub>t [x] \ subst_domain \ = {}" "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence *: "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p x) = {}" "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = \(set (map trms\<^sub>s\<^sub>t\<^sub>p S)) \\<^sub>s\<^sub>e\<^sub>t \" using Cons.IH(1) bvars\<^sub>s\<^sub>t_singleton[of x] by auto hence "trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = (trms\<^sub>s\<^sub>t\<^sub>p x) \\<^sub>s\<^sub>e\<^sub>t \" proof (cases x) case (Inequality X F) thus ?thesis using rm_vars_apply'[of \ "set X"] * by (metis (no_types, lifting) image_cong trm\<^sub>s\<^sub>t\<^sub>p_subst) qed simp_all thus ?case using * subst_all_insert by auto qed simp lemma subst_apply_fv_subset_strand_trm: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and fv_sub: "fv t \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst[OF P \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand_trm2: assumes fv_sub: "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst2[OF \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and P_subset: "P x \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm[OF P _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by force } ultimately show ?thesis by metis next case (Receive ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm[OF P _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ \(set (map P S)) \ (V \ set X)" "fv t' \ \(set (map P S)) \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm[OF P _ assms(3)] by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by auto hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma subst_apply_fv_subset_strand2: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q \ P = fv_r\<^sub>e\<^sub>q ac" and P_subset: "P x \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm2[OF _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Receive ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm2[OF _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = fv t'" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm2[OF _ assms(3)] P by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma strand_subst_fv_bounded_if_img_bounded: assumes "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using subst_sends_strand_fv_to_img[of S \] assms by blast lemma strand_fv_subst_subset_if_subst_elim: assumes "subst_elim \ v" and "v \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof (cases "v \ fv\<^sub>s\<^sub>t S") case True thus ?thesis proof (induction S) case (Cons x S) have *: "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using assms(1) proof (cases x) case (Inequality X F) hence "subst_elim (rm_vars (set X) \) v \ v \ set X" using assms(1) by blast moreover have "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) - set X" using Inequality by auto ultimately have "v \ fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_elim_def subst_apply_pairs_def) thus ?thesis using Inequality by simp qed (simp_all add: subst_elim_def) moreover have "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using Cons.IH proof (cases "v \ fv\<^sub>s\<^sub>t S") case False moreover have "v \ range_vars \" by (simp add: subst_elimD''[OF assms(1)] range_vars_alt_def) ultimately show ?thesis by (meson UnE subsetCE subst_sends_strand_fv_to_img) qed simp ultimately show ?case by auto qed simp next case False thus ?thesis using assms fv_strand_subst' unfolding subst_elim_def by (metis (mono_tags, opaque_lifting) fv\<^sub>s\<^sub>e\<^sub>t.simps imageE mem_simps(8) eval_term.simps(1)) qed lemma strand_fv_subst_subset_if_subst_elim': assumes "subst_elim \ v" "v \ fv\<^sub>s\<^sub>t S" "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using strand_fv_subst_subset_if_subst_elim[OF assms(1)] assms(2) strand_subst_fv_bounded_if_img_bounded[OF assms(3)] by blast lemma fv_ik_is_fv_rcv: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) = \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_ik_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_assignment_rhs_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) force+ lemma fv_ik_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_var_is_fv: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ x \ fv\<^sub>s\<^sub>t A" by (meson fv_ik_subset_fv_st'[of A] fv_subset_subterms subsetCE term.set_intros(3)) lemma fv_assignment_rhs_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>s\<^sub>t A" using fv_ik_subset_fv_st[of A] fv_assignment_rhs_subset_fv_st[of A] by simp+ lemma strand_step_id_subst[iff]: "x \\<^sub>s\<^sub>t\<^sub>p Var = x" by (cases x) auto lemma strand_id_subst[iff]: "S \\<^sub>s\<^sub>t Var = S" using strand_step_id_subst by (induct S) auto lemma strand_subst_vars_union_bound[simp]: "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) define \' where "\' \ rm_vars (set X) \" have 0: "range_vars \' \ range_vars \" using rm_vars_img[of "set X" \] by (auto simp add: \'_def subst_domain_def range_vars_alt_def) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: \'_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" proof (induction F) case (Cons f G) - obtain t t' where f: "f = (t,t')" by moura + obtain t t' where f: "f = (t,t')" by atomize_elim auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') = fv (t \ \') \ fv (t' \ \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \')" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (auto simp add: subst_apply_pairs_def) thus ?case using 0 Cons.IH subst_sends_fv_to_img[of t \'] subst_sends_fv_to_img[of t' \'] unfolding f by auto qed (simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto ultimately show ?case by auto qed simp lemma strand_vars_split: "vars\<^sub>s\<^sub>t (S@S') = vars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S'" "wfrestrictedvars\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S'" "fv\<^sub>s\<^sub>t (S@S') = fv\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S'" by auto lemma bvars_subst_ident: "bvars\<^sub>s\<^sub>t S = bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" unfolding bvars\<^sub>s\<^sub>t_def by (induct S) (simp_all add: subst_apply_strand_step_def split: strand_step.splits) lemma strand_subst_subst_idem: assumes "subst_idem \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" "subst_domain \ \ fv\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" shows "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" and "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" proof - from assms(2,3) have "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ subst_domain \ = {}" using subst_sends_strand_fv_to_img[of S \] by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" by (metis assms(1,4,5) bvars_subst_ident strand_subst_comp subst_idem_def) qed lemma strand_subst_img_bound: assumes "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "subst_domain \ \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" by (metis (no_types) fv\<^sub>s\<^sub>t_def Un_subset_iff assms(1)) thus ?thesis unfolding range_vars_alt_def fv\<^sub>s\<^sub>t_def by (metis subst_range.simps fv_set_mono fv_strand_subst Int_commute assms(2) image_Un le_iff_sup) qed lemma strand_subst_img_bound': assumes "subst_domain \ \ range_vars \ \ vars\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "(subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)) \ vars\<^sub>s\<^sub>t S = subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)" using assms(1) by (metis inf.absorb_iff1 range_vars_alt_def subst_range.simps) hence "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using vars_snd_rcv_strand fv_snd_rcv_strand assms(2) strand_subst_img_bound unfolding range_vars_alt_def by (metis (no_types) inf_le2 inf_sup_distrib1 subst_range.simps sup_bot.right_neutral) thus "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis fv_snd_rcv_strand le_supI1 vars_snd_rcv_strand) qed lemma strand_subst_all_fv_subset: assumes "fv t \ fv\<^sub>s\<^sub>t S" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms by (metis fv_strand_subst' Int_commute subst_apply_fv_subset) lemma strand_subst_not_dom_fixed: assumes "v \ fv\<^sub>s\<^sub>t S" and "v \ subst_domain \" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons x S') have 1: "\X. v \ subst_domain (rm_vars (set X) \)" using Cons.prems(2) rm_vars_dom_subset by force show ?case proof (cases "v \ fv\<^sub>s\<^sub>t S'") case True thus ?thesis using Cons.IH[OF _ Cons.prems(2)] by auto next case False hence 2: "v \ fv\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems(1) by simp hence "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(2) subst_not_dom_fixed proof (cases x) case (Inequality X F) hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using 2 by simp hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using subst_not_dom_fixed[OF _ 1] by (induct F) (auto simp add: subst_apply_pairs_def) thus ?thesis using Inequality 2 by auto qed (force simp add: subst_domain_def)+ thus ?thesis by auto qed qed simp lemma strand_vars_unfold: "v \ vars\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ vars\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ vars\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ vars\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma strand_fv_unfold: "v \ fv\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ fv\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ fv\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ fv\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma subterm_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ \ts. Receive ts \ set S \ t \\<^sub>s\<^sub>e\<^sub>t set ts" by (induct S rule: ik\<^sub>s\<^sub>t_induct) auto lemma fv_subset_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ fv t \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" proof - assume "t \ ik\<^sub>s\<^sub>t S" then obtain ts where "Receive ts \ set S" "t \\<^sub>s\<^sub>e\<^sub>t set ts" by (metis subterm_if_in_strand_ik) hence "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using subtermeq_vars_subset by auto thus ?thesis using in_strand_fv_subset_rcv[OF \Receive ts \ set S\] by auto qed lemma fv_subset_if_in_strand_ik': "t \ ik\<^sub>s\<^sub>t S \ fv t \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] fv_snd_rcv_strand_subset(2)[of S] by blast lemma vars_subset_if_in_strand_ik2: "t \ ik\<^sub>s\<^sub>t S \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] vars_snd_rcv_strand_subset2(2)[of S] by blast subsection \Lemmata: Simple Strands\ lemma simple_Cons[dest]: "simple (s#S) \ simple S" unfolding simple_def by auto lemma simple_split[dest]: assumes "simple (S@S')" shows "simple S" "simple S'" using assms unfolding simple_def by auto lemma simple_append[intro]: "\simple S; simple S'\ \ simple (S@S')" unfolding simple_def by auto lemma simple_append_sym[sym]: "simple (S@S') \ simple (S'@S)" by auto lemma not_simple_if_snd_fun: "Fun f T \ set ts \ S = S'@Send ts#S'' \ \simple S" by (metis simple_def list.set_cases list_all_append list_all_simps(1) simple\<^sub>s\<^sub>t\<^sub>p.simps(5,6)) lemma not_list_all_elim: "\list_all P A \ \B x C. A = B@x#C \ \P x \ list_all P B" proof (induction A rule: List.rev_induct) case (snoc a A) show ?case proof (cases "list_all P A") case True thus ?thesis using snoc.prems by auto next case False then obtain B x C where "A = B@x#C" "\P x" "list_all P B" using snoc.IH[OF False] by auto thus ?thesis by auto qed qed simp lemma not_simple\<^sub>s\<^sub>t\<^sub>p_elim: assumes "\simple\<^sub>s\<^sub>t\<^sub>p x" shows "(\ts. x = Send ts \ (\y. ts = [Var y])) \ (\a t t'. x = Equality a t t') \ (\X F. x = Inequality X F \ (\\. ineq_model \ X F))" using assms by (cases x) (fastforce elim: simple\<^sub>s\<^sub>t\<^sub>p.elims)+ lemma not_simple_elim: assumes "\simple S" shows "(\A B ts. S = A@Send ts#B \ (\x. ts = [Var x]) \ simple A) \ (\A B a t t'. S = A@Equality a t t'#B \ simple A) \ (\A B X F. S = A@Inequality X F#B \ (\\. ineq_model \ X F) \ simple A)" by (metis assms not_list_all_elim not_simple\<^sub>s\<^sub>t\<^sub>p_elim simple_def) lemma simple_snd_is_var: "\Send ts \ set S; simple S\ \ \v. ts = [Var v]" unfolding simple_def by (metis list_all_append list_all_simps(1) simple\<^sub>s\<^sub>t\<^sub>p.elims(2) split_list_first strand_step.distinct(1) strand_step.distinct(5) strand_step.inject(1)) subsection \Lemmata: Strand Measure\ lemma measure\<^sub>s\<^sub>t_wellfounded: "wf measure\<^sub>s\<^sub>t" unfolding measure\<^sub>s\<^sub>t_def by simp lemma strand_size_append[iff]: "size\<^sub>s\<^sub>t (S@S') = size\<^sub>s\<^sub>t S + size\<^sub>s\<^sub>t S'" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_map_fun_lt[simp]: "size\<^sub>s\<^sub>t (map Send1 X) < size (Fun f X)" "size\<^sub>s\<^sub>t (map Send1 X) < size\<^sub>s\<^sub>t [Send [Fun f X]]" "size\<^sub>s\<^sub>t (map Receive1 X) < size\<^sub>s\<^sub>t [Receive [Fun f X]]" "size\<^sub>s\<^sub>t [Send X] < size\<^sub>s\<^sub>t [Send [Fun f X]]" "size\<^sub>s\<^sub>t [Receive X] < size\<^sub>s\<^sub>t [Receive [Fun f X]]" by (induct X) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_rm_fun_lt[simp]: "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Send ts#S')" "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Receive ts#S')" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_fv_card_map_fun_eq: "card (fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S')) = card (fv\<^sub>s\<^sub>t (S@(map Send1 X)@S'))" proof - have "fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S') = fv\<^sub>s\<^sub>t (S@(map Send1 X)@S')" by auto thus ?thesis by simp qed lemma strand_fv_card_rm_fun_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S'))" by (force intro: card_mono) lemma strand_fv_card_rm_eq_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by (force intro: card_mono) subsection \Lemmata: Well-formed Strands\ lemma wf_prefix[dest]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_vars_mono[simp]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (V \ W) S" proof (induction S arbitrary: V) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ W) S" using Cons.prems(1) Cons.IH by simp thus ?thesis by (metis Send sup_assoc sup_commute wf\<^sub>s\<^sub>t.simps(3)) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ W) S" "fv t' \ V \ W" using Equality Cons.prems(1) Cons.IH by auto thus ?thesis using Equality Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using Equality Cons by auto qed qed auto qed simp lemma wf\<^sub>s\<^sub>tI[intro]: "wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t V S" "V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) = V" using Cons by auto thus ?thesis using Send by simp next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t V S" "fv t' \ V" using Equality Cons by auto thus ?thesis using wf_vars_mono Equality Assign by simp next case Check thus ?thesis using Equality Cons by auto qed qed simp_all qed simp lemma wf\<^sub>s\<^sub>tI'[intro]: "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S) \ \(fv_r\<^sub>e\<^sub>q assign ` set S) \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Equality a t t') thus ?thesis using Cons by (cases a) auto qed simp_all qed simp lemma wf_append_exec: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" proof (induction S arbitrary: V) case (Cons x S V) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Cons.prems Cons.IH by simp thus ?thesis using Send by (auto simp add: sup_assoc) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Assign by (auto simp add: sup_assoc) next case Check hence "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Check by (auto simp add: sup_assoc) qed qed auto qed simp lemma wf_append_suffix: "wf\<^sub>s\<^sub>t V S \ wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V ts S) hence *: "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using ConsSnd.prems(2) by fastforce thus ?case using ConsSnd.IH * by simp next case (ConsRcv V ts S) hence *: "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" "wf\<^sub>s\<^sub>t V S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using ConsRcv.prems(2) by fastforce thus ?case using ConsRcv.IH * by simp next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all moreover have "vars\<^sub>s\<^sub>t\<^sub>p (Equality Assign t t') = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t t'#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) by blast thus ?case using ConsEq.IH * by simp qed (simp_all add: wf\<^sub>s\<^sub>tI) lemma wf_append_suffix': assumes "wf\<^sub>s\<^sub>t V S" and "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S') \ \(fv_r\<^sub>e\<^sub>q assign ` set S') \ wfvarsoccs\<^sub>s\<^sub>t S \ V" shows "wf\<^sub>s\<^sub>t V (S@S')" using assms proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V ts S) hence *: "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (send\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using ConsSnd.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsSnd.IH[OF *] by auto next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (\assign: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsEq.IH[OF *(2)] *(1) by auto qed (auto simp add: wf\<^sub>s\<^sub>tI') lemma wf_send_compose: "wf\<^sub>s\<^sub>t V (S@(map Send1 X)@S') = wf\<^sub>s\<^sub>t V (S@Send1 (Fun f X)#S')" proof (induction S arbitrary: V) case Nil thus ?case proof (induction X arbitrary: V) case (Cons y Y) thus ?case by (simp add: sup_assoc) qed simp next case (Cons s S) thus ?case proof (cases s) case (Equality ac t t') thus ?thesis using Cons by (cases ac) auto qed auto qed lemma wf_snd_append[iff]: "wf\<^sub>s\<^sub>t V (S@[Send t]) = wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_snd_append': "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (Send t#S)" by simp lemma wf_rcv_append[dest]: "wf\<^sub>s\<^sub>t V (S@Receive t#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_rcv_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Receive ts#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V t' S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by (simp, fastforce) thus ?case using ConsRcv by auto next case (ConsEq V t' t'' S) hence "fv t'' \ V" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t' t''#S) = fv t' \ fv t'' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t')" using ConsEq.prems(2) by blast thus ?case using ConsEq by auto qed auto lemma wf_rcv_append''[intro]: "\wf\<^sub>s\<^sub>t V S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ wf\<^sub>s\<^sub>t V (S@[Receive ts])" by (induct S) (simp, metis vars_snd_rcv_strand_subset2(1) append_Nil2 le_supI1 order_trans wf_rcv_append') lemma wf_rcv_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Receive ts])" by (simp add: wf_rcv_append'[of _ _ "[]"]) lemma wf_eq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality a t t'#S') \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (Nil V) hence "wf\<^sub>s\<^sub>t (V \ fv t) S'" by (cases a) auto moreover have "V \ fv t = V" using Nil by auto ultimately show ?case by simp next case (ConsRcv V us S) hence "wf\<^sub>s\<^sub>t V (S @ Equality a t t' # S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv\<^sub>s\<^sub>e\<^sub>t (set us) \ V" by fastforce+ hence "wf\<^sub>s\<^sub>t V (S@S')" using ConsRcv.IH by auto thus ?case using \fv\<^sub>s\<^sub>e\<^sub>t (set us) \ V\ by simp next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@Equality a t t'#S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv u)" "fv u' \ V" by auto hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" using ConsEq.IH by auto thus ?case using \fv u' \ V\ by simp qed auto lemma wf_eq_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Equality a t t'#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t V (S@S')" by auto thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V us S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us)) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us))" by fastforce+ thus ?case using ConsSnd by auto qed auto lemma wf_eq_append''[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t']@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V us S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us)) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us))" by auto thus ?case using ConsSnd by auto qed auto lemma wf_eq_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t'])" by (simp add: wf_eq_append'[of _ _ "[]"]) lemma wf_eq_check_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_eq_check_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_eq_check_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Equality Check t t'])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Inequality X F#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_ineq_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Inequality X F#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Inequality X F])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_Receive1_prefix: assumes "wf\<^sub>s\<^sub>t X S" and "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" shows "wf\<^sub>s\<^sub>t X (map Receive1 ts@S)" using assms by (induct ts) simp_all lemma wf_Send1_prefix: assumes "wf\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" shows "wf\<^sub>s\<^sub>t X (map Send1 ts@S)" using assms by (induct ts arbitrary: X) (simp, force simp add: Un_assoc) lemma wf_rcv_fv_single[elim]: "wf\<^sub>s\<^sub>t V (Receive ts#S') \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp lemma wf_rcv_fv: "wf\<^sub>s\<^sub>t V (S@Receive ts#S') \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S \ V" proof (induction S arbitrary: V) case (Cons a S) thus ?case by (cases a) (auto split!: poscheckvariant.split) qed simp lemma wf_eq_fv: "wf\<^sub>s\<^sub>t V (S@Equality Assign t t'#S') \ fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" proof (induction S arbitrary: V) case (Cons a S) thus ?case by (cases a) (auto split!: poscheckvariant.split) qed simp lemma wf_simple_fv_occurrence: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "v \ wfrestrictedvars\<^sub>s\<^sub>t S" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f \ v \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e" using assms proof (induction S rule: List.rev_induct) case (snoc x S) from \wf\<^sub>s\<^sub>t {} (S@[x])\ have "wf\<^sub>s\<^sub>t {} S" "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]" using wf_append_exec[THEN wf_vars_mono, of "{}" S "[x]" "wfrestrictedvars\<^sub>s\<^sub>t S - wfvarsoccs\<^sub>s\<^sub>t S"] vars_snd_rcv_strand_subset2(4)[of S] Diff_partition[of "wfvarsoccs\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>s\<^sub>t S"] by auto from \simple (S@[x])\ have "simple S" "simple\<^sub>s\<^sub>t\<^sub>p x" unfolding simple_def by auto show ?case proof (cases "v \ wfrestrictedvars\<^sub>s\<^sub>t S") case False show ?thesis proof (cases x) case (Receive ts) hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S" using \wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]\ by simp hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ \x = Receive ts\ by auto thus ?thesis using \x = Receive ts\ snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] by fastforce next case (Send ts) hence "v \ vars\<^sub>s\<^sub>t\<^sub>p x" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ False by auto from Send obtain w where "ts = [Var w]" using \simple\<^sub>s\<^sub>t\<^sub>p x\ by (cases ts) (simp, metis in_set_conv_decomp simple_snd_is_var snoc.prems(2)) hence "v = w" using \x = Send ts\ \v \ vars\<^sub>s\<^sub>t\<^sub>p x\ by simp thus ?thesis using \x = Send ts\ \v \ wfrestrictedvars\<^sub>s\<^sub>t S\ \ts = [Var w]\ by auto next case (Equality ac t t') thus ?thesis using snoc.prems(2) unfolding simple_def by auto next case (Inequality t t') thus ?thesis using False snoc.prems(3) by auto qed qed (use snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] in fastforce) qed simp lemma Unifier_strand_fv_subset: assumes g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" by (metis (no_types) fv_subset_if_in_strand_ik[OF g_in_ik] disj \ fv_strand_subst subst_apply_fv_subset) lemma wf\<^sub>s\<^sub>t_induct'[consumes 1, case_names Nil ConsSnd ConsRcv ConsEq ConsEq2 ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "P []" "\ts S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Send ts])" "\ts S. \wf\<^sub>s\<^sub>t V S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Receive ts])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S; fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Equality Assign t t'])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Equality Check t t'])" "\X F S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: List.rev_induct) case (snoc x S) hence *: "wf\<^sub>s\<^sub>t V S" "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) [x]" by (metis wf_prefix, metis wf_append_exec) have IH: "P S" using snoc.IH[OF *(1)] snoc.prems by auto note ** = snoc.prems(3,4,5,6,7)[OF *(1) IH] *(2) show ?case using **(1,2,4,5,6) proof (cases x) case (Equality ac t t') then show ?thesis using **(3,4,6) by (cases ac) auto qed auto qed simp lemma wf_subst_apply: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" proof (induction S arbitrary: V rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V ts S) hence "wf\<^sub>s\<^sub>t V S" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp_all hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsRcv.IH subst_apply_fv_subset by (simp, force) thus ?case by simp next case (ConsSnd V ts S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)))) (S \\<^sub>s\<^sub>t \)" using ConsSnd.IH by metis hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts))) (S \\<^sub>s\<^sub>t \)" by simp hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) (S \\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_unfold_set) thus ?case by simp next case (ConsEq V t t' S) hence "wf\<^sub>s\<^sub>t (V \ fv t) S" "fv t' \ V" by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>t \)" and *: "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsEq.IH subst_apply_fv_subset by force+ hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>t \)" using subst_apply_fv_union by metis thus ?case using * by simp qed simp_all lemma wf_unify: assumes wf: "wf\<^sub>s\<^sub>t V (S@Send [Fun f X]#S')" and g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case (snoc x S' V) have fun_fv_bound: "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" using snoc.prems(4) bvars\<^sub>s\<^sub>t_split Unifier_strand_fv_subset[OF g_in_ik \] by auto hence "fv (Fun f X \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" using fv_ik_is_fv_rcv by metis hence "fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using fv_ik_subset_fv_st[of "S \\<^sub>s\<^sub>t \"] by blast hence *: "fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by fastforce from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Send [Fun f X]#S')" using wf_prefix[of V "S@Send [Fun f X]#S'" "[x]"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using snoc.IH[OF _ snoc.prems(2,3)] snoc.prems(4) by auto from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')) [x]" using wf_append_exec[of V "(S@Send [Fun f X]#S')" "[x]"] by simp from snoc.prems(4) have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" by auto show ?case proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive t) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send [Fun f X]#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis (no_types) inf_sup_aci(5) subst_apply_fv_subset_strand2 subst_apply_fv_union disj') hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv\<^sub>s\<^sub>e\<^sub>t (set t \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) " using \x = Receive t\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)])" using wf_rcv_append'''[OF **, of "t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by simp thus ?thesis using \x = Receive t\ by auto next case (Equality ac s s') show ?thesis proof (cases ac) case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using Equality *** by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send [Fun f X]#S'"] by blast hence "fv s' \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv\<^sub>e\<^sub>q ac" ac x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac s s'\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Equality ac (s \ \) (s' \ \)])" using wf_eq_append'''[OF **] by metis thus ?thesis using \x = Equality ac s s'\ by auto next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality t t') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (auto dest: wf_subst_apply) lemma wf_equality: assumes wf: "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" and \: "mgu t t' = Some \" and disj: "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case Nil thus ?case using wf_prefix[of V S "[Equality ac t t']"] wf_subst_apply[of V S \] by auto next case (snoc x S' V) show ?case proof (cases ac) case Assign hence "fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S" using wf_eq_fv[of V, of S t t' "S'@[x]"] snoc by auto hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t S" using vars_snd_rcv_strand_subset2(4)[of S] by blast hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by force moreover have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" using snoc.prems(3) by auto ultimately have "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis inf_sup_aci(5) subst_apply_fv_subset_strand_trm2) moreover have "fv (t \ \) = fv (t' \ \)" by (metis MGU_is_Unifier[OF mgu_gives_MGU[OF \]]) ultimately have *: "fv (t \ \) \ fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by simp from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" using wf_prefix[of V "S@Equality ac t t'#S'"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" by (metis snoc.IH \ disj'(3)) from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')) [x]" using wf_append_exec[of V "(S@Equality ac t t'#S')" "[x]"] by simp show ?thesis proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive s) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of fv\<^sub>s\<^sub>t\<^sub>p] by (metis (no_types) inf_sup_aci(5) subst_apply_fv_union disj'(1,2)) hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" when "ac = Assign" using * that by blast hence "fv\<^sub>s\<^sub>e\<^sub>t (set s \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V))" when "ac = Assign" using \x = Receive s\ that by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (s \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)])" when "ac = Assign" using wf_rcv_append'''[OF **, of "s \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] that by simp thus ?thesis using \x = Receive s\ Assign by auto next case (Equality ac' s s') show ?thesis proof (cases ac') case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** Equality by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv s' \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac' x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac' (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv_r\<^sub>e\<^sub>q ac'" ac' x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * \ac = Assign\ by blast hence ****: "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac' s s'\ \ac = Assign\ by auto thus ?thesis using \x = Equality ac' s s'\ ** **** wf_eq_append' \ac = Assign\ by (metis (no_types, lifting) append.assoc append_Nil2 strand_step.case(3) strand_subst_hom subst_apply_strand_step_def) next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality s s') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (metis snoc.prems(1) wf_eq_check_append wf_subst_apply) qed lemma wf_rcv_prefix_ground: "wf\<^sub>s\<^sub>t {} ((map Receive M)@S) \ vars\<^sub>s\<^sub>t (map Receive M) = {}" by (induct M) auto lemma simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d: assumes "simple S" shows "wfvarsoccs\<^sub>s\<^sub>t S = \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" using assms unfolding simple_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma wf\<^sub>s\<^sub>t_simple_induct[consumes 2, case_names Nil ConsSnd ConsRcv ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "simple S" "P []" "\v S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Send [Var v]])" "\ts S. \wf\<^sub>s\<^sub>t V S; simple S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ P (S@[Receive ts])" "\X F S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd t S) hence "P S" by auto obtain v where "t = [Var v]" using simple_snd_is_var[OF _ \simple (S@[Send t])\] by auto thus ?case using ConsSnd.prems(3)[OF \wf\<^sub>s\<^sub>t V S\ _ \P S\] \simple (S@[Send t])\ by auto next case (ConsRcv t S) thus ?case using simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d[of "S@[Receive t]"] by auto qed (auto simp add: simple_def) lemma wf_trm_stp_dom_fv_disjoint: "\wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \; t \ trms\<^sub>s\<^sub>t S\ \ subst_domain \ \ fv t = {}" unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by force lemma wf_constr_bvars_disj: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce lemma wf_constr_bvars_disj': assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" shows "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" (is ?A) and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = {}" (is ?B) proof - have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using assms(1) unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce+ thus ?A and ?B using assms(2) bvars_subst_ident[of S \] by blast+ qed lemma (in intruder_model) wf_simple_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" shows "\v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ = \ v \ \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ = \ w)" (is "?P S") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ w") case True thus ?thesis using ConsSnd.IH by fastforce next case False thus ?thesis using ConsSnd.prems by auto qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) vars_snd_rcv_strand_subset2(1) by force hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" using ConsRcv.prems(1) by fastforce hence "?P S" by (metis ConsRcv.IH) thus ?case by fastforce next case (ConsIneq X F S) moreover have "wfrestrictedvars\<^sub>s\<^sub>t (S @ [Inequality X F]) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "?P S" by blast thus ?case by fastforce qed simp lemma (in intruder_model) wf_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ \ \ w) \ ((\t'. S = S\<^sub>p\<^sub>r\<^sub>e@Send t'#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \\<^sub>s\<^sub>e\<^sub>t set t' \\<^sub>s\<^sub>e\<^sub>t \) \ (\t' t''. S = S\<^sub>p\<^sub>r\<^sub>e@Equality Assign t' t''#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ t' \ \))" (is "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. ?P S\<^sub>p\<^sub>r\<^sub>e \ ?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd ts' S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ w") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsSnd.IH by moura + using ConsSnd.IH by atomize_elim auto thus ?thesis by fastforce next case False then obtain v where v: "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts')" "t \ \ \ \ v" using ConsSnd.prems by auto then obtain t' where t': "t' \ set ts'" "v \ fv t'" by auto have "t \ \ \ t' \ \" using v(2) t'(2) subst_mono[of "Var v" t' \] vars_iff_subtermeq[of v] term.order_trans by auto hence "t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts' \\<^sub>s\<^sub>e\<^sub>t \" using v(1) t'(1) by blast thus ?thesis using False v by auto qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsRcv.prems by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsRcv.IH by moura + using ConsRcv.IH by atomize_elim auto thus ?case by fastforce next case (ConsEq s s' S) have *: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsEq.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsEq.IH by moura + using ConsEq.IH by atomize_elim auto thus ?thesis by fastforce next case False then obtain v where "v \ fv s" "t \ \ \ \ v" using ConsEq.prems * by auto hence "t \ \ \ s \ \" using vars_iff_subtermeq[of v s] subst_mono[of "Var v" s \] term.order_trans by auto thus ?thesis using False by fastforce qed next case (ConsEq2 s s' S) have "wfrestrictedvars\<^sub>s\<^sub>t (S@[Equality Check s s']) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsEq2.prems by metis then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsEq2.IH by moura + using ConsEq2.IH by atomize_elim auto thus ?case by fastforce next case (ConsIneq X F S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" - using ConsIneq.IH by moura + using ConsIneq.IH by atomize_elim auto thus ?case by fastforce qed simp subsection \Constraint Semantics\ context intruder_model begin subsubsection \Definitions\ text \The constraint semantics in which the intruder is limited to composition only\ fun strand_sem_c::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c") where "\M; []\\<^sub>c = (\\. True)" | "\M; Send ts#S\\<^sub>c = (\\. (\t \ set ts. M \\<^sub>c t \ \) \ \M; S\\<^sub>c \)" | "\M; Receive ts#S\\<^sub>c = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>c \)" | "\M; Equality _ t t'#S\\<^sub>c = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c \)" | "\M; Inequality X F#S\\<^sub>c = (\\. ineq_model \ X F \ \M; S\\<^sub>c \)" definition constr_sem_c ("_ \\<^sub>c \_,_\") where "\ \\<^sub>c \S,\\ \ (\ supports \ \ \{}; S\\<^sub>c \)" abbreviation constr_sem_c' ("_ \\<^sub>c \_\" 90) where "\ \\<^sub>c \S\ \ \ \\<^sub>c \S,Var\" text \The full constraint semantics\ fun strand_sem_d::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d") where "\M; []\\<^sub>d = (\\. True)" | "\M; Send ts#S\\<^sub>d = (\\. (\t \ set ts. M \ t \ \) \ \M; S\\<^sub>d \)" | "\M; Receive ts#S\\<^sub>d = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>d \)" | "\M; Equality _ t t'#S\\<^sub>d = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d \)" | "\M; Inequality X F#S\\<^sub>d = (\\. ineq_model \ X F \ \M; S\\<^sub>d \)" definition constr_sem_d ("_ \ \_,_\") where "\ \ \S,\\ \ (\ supports \ \ \{}; S\\<^sub>d \)" abbreviation constr_sem_d' ("_ \ \_\" 90) where "\ \ \S\ \ \ \ \S,Var\" lemmas strand_sem_induct = strand_sem_c.induct[case_names Nil ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lemmata\ lemma strand_sem_d_if_c: "\ \\<^sub>c \S,\\ \ \ \ \S,\\" proof - assume *: "\ \\<^sub>c \S,\\" { fix M have "\M; S\\<^sub>c \ \ \M; S\\<^sub>d \" proof (induction S rule: strand_sem_induct) case (ConsSnd M ts S) hence "\t \ set ts. M \\<^sub>c t \ \" "\M; S\\<^sub>d \" by auto thus ?case using strand_sem_d.simps(2)[of M _ S] by auto qed (auto simp add: ineq_model_def) } thus ?thesis using * by (simp add: constr_sem_c_def constr_sem_d_def) qed lemma strand_sem_mono_ik: "\M \ M'; \M; S\\<^sub>c \\ \ \M'; S\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\M \ M'; \M; S\\<^sub>d \\ \ \M'; S\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") proof - show "\?A'; ?A''\ \ ?A" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case using ConsRcv.IH[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M" "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M'"] by auto next case (ConsSnd M ts S) hence "\t \ set ts. M \\<^sub>c t \ \" "\M'; S\\<^sub>c \" by auto hence "\t \ set ts. M' \\<^sub>c t \ \" using ideduct_synth_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>c \\ by simp qed auto show "\?B'; ?B''\ \ ?B" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case using ConsRcv.IH[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M" "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M'"] by auto next case (ConsSnd M ts S) hence "\t \ set ts. M \ t \ \" "\M'; S\\<^sub>d \" by auto hence "\t \ set ts. M' \ t \ \" using ideduct_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>d \\ by simp qed auto qed context begin private lemma strand_sem_split_left: "\M; S@S'\\<^sub>c \ \ \M; S\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M; S\\<^sub>d \" proof (induct S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons by (cases x) simp_all } { case 2 thus ?case using Cons by (cases x) simp_all } qed simp_all private lemma strand_sem_split_right: "\M; S@S'\\<^sub>c \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \" proof (induction S arbitrary: M rule: ik\<^sub>s\<^sub>t_induct) case (ConsRcv ts S) { case 1 thus ?case using ConsRcv.IH(1)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] by (simp add: Un_commute Un_left_commute image_Un) } { case 2 thus ?case using ConsRcv.IH(2)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] by (simp add: Un_commute Un_left_commute image_Un) } qed simp_all lemmas strand_sem_split[dest] = strand_sem_split_left(1) strand_sem_split_right(1) strand_sem_split_left(2) strand_sem_split_right(2) end lemma strand_sem_Send_split[dest]: "\\M; map Send T\\<^sub>c \; ts \ set T\ \ \M; [Send ts]\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\\M; map Send T\\<^sub>d \; ts \ set T\ \ \M; [Send ts]\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") "\\M; map Send T@S\\<^sub>c \; ts \ set T\ \ \M; Send ts#S\\<^sub>c \" (is "\?C'; ?C''\ \ ?C") "\\M; map Send T@S\\<^sub>d \; ts \ set T\ \ \M; Send ts#S\\<^sub>d \" (is "\?D'; ?D''\ \ ?D") "\\M; map Send1 T'\\<^sub>c \; t \ set T'\ \ \M; [Send1 t]\\<^sub>c \" (is "\?E'; ?E''\ \ ?E") "\\M; map Send1 T'\\<^sub>d \; t \ set T'\ \ \M; [Send1 t]\\<^sub>d \" (is "\?F'; ?F''\ \ ?F") "\\M; map Send1 T'@S\\<^sub>c \; t \ set T'\ \ \M; Send1 t#S\\<^sub>c \" (is "\?G'; ?G''\ \ ?G") "\\M; map Send1 T'@S\\<^sub>d \; t \ set T'\ \ \M; Send1 t#S\\<^sub>d \" (is "\?H'; ?H''\ \ ?H") proof - show A: "\?A'; ?A''\ \ ?A" by (induct "map Send T" arbitrary: T rule: strand_sem_c.induct) auto show B: "\?B'; ?B''\ \ ?B" by (induct "map Send T" arbitrary: T rule: strand_sem_d.induct) auto show "\?C'; ?C''\ \ ?C" "\?D'; ?D''\ \ ?D" using list.set_map list.simps(8) set_empty ik_snd_empty sup_bot.right_neutral by (metis (no_types, lifting) A strand_sem_split(1,2) strand_sem_c.simps(2), metis (no_types, lifting) B strand_sem_split(3,4) strand_sem_d.simps(2)) show "\?E'; ?E''\ \ ?E" by (induct "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) auto show "\?F'; ?F''\ \ ?F" by (induct "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) auto show "\?G'; ?G''\ \ ?G" proof (induction "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) case (2 M ts S') obtain t' T'' where ts: "ts = [t']" "T' = t'#T''" "S' = map Send1 T''" using "2.hyps"(2) by blast thus ?case using "2.prems" "2.hyps"(1) proof (cases "t = t'") case True have "ik\<^sub>s\<^sub>t (map Send1 T') \\<^sub>s\<^sub>e\<^sub>t \ = {}" by force hence "\M; [Send1 t]\\<^sub>c \" "\M; S\\<^sub>c \" using "2.prems"(1) unfolding ts(2) True by auto thus ?thesis by simp qed auto qed auto show "\?H'; ?H''\ \ ?H" proof (induction "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) case (2 M ts S') obtain t' T'' where ts: "ts = [t']" "T' = t'#T''" "S' = map Send1 T''" using "2.hyps"(2) by blast thus ?case using "2.prems" "2.hyps"(1) proof (cases "t = t'") case True have "ik\<^sub>s\<^sub>t (map Send1 T') \\<^sub>s\<^sub>e\<^sub>t \ = {}" by force hence "\M; [Send1 t]\\<^sub>d \" "\M; S\\<^sub>d \" using "2.prems"(1) unfolding ts(2) True by auto thus ?thesis by simp qed auto qed auto qed lemma strand_sem_Send_map: "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>c \) \ \M; map Send T\\<^sub>c \" "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>d \) \ \M; map Send T\\<^sub>d \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>c \) \ \M; map Send1 T'\\<^sub>c \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>d \) \ \M; map Send1 T'\\<^sub>d \" "\M; map Send1 T'\\<^sub>c \ \ \M; [Send T']\\<^sub>c \" "\M; map Send1 T'\\<^sub>d \ \ \M; [Send T']\\<^sub>d \" proof - show "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>c \) \ \M; map Send T\\<^sub>c \" "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>d \) \ \M; map Send T\\<^sub>d \" by (induct T) auto show "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>c \) \ \M; map Send1 T'\\<^sub>c \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>d \) \ \M; map Send1 T'\\<^sub>d \" by (induct T') auto show "\M; map Send1 T'\\<^sub>c \ \ \M; [Send T']\\<^sub>c \" "\M; map Send1 T'\\<^sub>d \ \ \M; [Send T']\\<^sub>d \" by (induct T') auto qed lemma strand_sem_Receive_map: "\M; map Receive T\\<^sub>c \" "\M; map Receive T\\<^sub>d \" "\M; map Receive1 T'\\<^sub>c \" "\M; map Receive1 T'\\<^sub>d \" "\M; [Receive T']\\<^sub>c \" "\M; [Receive T']\\<^sub>d \" proof - show "\M; map Receive T\\<^sub>c \" "\M; map Receive T\\<^sub>d \" by (induct T arbitrary: M) auto show "\M; map Receive1 T'\\<^sub>c \" "\M; map Receive1 T'\\<^sub>d \" by (induct T' arbitrary: M) auto show "\M; [Receive T']\\<^sub>c \" "\M; [Receive T']\\<^sub>d \" by (induct T' arbitrary: M) auto qed lemma strand_sem_append[intro]: "\\M; S\\<^sub>c \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \\ \ \M; S@S'\\<^sub>c \" "\\M; S\\<^sub>d \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \\ \ \M; S@S'\\<^sub>d \" proof (induction S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons proof (cases x) case (Receive ts) thus ?thesis using 1 Cons.IH(1)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] strand_sem_c.simps(3)[of M ts] image_Un[of "\t. t \ \" "set ts" "ik\<^sub>s\<^sub>t S"] by (metis (no_types, lifting) ik\<^sub>s\<^sub>t.simps(2) Un_assoc Un_commute append_Cons) qed auto } { case 2 thus ?case using Cons proof (cases x) case (Receive ts) thus ?thesis using 2 Cons.IH(2)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] strand_sem_d.simps(3)[of M ts] image_Un[of "\t. t \ \" "set ts" "ik\<^sub>s\<^sub>t S"] by (metis (no_types, lifting) ik\<^sub>s\<^sub>t.simps(2) Un_assoc Un_commute append_Cons) qed auto } qed simp_all lemma ineq_model_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model (\ \\<^sub>s \) X F" shows "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using * by (induct F) force+ have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "(fst f \ \) \ \ \\<^sub>s \ \ (snd f \ \) \ \ \\<^sub>s \" using f by auto moreover have "(fst f \ \, snd f \ \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_subst': fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" obtain f where f: "f \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) auto then obtain g where g: "g \ set F" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "fst g \ \ \\<^sub>s (\ \\<^sub>s \) \ snd g \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) g by (simp add: prod.case_eq_if) hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using g Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_ground_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" and "ground (subst_range \)" and "ineq_model \ X F" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \ )" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) force+ hence "fv (fst f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "fv (snd f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto hence "fv (fst f) - set X \ subst_domain \" "fv (snd f) - set X \ subst_domain \" using assms(1) by auto hence "fv (fst f \ \) \ subst_domain \" "fv (snd f \ \) \ subst_domain \" using \ by (simp_all add: range_vars_alt_def subst_fv_unfold_ground_img) hence "fv (fst f \ \ \\<^sub>s \) = {}" "fv (snd f \ \ \\<^sub>s \) = {}" using assms(2) by (simp_all add: subst_fv_dom_ground_if_ground_img) hence "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) subst_ground_ident by fastforce hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed context begin private lemma strand_sem_subst_c: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>c \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" "\t \ set ts. M \\<^sub>c t \ (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \\<^sub>c (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of _ \ \] by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \\<^sub>c t \ \" by fastforce thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>c \\ unfolding subst_apply_strand_def by simp next case (ConsRcv M ts S) have *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; S\\<^sub>c (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp have "bvars\<^sub>s\<^sub>t (Receive ts#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)#(S \\<^sub>s\<^sub>t \)\\<^sub>c \" using ConsRcv.IH[OF * **] by (metis (no_types) image_set strand_sem_c.simps(3) subst_comp_all) thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto qed simp_all private lemma strand_sem_subst_c': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>c \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" by auto hence "\M; S\\<^sub>c (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send ts]\\<^sub>c (\ \\<^sub>s \)" proof - have "\M; [send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t]\\<^sub>c \" using \\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>c \\ by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \\<^sub>c t \ \" by simp hence "\t \ set ts. M \\<^sub>c t \ \ \ \" by auto hence "\t \ set ts. M \\<^sub>c t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send ts]\\<^sub>c (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>c \" by (simp add: subst_all_insert) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>c \" by (metis subst_comp_all) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>c (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>d (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>d \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" "\t \ set ts. M \ t \ (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \ (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of _ \ \] by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>d \\ by simp next case (ConsRcv M ts S) have "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; S\\<^sub>d (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp hence *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>d (\ \\<^sub>s \)" by (metis subst_comp_all) have "bvars\<^sub>s\<^sub>t (Receive ts#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)#(S \\<^sub>s\<^sub>t \)\\<^sub>d \" using ConsRcv.IH[OF * **] by simp thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \ \ \M; S\\<^sub>d (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>d \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" by auto hence "\M; S\\<^sub>d (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send ts]\\<^sub>d (\ \\<^sub>s \)" proof - have "\M; [send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t]\\<^sub>d \" using \\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>d \\ by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp hence "\t \ set ts. M \ t \ \ \ \" by auto hence "\t \ set ts. M \ t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send ts]\\<^sub>d (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>d \" by (simp add: subst_all_insert) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>d \" by (metis subst_comp_all) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>d (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all lemmas strand_sem_subst = strand_sem_subst_c strand_sem_subst_c' strand_sem_subst_d strand_sem_subst_d' end lemma strand_sem_subst_subst_idem: assumes \: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\\M; S \\<^sub>s\<^sub>t \\\<^sub>c (\ \\<^sub>s \); subst_idem \\ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using strand_sem_subst(2)[OF assms, of M "\ \\<^sub>s \"] subst_compose_assoc[of \ \ \] unfolding subst_idem_def by argo lemma strand_sem_subst_comp: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" and "\M; S\\<^sub>c \" "subst_domain \ \ (vars\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>e\<^sub>t M) = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \)" proof - from assms(3) have "subst_domain \ \ vars\<^sub>s\<^sub>t S = {}" "subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t M = {}" by auto hence "S \\<^sub>s\<^sub>t \ = S" "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using strand_substI set_subst_ident[of M \] by (blast, blast) thus ?thesis using assms(2) by (auto simp add: strand_sem_subst(2)[OF assms(1)]) qed lemma strand_sem_c_imp_ineqs_neq: assumes "\M; S\\<^sub>c \" "Inequality X [(t,t')] \ set S" shows "t \ t' \ (\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ t' \ \ \ t \ \ \ \ \ t' \ \ \ \)" using assms proof (induction rule: strand_sem_induct) case (ConsIneq M Y F S) thus ?case proof (cases "Inequality X [(t,t')] \ set S") case False hence "X = Y" "F = [(t,t')]" using ConsIneq by auto hence *: "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using ConsIneq by (auto simp add: ineq_model_def) then obtain \ where \: "subst_domain \ = set X" "ground (subst_range \)" "t \ \ \ \ \ t' \ \ \ \" - using interpretation_subst_exists'[of "set X"] by moura + using interpretation_subst_exists'[of "set X"] by atomize_elim auto hence "t \ t'" by auto moreover have "\\ \. t \ \ \ \ \ t' \ \ \ \ \ t \ \ \ t' \ \" by auto ultimately show ?thesis using * by auto qed simp qed simp_all lemma strand_sem_c_imp_ineq_model: assumes "\M; S\\<^sub>c \" "Inequality X F \ set S" shows "ineq_model \ X F" using assms by (induct S rule: strand_sem_induct) force+ lemma strand_sem_wf_simple_fv_sat: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\{}; S\\<^sub>c \" shows "\v. v \ wfrestrictedvars\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsRcv t S) have "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) ConsRcv.prems(1) vars_snd_rcv_strand2 by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Receive t]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Receive t]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsRcv.IH ideduct_synth_mono by meson next case (ConsIneq X F S) hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Inequality X F]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Inequality X F]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsIneq.IH ideduct_synth_mono by meson next case (ConsSnd w S) hence *: "\{}; S\\<^sub>c \" "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ w" by auto have **: "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Send [Var w]]) \\<^sub>s\<^sub>e\<^sub>t \" by simp show ?case proof (cases "v = w") case True thus ?thesis using *(2) ideduct_synth_mono[OF _ **] by meson next case False hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsSnd.prems(1) by auto thus ?thesis using ConsSnd.IH[OF _ *(1)] ideduct_synth_mono[OF _ **] by metis qed qed simp lemma strand_sem_wf_ik_or_assignment_rhs_fun_subterm: assumes "wf\<^sub>s\<^sub>t {} A" "\{}; A\\<^sub>c \" "Var x \ ik\<^sub>s\<^sub>t A" "\ x = Fun f T" "t\<^sub>i \ set T" "\ik\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" obtains S where "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t A)" "Fun f T = Fun f S \ \" proof - have "x \ wfrestrictedvars\<^sub>s\<^sub>t A" by (metis (no_types) assms(3) set_rev_mp term.set_intros(3) vars_subset_if_in_strand_ik2) moreover have "Fun f T \ \ = Fun f T" by (metis subst_ground_ident interpretation_grounds_all assms(4,7)) ultimately obtain A\<^sub>p\<^sub>r\<^sub>e A\<^sub>s\<^sub>u\<^sub>f where *: "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ w)" "(\t. A = A\<^sub>p\<^sub>r\<^sub>e@Send t#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \\<^sub>s\<^sub>e\<^sub>t set t \\<^sub>s\<^sub>e\<^sub>t \) \ (\t t'. A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \ t \ \)" using wf_strand_first_Send_var_split[OF assms(1)] assms(4) subtermeqI' by metis moreover { fix ts assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Send ts#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" hence ***: "\t \ set ts. ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" "\ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using assms(2,6) by (auto intro: ideduct_synth_mono) then obtain t where t: "t \ set ts" "Fun f T \ t \ \" "ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using **(2) by blast obtain s where s: "s \ ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" "Fun f T \ s \ \" using t(3,2) ***(2) assms(5) by (induct rule: intruder_synth_induct) auto then obtain g S where gS: "Fun g S \ s" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF s(2)] *(1) by force hence ?thesis using that **(1) s(1) by force } moreover { fix t t' assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ t \ \" with assms(2) have "t \ \ = t' \ \" by auto hence "Fun f T \ t' \ \" using **(2) by auto from assms(1) **(1) have "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" using wf_eq_fv[of "{}" A\<^sub>p\<^sub>r\<^sub>e t t' A\<^sub>s\<^sub>u\<^sub>f] vars_snd_rcv_strand_subset2(4)[of A\<^sub>p\<^sub>r\<^sub>e] by blast then obtain g S where gS: "Fun g S \ t'" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF \Fun f T \ t' \ \\] *(1) by fastforce hence ?thesis using that **(1) by auto } ultimately show ?thesis by auto qed lemma ineq_model_not_unif_is_sat_ineq: assumes "\\. Unifier \ t t'" shows "ineq_model \ X [(t, t')]" using assms list.set_intros(1)[of "(t,t')" "[]"] unfolding ineq_model_def by blast lemma strand_sem_not_unif_is_sat_ineq: assumes "\\. Unifier \ t t'" shows "\M; [Inequality X [(t,t')]]\\<^sub>c \" "\M; [Inequality X [(t,t')]]\\<^sub>d \" using ineq_model_not_unif_is_sat_ineq[OF assms] strand_sem_c.simps(1,5)[of M] strand_sem_d.simps(1,5)[of M] by presburger+ lemma ineq_model_singleI[intro]: assumes "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" shows "ineq_model \ X [(t,t')]" using assms unfolding ineq_model_def by auto lemma ineq_model_singleE: assumes "ineq_model \ X [(t,t')]" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using assms unfolding ineq_model_def by auto lemma ineq_model_single_iff: fixes F::"(('a,'b) term \ ('a,'b) term) list" shows "ineq_model \ X F \ ineq_model \ X [(Fun f (Fun c []#map fst F),Fun f (Fun c []#map snd F))]" (is "?A \ ?B") proof - let ?P = "\\ f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)" let ?Q = "\\ t t'. t \ (\ \\<^sub>s \) \ t' \ (\ \\<^sub>s \)" let ?T = "\g. Fun c []#map g F" let ?S = "\\ g. map (\x. x \ (\ \\<^sub>s \)) (Fun c []#map g F)" let ?t = "Fun f (?T fst)" let ?t' = "Fun f (?T snd)" have len: "\g h. length (?T g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \ \. length (?S \ g) = length (?S \ h)" by simp_all { fix \::"('a,'b) subst" assume \: "subst_domain \ = set X" "ground (subst_range \)" have "list_ex (?P \) F \ ?Q \ ?t ?t'" proof assume "list_ex (?P \) F" then obtain a where a: "a \ set F" "?P \ a" by (metis (mono_tags, lifting) Bex_set) thus "?Q \ ?t ?t'" by auto qed (fastforce simp add: Bex_set) } thus ?thesis unfolding ineq_model_def case_prod_unfold by auto qed subsection \Constraint Semantics (Alternative, Equivalent Version)\ text \These are the constraint semantics used in the CSF 2017 paper\ fun strand_sem_c'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c''") where "\M; []\\<^sub>c' = (\\. True)" | "\M; Send ts#S\\<^sub>c' = (\\. (\t \ set ts. M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \) \ \M; S\\<^sub>c' \)" | "\M; Receive ts#S\\<^sub>c' = \set ts \ M; S\\<^sub>c'" | "\M; Equality _ t t'#S\\<^sub>c' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c' \)" | "\M; Inequality X F#S\\<^sub>c' = (\\. ineq_model \ X F \ \M; S\\<^sub>c' \)" fun strand_sem_d'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d''") where "\M; []\\<^sub>d' = (\\. True)" | "\M; Send ts#S\\<^sub>d' = (\\. (\t \ set ts. M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \) \ \M; S\\<^sub>d' \)" | "\M; Receive ts#S\\<^sub>d' = \set ts \ M; S\\<^sub>d'" | "\M; Equality _ t t'#S\\<^sub>d' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d' \)" | "\M; Inequality X F#S\\<^sub>d' = (\\. ineq_model \ X F \ \M; S\\<^sub>d' \)" lemma strand_sem_eq_defs: "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" proof - have 1: "\M; \\\<^sub>c' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" proof (induction \ arbitrary: M rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 2: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \ \ \M; \\\<^sub>c' \" proof (induction \ arbitrary: M rule: strand_sem_c'.induct) case (3 M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 3: "\M; \\\<^sub>d' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" proof (induction \ arbitrary: M rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 4: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \ \ \M; \\\<^sub>d' \" proof (induction \ arbitrary: M rule: strand_sem_d'.induct) case (3 M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all show "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" by (metis 1 2, metis 3 4) qed lemma strand_sem_split'[dest]: "\M; S@S'\\<^sub>c' \ \ \M; S\\<^sub>c' \" "\M; S@S'\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \" "\M; S@S'\\<^sub>d' \ \ \M; S\\<^sub>d' \" "\M; S@S'\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] strand_sem_split(2,4) by (auto simp add: image_Un) lemma strand_sem_append'[intro]: "\M; S\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \ \ \M; S@S'\\<^sub>c' \" "\M; S\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \ \ \M; S@S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] by (auto simp add: image_Un) end subsection \Dual Strands\ fun dual\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) strand" where "dual\<^sub>s\<^sub>t [] = []" | "dual\<^sub>s\<^sub>t (Receive t#S) = Send t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (Send t#S) = Receive t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (x#S) = x#(dual\<^sub>s\<^sub>t S)" lemma dual\<^sub>s\<^sub>t_append: "dual\<^sub>s\<^sub>t (A@B) = (dual\<^sub>s\<^sub>t A)@(dual\<^sub>s\<^sub>t B)" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_self_inverse: "dual\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_trms_eq: "trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = trms\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_fv: "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = fv\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_bvars: "bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = bvars\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) fastforce+ end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy @@ -1,2208 +1,2208 @@ (* Title: Typed_Model.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Typed Model\ theory Typed_Model imports Lazy_Intruder begin text \Term types\ type_synonym ('f,'v) term_type = "('f,'v) term" text \Constructors for term types\ abbreviation (input) TAtom::"'v \ ('f,'v) term_type" where "TAtom a \ Var a" abbreviation (input) TComp::"['f, ('f,'v) term_type list] \ ('f,'v) term_type" where "TComp f ts \ Fun f ts" text \ The typed model extends the intruder model with a typing function \\\ that assigns types to terms. \ locale typed_model = intruder_model arity public Ana for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" + fixes \::"('fun,'var) term \ ('fun,'atom::finite) term_type" assumes const_type: "\c. arity c = 0 \ \a. \ts. \ (Fun c ts) = TAtom a" and fun_type: "\f ts. arity f > 0 \ \ (Fun f ts) = TComp f (map \ ts)" and \_wf: "\x f ts. TComp f ts \ \ (Var x) \ arity f > 0" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" begin subsection \Definitions\ text \The set of atomic types\ abbreviation "\\<^sub>a \ UNIV::('atom set)" text \Well-typed substitutions\ definition wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. \ (Var v) = \ (\ v))" text \The set of sub-message patterns (SMP)\ inductive_set SMP::"('fun,'var) terms \ ('fun,'var) terms" for M where MP[intro]: "t \ M \ t \ SMP M" | Subterm[intro]: "\t \ SMP M; t' \ t\ \ t' \ SMP M" | Substitution[intro]: "\t \ SMP M; wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ \ (t \ \) \ SMP M" | Ana[intro]: "\t \ SMP M; Ana t = (K,T); k \ set K\ \ k \ SMP M" text \ Type-flaw resistance for sets: Unifiable sub-message patterns must have the same type (unless they are variables) \ definition tfr\<^sub>s\<^sub>e\<^sub>t where "tfr\<^sub>s\<^sub>e\<^sub>t M \ (\s \ SMP M - (Var`\). \t \ SMP M - (Var`\). (\\. Unifier \ s t) \ \ s = \ t)" text \ Type-flaw resistance for strand steps: - The terms in a satisfiable equality step must have the same types - Inequality steps must satisfy the conditions of the "inequality lemma"\ fun tfr\<^sub>s\<^sub>t\<^sub>p where "tfr\<^sub>s\<^sub>t\<^sub>p (Equality a t t') = ((\\. Unifier \ t t') \ \ t = \ t')" | "tfr\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X)))" | "tfr\<^sub>s\<^sub>t\<^sub>p _ = True" text \ Type-flaw resistance for strands: - The set of terms in strands must be type-flaw resistant - The steps of strands must be type-flaw resistant \ definition tfr\<^sub>s\<^sub>t where "tfr\<^sub>s\<^sub>t S \ tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" subsection \Small Lemmata\ lemma tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ ((\a t t'. Equality a t t' \ set S \ (\\. Unifier \ t t') \ \ t = \ t') \ (\X F. Inequality X F \ set S \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))))" (is "?P S \ ?Q S") proof show "?P S \ ?Q S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp show "?Q S \ ?P S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp qed lemma \_wf'': "TComp f T \ \ t \ arity f > 0" proof (induction t) case (Var x) thus ?case using \_wf(1)[of f T x] by blast next case (Fun g S) thus ?case using fun_type[of g S] const_type[of g] by (cases "arity g") auto qed lemma \_wf': "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" proof (induction t) case (Fun f T) hence *: "arity f = length T" "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto { assume "arity f = 0" hence ?case using const_type[of f] by auto } moreover { assume "arity f > 0" hence ?case using fun_type[of f] * by force } ultimately show ?case by auto qed (metis \_wf(2)) lemma fun_type_inv: assumes "\ t = TComp f T" shows "arity f > 0" using \_wf''(1)[of f T t] assms by simp_all lemma fun_type_inv_wf: assumes "\ t = TComp f T" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "arity f = length T" using \_wf'[OF assms(2)] assms(1) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma const_type_inv: "\ (Fun c X) = TAtom a \ arity c = 0" by (rule ccontr, simp add: fun_type) lemma const_type_inv_wf: assumes "\ (Fun c X) = TAtom a" and "wf\<^sub>t\<^sub>r\<^sub>m (Fun c X)" shows "X = []" by (metis assms const_type_inv length_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) lemma const_type': "\c \ \. \a \ \\<^sub>a. \X. \ (Fun c X) = TAtom a" using const_type by simp lemma fun_type': "\f \ \\<^sub>f. \X. \ (Fun f X) = TComp f (map \ X)" using fun_type by simp lemma fun_type_id_eq: "\ (Fun f X) = TComp g Y \ f = g" by (metis const_type fun_type neq0_conv "term.inject"(2) "term.simps"(4)) lemma fun_type_length_eq: "\ (Fun f X) = TComp g Y \ length X = length Y" by (metis fun_type fun_type_id_eq fun_type_inv(1) length_map term.inject(2)) lemma pgwt_type_map: assumes "public_ground_wf_term t" shows "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" proof - let ?A = "\ t = TAtom a \ (\f. t = Fun f [])" let ?B = "\ t = TComp g Y \ (\X. t = Fun g X \ map \ X = Y)" have "?A \ ?B" proof (cases "\ t") case (Var a) obtain f X where "t = Fun f X" "arity f = length X" using pgwt_fun[OF assms(1)] pgwt_arity[OF assms(1)] by fastforce+ thus ?thesis using const_type_inv \\ t = TAtom a\ by auto next case (Fun g Y) obtain f X where *: "t = Fun f X" using pgwt_fun[OF assms(1)] by force hence "f = g" "map \ X = Y" using fun_type_id_eq \\ t = TComp g Y\ fun_type[OF fun_type_inv(1)[OF \\ t = TComp g Y\]] by fastforce+ thus ?thesis using *(1) \\ t = TComp g Y\ by auto qed thus "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" by auto qed lemma wt_subst_Var[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm: "(\v. v \ fv t \ \ (Var v) = \ (\ v)) \ \ t = \ (t \ \)" proof (induction t) case (Fun f X) hence *: "\x. x \ set X \ \ x = \ (x \ \)" by auto show ?case proof (cases "f \ \\<^sub>f") case True hence "\X. \ (Fun f X) = TComp f (map \ X)" using fun_type' by auto thus ?thesis using * by auto next case False hence "\a \ \\<^sub>a. \X. \ (Fun f X) = TAtom a" using const_type' by auto thus ?thesis by auto qed qed auto lemma wt_subst_trm': "\wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; \ s = \ t\ \ \ (s \ \) = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm'': "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ t = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_compose: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have "\v. \ (\ v) = \ (\ v \ \)" using wt_subst_trm \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis moreover have "\v. \ (Var v) = \ (\ v)" using \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis ultimately have "\v. \ (Var v) = \ (\ v \ \)" by metis thus ?thesis unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_compose_def by metis qed lemma wt_subst_TAtom_Var_cases: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and x: "\ (Var x) = TAtom a" shows "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" proof (cases "(\y. \ x = Var y)") case False then obtain c T where c: "\ x = Fun c T" by (cases "\ x") simp_all hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun c T)" using \(2) by fastforce hence "T = []" using const_type_inv_wf[of c T a] x c wt_subst_trm''[OF \(1), of "Var x"] by fastforce thus ?thesis using c by blast qed simp lemma wt_subst_TAtom_fv: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and "\x \ fv t - X. \a. \ (Var x) = TAtom a" shows "\x \ fv (t \ \) - fv\<^sub>s\<^sub>e\<^sub>t (\ ` X). \a. \ (Var x) = TAtom a" using assms(3) proof (induction t) case (Var x) thus ?case proof (cases "x \ X") case False - with Var obtain a where "\ (Var x) = TAtom a" by moura + with Var obtain a where "\ (Var x) = TAtom a" by atomize_elim auto hence *: "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto show ?thesis proof (cases "\ x") case (Var y) thus ?thesis using * by auto next case (Fun f T) hence "T = []" using * const_type_inv[of f T a] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?thesis using Fun by auto qed qed auto qed fastforce lemma wt_subst_TAtom_subterms_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv t. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv t)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) proof (induction t) case (Var x) - obtain a where a: "\ (Var x) = TAtom a" using Var.prems(1) by moura + obtain a where a: "\ (Var x) = TAtom a" using Var.prems(1) by atomize_elim auto hence "\ (\ x) = TAtom a" using wt_subst_trm''[OF assms(1), of "Var x"] by simp hence "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" using const_type_inv_wf Var.prems(2) by (cases "\ x") auto thus ?case by auto next case (Fun f T) have "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" when "t \ set T" for t using that Fun.prems(1,2) Fun.IH[OF that] by auto thus ?case by auto qed lemma wt_subst_TAtom_subterms_set_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv\<^sub>s\<^sub>e\<^sub>t M)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof show "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" then obtain s where s: "s \ M" "t \ subterms (s \ \)" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed show "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ M" "t \ subterms s \\<^sub>s\<^sub>e\<^sub>t \" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed qed lemma wt_subst_subst_upd: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ (Var x) = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\(x := t))" using assms unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis fun_upd_other fun_upd_same) lemma wt_subst_const_fv_type_eq: assumes "\x \ fv t. \a. \ (Var x) = TAtom a" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\x \ fv (t \ \). \y \ fv t. \ (Var x) = \ (Var y)" using assms(1) proof (induction t) case (Var x) - then obtain a where a: "\ (Var x) = TAtom a" by moura + then obtain a where a: "\ (Var x) = TAtom a" by atomize_elim auto show ?case proof (cases "\ x") case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "\ (Fun f T) = TAtom a" using a wt_subst_trm''[OF \(1), of "Var x"] \(2) by fastforce+ thus ?thesis using const_type_inv_wf Fun by fastforce qed (use a wt_subst_trm''[OF \(1), of "Var x"] in simp) qed fastforce lemma TComp_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TComp f T" shows "(\v. t = Var v) \ (\T'. t = Fun f T' \ T = map \ T' \ T' \ [])" proof (cases "\v. t = Var v") case False then obtain T' where T': "t = Fun f T'" "T = map \ T'" using assms fun_type[OF fun_type_inv(1)[OF assms(2)]] fun_type_id_eq by (cases t) force+ thus ?thesis using assms fun_type_inv(1) fun_type_inv_wf by fastforce qed metis lemma TAtom_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "(\v. t = Var v) \ (\f. t = Fun f [])" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) auto lemma subtermeq_imp_subtermtypeeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "\ s \ \ t" using assms(2,1) proof (induction t) case (Fun f T) thus ?case proof (cases "s = Fun f T") case False then obtain x where x: "x \ set T" "s \ x" using Fun.prems(1) by auto hence "wf\<^sub>t\<^sub>r\<^sub>m x" using wf_trm_subtermeq[OF Fun.prems(2)] Fun_param_is_subterm[of _ T f] by auto hence "\ s \ \ x" using Fun.IH[OF x] by simp moreover have "arity f > 0" using x fun_type_inv_wf Fun.prems by (metis length_pos_if_in_set term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using x Fun.prems fun_type[of f T] by auto qed simp qed simp lemma subterm_funs_term_in_type: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "Fun f T \ t" "\ (Fun f T) = TComp f (map \ T)" shows "f \ funs_term (\ t)" using assms(2,1,3) proof (induction t) case (Fun f' T') hence [simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by (metis wf_trm_subtermeq) { fix a assume \: "\ (Fun f' T') = TAtom a" hence "Fun f T = Fun f' T'" using Fun TAtom_term_cases subtermeq_Var_const by metis hence False using Fun.prems(3) \ by simp } moreover { fix g S assume \: "\ (Fun f' T') = TComp g S" hence "g = f'" "S = map \ T'" using Fun.prems(2) fun_type_id_eq[OF \] fun_type[OF fun_type_inv(1)[OF \]] by auto hence \': "\ (Fun f' T') = TComp f' (map \ T')" using \ by auto hence "g \ funs_term (\ (Fun f' T'))" using \ by auto moreover { assume "Fun f T \ Fun f' T'" then obtain x where "x \ set T'" "Fun f T \ x" using Fun.prems(1) by auto hence "f \ funs_term (\ x)" using Fun.IH[OF _ _ _ Fun.prems(3), of x] wf_trm_subtermeq[OF \wf\<^sub>t\<^sub>r\<^sub>m (Fun f' T')\, of x] by force moreover have "\ x \ set (map \ T')" using \' \x \ set T'\ by auto ultimately have "f \ funs_term (\ (Fun f' T'))" using \' by auto } ultimately have ?case by (cases "Fun f T = Fun f' T'") (auto simp add: \g = f'\) } ultimately show ?case by (cases "\ (Fun f' T')") auto qed simp lemma wt_subst_fv_termtype_subterm: assumes "x \ fv (\ y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m (\ y)" shows "\ (Var x) \ \ (Var y)" using subtermeq_imp_subtermtypeeq[OF assms(3) var_is_subterm[OF assms(1)]] wt_subst_trm''[OF assms(2), of "Var y"] by auto lemma wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\y \ Y. \ (Var x) \ \ (Var y)" using wt_subst_fv_termtype_subterm[OF _ assms(2), of x] assms(1,3) by fastforce lemma funs_term_type_iff: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" and f: "arity f > 0" shows "f \ funs_term (\ t) \ (f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x))))" (is "?P t \ ?Q t") using t proof (induction t) case (Fun g T) hence IH: "?P s \ ?Q s" when "s \ set T" for s using that wf_trm_subterm[OF _ Fun_param_is_subterm] by blast have 0: "arity g = length T" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto show ?case proof (cases "f = g") case True thus ?thesis using fun_type[OF f] by simp next case False have "?P (Fun g T) \ (\s \ set T. ?P s)" proof assume *: "?P (Fun g T)" hence "\ (Fun g T) = TComp g (map \ T)" using const_type[of g] fun_type[of g] by force thus "\s \ set T. ?P s" using False * by force next assume *: "\s \ set T. ?P s" hence "\ (Fun g T) = TComp g (map \ T)" using 0 const_type[of g] fun_type[of g] by force thus "?P (Fun g T)" using False * by force qed thus ?thesis using False f IH by auto qed qed simp lemma funs_term_type_iff': assumes M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and f: "arity f > 0" shows "f \ \(funs_term ` \ ` M) \ (f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x))))" (is "?A \ ?B") proof assume ?A - then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term (\ t)" using M by moura + then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term (\ t)" using M by atomize_elim auto thus ?B using funs_term_type_iff[OF _ f, of t] by auto next assume ?B then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" using M by auto thus ?A using funs_term_type_iff[OF _ f, of t] by blast qed lemma Ana_subterm_type: assumes "Ana t = (K,M)" and "wf\<^sub>t\<^sub>r\<^sub>m t" and "m \ set M" shows "\ m \ \ t" proof - have "m \ t" using Ana_subterm[OF assms(1)] assms(3) by auto thus ?thesis using subtermeq_imp_subtermtypeeq[OF assms(2)] by simp qed lemma wf_trm_TAtom_subterms: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "subterms t = {t}" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) force+ lemma wf_trm_TComp_subterm: assumes "wf\<^sub>t\<^sub>r\<^sub>m s" "t \ s" obtains f T where "\ s = TComp f T" proof (cases s) case (Var x) thus ?thesis using \t \ s\ by simp next case (Fun g S) hence "length S > 0" using assms Fun_subterm_inside_params[of t g S] by auto hence "arity g > 0" by (metis \wf\<^sub>t\<^sub>r\<^sub>m s\ \s = Fun g S\ term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) thus ?thesis using fun_type \s = Fun g S\ that by auto qed lemma SMP_empty[simp]: "SMP {} = {}" proof (rule ccontr) assume "SMP {} \ {}" then obtain t where "t \ SMP {}" by auto thus False by (induct t rule: SMP.induct) auto qed lemma SMP_I: assumes "s \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t \ s \ \" "\v. wf\<^sub>t\<^sub>r\<^sub>m (\ v)" shows "t \ SMP M" using SMP.Substitution[OF SMP.MP[OF assms(1)] assms(2)] SMP.Subterm[of "s \ \" M t] assms(3,4) by (cases "t = s \ \") simp_all lemma SMP_wf_trm: assumes "t \ SMP M" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m t" using assms(1) by (induct t rule: SMP.induct) (use assms(2) in blast, use wf_trm_subtermeq in blast, use wf_trm_subst in blast, use Ana_keys_wf' in blast) lemma SMP_ikI[intro]: "t \ ik\<^sub>s\<^sub>t S \ t \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma MP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by force lemma SMP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma SMP_subset_I: assumes M: "\t \ M. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" shows "SMP M \ SMP N" proof fix t show "t \ SMP M \ t \ SMP N" proof (induction t rule: SMP.induct) case (MP t) then obtain s \ where s: "s \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = s \ \" - using M by moura + using M by atomize_elim auto show ?case using SMP_I[OF s(1,2), of "s \ \"] s(3,4) wf_trm_subst_range_iff by fast qed (auto intro!: SMP.Substitution[of _ N]) qed lemma SMP_union: "SMP (A \ B) = SMP A \ SMP B" proof show "SMP (A \ B) \ SMP A \ SMP B" proof fix t assume "t \ SMP (A \ B)" thus "t \ SMP A \ SMP B" by (induct rule: SMP.induct) blast+ qed { fix t assume "t \ SMP A" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } moreover { fix t assume "t \ SMP B" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } ultimately show "SMP A \ SMP B \ SMP (A \ B)" by blast qed lemma SMP_append[simp]: "SMP (trms\<^sub>s\<^sub>t (S@S')) = SMP (trms\<^sub>s\<^sub>t S) \ SMP (trms\<^sub>s\<^sub>t S')" (is "?A = ?B") using SMP_union by simp lemma SMP_mono: "A \ B \ SMP A \ SMP B" proof - assume "A \ B" - then obtain C where "B = A \ C" by moura + then obtain C where "B = A \ C" by atomize_elim auto thus "SMP A \ SMP B" by (simp add: SMP_union) qed lemma SMP_Union: "SMP (\m \ M. f m) = (\m \ M. SMP (f m))" proof show "SMP (\m\M. f m) \ (\m\M. SMP (f m))" proof fix t assume "t \ SMP (\m\M. f m)" thus "t \ (\m\M. SMP (f m))" by (induct t rule: SMP.induct) force+ qed show "(\m\M. SMP (f m)) \ SMP (\m\M. f m)" proof fix t assume "t \ (\m\M. SMP (f m))" - then obtain m where "m \ M" "t \ SMP (f m)" by moura + then obtain m where "m \ M" "t \ SMP (f m)" by atomize_elim auto thus "t \ SMP (\m\M. f m)" using SMP_mono[of "f m" "\m\M. f m"] by auto qed qed lemma SMP_singleton_ex: "t \ SMP M \ (\m \ M. t \ SMP {m})" "m \ M \ t \ SMP {m} \ t \ SMP M" using SMP_Union[of "\t. {t}" M] by auto lemma SMP_Cons: "SMP (trms\<^sub>s\<^sub>t (x#S)) = SMP (trms\<^sub>s\<^sub>t [x]) \ SMP (trms\<^sub>s\<^sub>t S)" using SMP_append[of "[x]" S] by auto lemma SMP_Nil[simp]: "SMP (trms\<^sub>s\<^sub>t []) = {}" proof - { fix t assume "t \ SMP (trms\<^sub>s\<^sub>t [])" hence False by induct auto } thus ?thesis by blast qed lemma SMP_subset_union_eq: assumes "M \ SMP N" shows "SMP N = SMP (M \ N)" proof - { fix t assume "t \ SMP (M \ N)" hence "t \ SMP N" using assms by (induction rule: SMP.induct) blast+ } thus ?thesis using SMP_union by auto qed lemma SMP_subterms_subset: "subterms\<^sub>s\<^sub>e\<^sub>t M \ SMP M" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" then obtain m where "m \ M" "t \ m" by auto thus "t \ SMP M" using SMP_I[of _ _ Var] by auto qed lemma SMP_SMP_subset: "N \ SMP M \ SMP N \ SMP M" by (metis SMP_mono SMP_subset_union_eq Un_commute Un_upper2) lemma wt_subst_rm_vars: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_vars X \)" using rm_vars_dom unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma wt_subst_SMP_subset: assumes "trms\<^sub>s\<^sub>t S \ SMP S'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ SMP S'" proof fix t assume *: "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" show "t \ SMP S'" using trm_strand_subst_cong(2)[OF *] proof assume "\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S" thus "t \ SMP S'" using assms SMP.Substitution by auto next assume "\X F. Inequality X F \ set S \ (\t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \)" then obtain X F t' where **: "Inequality X F \ set S" "t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "t = t' \ rm_vars (set X) \" by force - then obtain s where s: "s \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F)" "t = s \ rm_vars (set X) \" by moura + then obtain s where s: "s \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F)" "t = s \ rm_vars (set X) \" by atomize_elim auto hence "s \ SMP (trms\<^sub>s\<^sub>t S)" using **(1) by force hence "t \ SMP (trms\<^sub>s\<^sub>t S)" using SMP.Substitution[OF _ wt_subst_rm_vars[OF assms(2)] wf_trms_subst_rm_vars'[OF assms(3)]] unfolding s(2) by blast thus "t \ SMP S'" by (metis SMP_union SMP_subset_union_eq UnCI assms(1)) qed qed lemma MP_subset_SMP: "\(trms\<^sub>s\<^sub>t\<^sub>p ` set S) \ SMP (trms\<^sub>s\<^sub>t S)" "trms\<^sub>s\<^sub>t S \ SMP (trms\<^sub>s\<^sub>t S)" "M \ SMP M" by auto lemma SMP_fun_map_snd_subset: "SMP (trms\<^sub>s\<^sub>t (map Send1 X)) \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof fix t assume "t \ SMP (trms\<^sub>s\<^sub>t (map Send1 X))" thus "t \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof (induction t rule: SMP.induct) case (MP t) hence "t \ set X" by auto hence "t \ Fun f X" by (metis subtermI') thus ?case using SMP.Subterm[of "Fun f X" "trms\<^sub>s\<^sub>t [Send1 (Fun f X)]" t] using SMP.MP by auto qed blast+ qed lemma SMP_wt_subst_subset: assumes "t \ SMP (M \\<^sub>s\<^sub>e\<^sub>t \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "t \ SMP M" using assms wf_trm_subst_range_iff[of \] by (induct t rule: SMP.induct) blast+ lemma SMP_wt_instances_subset: assumes "\t \ M. \s \ N. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "t \ SMP M" shows "t \ SMP N" proof - obtain m where m: "m \ M" "t \ SMP {m}" using SMP_singleton_ex(1)[OF assms(2)] by blast then obtain n \ where n: "n \ N" "m = n \ \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms(1) by fast have "t \ SMP (N \\<^sub>s\<^sub>e\<^sub>t \)" using n(1,2) SMP_singleton_ex(2)[of m "N \\<^sub>s\<^sub>e\<^sub>t \", OF _ m(2)] by fast thus ?thesis using SMP_wt_subst_subset[OF _ n(3,4)] by blast qed lemma SMP_consts: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" shows "SMP M = M" proof show "SMP M \ M" proof fix t show "t \ SMP M \ t \ M" apply (induction t rule: SMP.induct) by (use assms in auto) qed qed auto lemma SMP_subterms_eq: "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) = SMP M" proof show "SMP M \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M)" using SMP_mono[of M "subterms\<^sub>s\<^sub>e\<^sub>t M"] by blast show "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ SMP M" proof fix t show "t \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ t \ SMP M" by (induction t rule: SMP.induct) blast+ qed qed lemma SMP_funs_term: assumes t: "t \ SMP M" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" and f: "arity f > 0" and M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and Ana_f: "\s K T. Ana s = (K,T) \ f \ \(funs_term ` set K) \ f \ funs_term s" shows "f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x)))" using t proof (induction t rule: SMP.induct) case (Subterm t t') thus ?case by (metis UN_I vars_iff_subtermeq funs_term_subterms_eq(1) term.order_trans) next case (Substitution t \) show ?case using M SMP_wf_trm[OF Substitution.hyps(1)] wf_trm_subst[of \ t, OF Substitution.hyps(3)] funs_term_type_iff[OF _ f] wt_subst_trm''[OF Substitution.hyps(2), of t] Substitution.prems Substitution.IH by metis next case (Ana t K T t') thus ?case using Ana_f[OF Ana.hyps(2)] Ana_keys_fv[OF Ana.hyps(2)] by fastforce qed auto lemma id_type_eq: assumes "\ (Fun f X) = \ (Fun g Y)" shows "f \ \ \ g \ \" "f \ \\<^sub>f \ g \ \\<^sub>f" using assms const_type' fun_type' id_union_univ(1) by (metis UNIV_I UnE "term.distinct"(1))+ lemma fun_type_arg_cong: assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (x#X)) = \ (Fun g (y#Y))" shows "\ x = \ y" "\ (Fun f X) = \ (Fun g Y)" using assms fun_type' by auto lemma fun_type_arg_cong': assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" shows "\ x = \ y" using assms proof (induction X arbitrary: Y) case Nil thus ?case using fun_type_arg_cong(1)[of f g x X' y Y'] by auto next case (Cons x' X Y'') then obtain y' Y where "Y'' = y'#Y" by (metis length_Suc_conv) hence "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" using Cons.prems(3,4) fun_type_arg_cong(2)[OF Cons.prems(1,2), of x' "X@x#X'"] by auto thus ?thesis using Cons.IH[OF Cons.prems(1,2)] by auto qed lemma fun_type_param_idx: "\ (Fun f T) = Fun g S \ i < length T \ \ (T ! i) = S ! i" by (metis fun_type fun_type_id_eq fun_type_inv(1) nth_map term.inject(2)) lemma fun_type_param_ex: assumes "\ (Fun f T) = Fun g (map \ S)" "t \ set S" shows "\s \ set T. \ s = \ t" using fun_type_length_eq[OF assms(1)] length_map[of \ S] assms(2) fun_type_param_idx[OF assms(1)] nth_map in_set_conv_nth by metis lemma tfr_stp_all_split: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p [x]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@x#S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by fastforce+ lemma tfr_stp_all_append: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" using assms by fastforce lemma tfr_stp_all_wt_subst_apply: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "bvars\<^sub>s\<^sub>t S \ range_vars \ = {}" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using assms(1,4) proof (induction S) case (Cons x S) hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using tfr_stp_all_split(2)[of x S] unfolding range_vars_alt_def by fastforce thus ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons.prems by auto hence "(\\. Unifier \ (t \ \) (t' \ \)) \ \ (t \ \) = \ (t' \ \)" by (metis Unifier_comp' wt_subst_trm'[OF assms(2)]) moreover have "(x#S) \\<^sub>s\<^sub>t \ = Equality a (t \ \) (t' \ \)#(S \\<^sub>s\<^sub>t \)" using \x = Equality a t t'\ by auto ultimately show ?thesis using IH by auto next case (Inequality X F) let ?\ = "rm_vars (set X) \" let ?G = "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\" let ?P = "\F X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a" let ?Q = "\F X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X)" have 0: "set X \ range_vars ?\ = {}" using Cons.prems(2) Inequality rm_vars_img_subset[of "set X"] by (auto simp add: subst_domain_def range_vars_alt_def) have 1: "?P F X \ ?Q F X" using Inequality Cons.prems by simp have 2: "fv\<^sub>s\<^sub>e\<^sub>t (?\ ` set X) = set X" by auto have "?P ?G X" when "?P F X" using that proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) have "\x \ (fv (t \ ?\) \ fv (t' \ ?\)) - set X. \a. \ (Var x) = Var a" proof - have *: "\x \ fv t - set X. \a. \ (Var x) = Var a" "\x \ fv t' - set X. \a. \ (Var x) = Var a" using g Cons.prems by simp_all have **: "\x. wf\<^sub>t\<^sub>r\<^sub>m (?\ x)" using \(2) wf_trm_subst_range_iff[of \] wf_trm_subst_rm_vars'[of \ _ "set X"] by simp show ?thesis using wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(1)] wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(2)] wt_subst_trm'[OF wt_subst_rm_vars[OF \(1), of "set X"]] 2 by blast qed moreover have "\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) - set X. \a. \ (Var x) = Var a" using Cons by auto ultimately show ?case using g by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) hence "?P ?G X \ ?Q ?G X" using 1 ineq_subterm_inj_cond_subst[OF 0, of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of F ?\] by presburger moreover have "(x#S) \\<^sub>s\<^sub>t \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)#(S \\<^sub>s\<^sub>t \)" using \x = Inequality X F\ by auto ultimately show ?thesis using IH by simp qed auto qed simp lemma tfr_stp_all_same_type: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S') \ Unifier \ t t' \ \ t = \ t'" by force+ lemma tfr_subset: "\A B. tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - show 1: "tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" for A B using SMP_union[of A B] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp fix A B assume B: "tfr\<^sub>s\<^sub>e\<^sub>t B" show "A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "A \ B" - then obtain C where "B = A \ C" by moura + then obtain C where "B = A \ C" by atomize_elim auto thus ?thesis using B 1 by blast qed show "SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "SMP A \ SMP B" - then obtain C where "SMP B = SMP A \ C" by moura + then obtain C where "SMP B = SMP A \ C" by atomize_elim auto thus ?thesis using B unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed qed lemma tfr_empty[simp]: "tfr\<^sub>s\<^sub>e\<^sub>t {}" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp lemma tfr_consts_mono: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" and "tfr\<^sub>s\<^sub>e\<^sub>t N" shows "tfr\<^sub>s\<^sub>e\<^sub>t (N \ M)" proof - { fix s t assume *: "s \ SMP (N \ M) - range Var" "t \ SMP (N \ M) - range Var" "\\. Unifier \ s t" hence **: "is_Fun s" "is_Fun t" "s \ SMP N \ s \ M" "t \ SMP N \ t \ M" using assms(3) SMP_consts[OF assms(1,2)] SMP_union[of N M] by auto moreover have "\ s = \ t" when "s \ SMP N" "t \ SMP N" using that assms(3) *(3) **(1,2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "\ s = \ t" when st: "s \ M" "t \ M" proof - - obtain c d where "s = Fun c []" "t = Fun d []" using st assms(1) by moura + obtain c d where "s = Fun c []" "t = Fun d []" using st assms(1) by atomize_elim auto hence "s = t" using *(3) by fast thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ SMP N" "t \ M" proof - - obtain c where "t = Fun c []" using st assms(1) by moura + obtain c where "t = Fun c []" using st assms(1) by atomize_elim auto hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ M" "t \ SMP N" proof - - obtain c where "s = Fun c []" using st assms(1) by moura + obtain c where "s = Fun c []" using st assms(1) by atomize_elim auto hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed ultimately have "\ s = \ t" by metis } thus ?thesis by (metis tfr\<^sub>s\<^sub>e\<^sub>t_def) qed lemma dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" proof (induction S) case (Cons x S) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Cons.prems by simp hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" using Cons.IH by metis from Cons show ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons by auto thus ?thesis using Equality IH by fastforce next case (Inequality X F) have "set (dual\<^sub>s\<^sub>t (x#S)) = insert x (set (dual\<^sub>s\<^sub>t S))" using Inequality by auto moreover have "(\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = Var a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" using Cons.prems Inequality by auto ultimately show ?thesis using Inequality IH by auto qed auto qed simp lemma subst_var_inv_wt: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst_var_inv \ X)" using assms f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by presburger lemma subst_var_inv_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (subst_var_inv \ X))" using f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by auto lemma unify_list_wt_if_same_type: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" and "\(v,t) \ set B. \ (Var v) = \ t" shows "\(v,t) \ set U. \ (Var v) = \ t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (2 f X g Y E B U) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)" "\ (Fun f X) = \ (Fun g Y)" by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) have "\(s,t) \ set E'. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" proof - { fix s t assume "(s,t) \ set E'" then obtain X' X'' Y' Y'' where "X = X'@s#X''" "Y = Y'@t#Y''" "length X' = length Y'" using zip_arg_subterm_split[of s t X Y] \E' = zip X Y\ by metis hence "\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))" by (metis \\ (Fun f X) = \ (Fun g Y)\) from \E' = zip X Y\ have "\(s,t) \ set E'. s \ Fun f X \ t \ Fun g Y" using zip_arg_subterm[of _ _ X Y] by blast with \(s,t) \ set E'\ have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subterm \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)\ by (blast,blast) moreover have "f \ \\<^sub>f" proof (rule ccontr) assume "f \ \\<^sub>f" hence "f \ \" "arity f = 0" using const_arity_eq_zero[of f] by simp_all thus False using \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ * \(s,t) \ set E'\ unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed hence "\ s = \ t" using fun_type_arg_cong' \f \ \\<^sub>f\ \\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))\ \length X' = length Y'\ \f = g\ by metis ultimately have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" by metis+ } thus ?thesis by blast qed moreover have "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" using "2.prems"(2) by auto ultimately show ?case using "2.IH"[OF * ** _ "2.prems"(3)] by fastforce next case (3 v t E B U) hence "\ (Var v) = \ t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)" and *: "\(v, t) \ set ((v,t)#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "3.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto show ?case proof (cases "t = Var v") assume "t = Var v" thus ?case using 3 by auto next assume "t \ Var v" hence "v \ fv t" using "3.prems"(1) by auto hence **: "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using Unification.unify.simps(3)[of v t E B] "3.prems"(1) \t \ Var v\ by auto have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)\]] by (simp add: subst_list_def) ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ ** _ *(1)] by auto qed next case (4 f X v E B U) hence "\ (Var v) = \ (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))" and *: "\(v, t) \ set ((v,(Fun f X))#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "4.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto have "v \ fv (Fun f X)" using "4.prems"(1) by force hence **: "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, (Fun f X))#B) = Some U" using Unification.unify.simps(3)[of v "Fun f X" E B] "4.prems"(1) by auto have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\] "4.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))\]] by (simp add: subst_list_def) ultimately show ?case using "4.IH"[OF \v \ fv (Fun f X)\ ** _ *(1)] by auto qed auto lemma mgu_wt_if_same_type: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - let ?fv_disj = "\v t S. \(\(v',t') \ S - {(v,t)}. (insert v (fv t)) \ (insert v' (fv t')) \ {})" from assms(1) obtain \' where "Unification.unify [(s,t)] [] = Some \'" "subst_of \' = \" by (auto simp: mgu_def split: option.splits) hence "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using assms(2,3,4) unify_list_wt_if_same_type unify_list_distinct[of "[(s,t)]"] by auto thus "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \subst_of \' = \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \' arbitrary: \ rule: List.rev_induct) case (snoc tt \' \) then obtain v t where tt: "tt = (v,t)" by (metis surj_pair) hence \: "\ = subst v t \\<^sub>s subst_of \'" using snoc.prems(3) by simp have "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using snoc.prems(1,2) by auto then obtain \'' where \'': "subst_of \' = \''" "\v. \ (Var v) = \ (\'' v)" by (metis snoc.IH) hence "\ t = \ (t \ \'')" for t using wt_subst_trm by blast hence "\ (Var v) = \ (\'' v)" "\ t = \ (t \ \'')" using \''(2) by auto moreover have "\ (Var v) = \ t" using snoc.prems(1) tt by simp moreover have \2: "\ = Var(v := t) \\<^sub>s \'' " using \ \''(1) unfolding subst_def by simp ultimately have "\ (Var v) = \ (\ v)" unfolding subst_compose_def by simp have "subst_domain (subst v t) \ {v}" unfolding subst_def by (auto simp add: subst_domain_def) hence *: "subst_domain \ \ insert v (subst_domain \'')" using tt \ \''(1) snoc.prems(2) subst_domain_compose[of _ \''] by (auto simp add: subst_domain_def) have "v \ set (map fst \')" using tt snoc.prems(2) by auto hence "v \ subst_domain \''" using \''(1) subst_of_dom_subset[of \'] by auto { fix w assume "w \ subst_domain \''" hence "\ w = \'' w" using \2 \''(1) \v \ subst_domain \''\ unfolding subst_compose_def by auto hence "\ (Var w) = \ (\ w)" using \''(2) by simp } thus ?case using \\ (Var v) = \ (\ v)\ * by force qed simp qed lemma wt_Unifier_if_Unifier: assumes s_t: "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" and \: "Unifier \ s t" shows "\\. Unifier \ s t \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_always_unifies[OF \] mgu_gives_MGU[THEN MGU_is_Unifier[of s _ t]] mgu_wt_if_same_type[OF _ s_t] mgu_wf_trm[OF _ s_t(1,2)] wf_trm_subst_range_iff by fast end subsection \Automatically Proving Type-Flaw Resistance\ subsubsection \Definitions: Variable Renaming\ abbreviation "max_var t \ Max (insert 0 (snd ` fv t))" abbreviation "max_var_set X \ Max (insert 0 (snd ` X))" definition "var_rename n v \ Var (fst v, snd v + Suc n)" definition "var_rename_inv n v \ Var (fst v, snd v - Suc n)" subsubsection \Definitions: Computing a Finite Representation of the Sub-Message Patterns\ text \A sufficient requirement for a term to be a well-typed instance of another term\ definition is_wt_instance_of_cond where "is_wt_instance_of_cond \ t s \ ( \ t = \ s \ (case mgu t s of None \ False | Some \ \ inj_on \ (fv t) \ (\x \ fv t. is_Var (\ x))))" definition has_all_wt_instances_of where "has_all_wt_instances_of \ N M \ \t \ N. \s \ M. is_wt_instance_of_cond \ t s" text \This function computes a finite representation of the set of sub-message patterns\ definition SMP0 where "SMP0 Ana \ M \ let f = \t. Fun (the_Fun (\ t)) (map Var (zip (args (\ t)) [0.. t))])); g = \M'. map f (filter (\t. is_Var t \ is_Fun (\ t)) M')@ concat (map (fst \ Ana) M')@concat (map subterms_list M'); h = remdups \ g in while (\A. set (h A) \ set A) h M" text \These definitions are useful to refine an SMP representation set\ fun generalize_term where "generalize_term _ _ n (Var x) = (Var x, n)" | "generalize_term \ p n (Fun f T) = (let \ = \ (Fun f T) in if p \ then (Var (\, n), Suc n) else let (T',n') = foldr (\t (S,m). let (t',m') = generalize_term \ p m t in (t'#S,m')) T ([],n) in (Fun f T', n'))" definition generalize_terms where "generalize_terms \ p \ map (fst \ generalize_term \ p 0)" definition remove_superfluous_terms where "remove_superfluous_terms \ T \ let f = \S t R. \s \ set S - R. s \ t \ is_wt_instance_of_cond \ t s; g = \S t (U,R). if f S t R then (U, insert t R) else (t#U, R); h = \S. remdups (fst (foldr (g S) S ([],{}))) in while (\S. h S \ S) h T" subsubsection \Definitions: Checking Type-Flaw Resistance\ definition is_TComp_var_instance_closed where "is_TComp_var_instance_closed \ M \ \x \ fv\<^sub>s\<^sub>e\<^sub>t M. is_Fun (\ (Var x)) \ (\t \ M. is_Fun t \ \ t = \ (Var x) \ list_all is_Var (args t) \ distinct (args t))" definition finite_SMP_representation where "finite_SMP_representation arity Ana \ M \ (M = {} \ card M > 0) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M \ has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M \ has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M \ is_TComp_var_instance_closed \ M" definition comp_tfr\<^sub>s\<^sub>e\<^sub>t where "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M \ finite_SMP_representation arity Ana \ M \ (let \ = var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)) in \s \ M. \t \ M. is_Fun s \ is_Fun t \ \ s \ \ t \ mgu s (t \ \) = None)" fun comp_tfr\<^sub>s\<^sub>t\<^sub>p where "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\_: t \ t'\\<^sub>s\<^sub>t) = (mgu t t' \ None \ \ t = \ t')" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\X\\\: F\\<^sub>s\<^sub>t) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. is_Var (\ (Var x))) \ (\u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F). is_Fun u \ (args u = [] \ (\s \ set (args u). s \ Var ` set X))))" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p _ _ = True" definition comp_tfr\<^sub>s\<^sub>t where "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S \ list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (trms_list\<^sub>s\<^sub>t S) \ has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" subsubsection \Small Lemmata\ lemma max_var_set_mono: assumes "finite N" and "M \ N" shows "max_var_set M \ max_var_set N" by (meson assms Max.subset_imp finite.insertI finite_imageI image_mono insert_mono insert_not_empty) lemma less_Suc_max_var_set: assumes z: "z \ X" and X: "finite X" shows "snd z < Suc (max_var_set X)" proof - have "snd z \ snd ` X" using z by simp hence "snd z \ Max (insert 0 (snd ` X))" using X by simp thus ?thesis using X by simp qed lemma (in typed_model) finite_SMP_representationD: assumes "finite_SMP_representation arity Ana \ M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" and "is_TComp_var_instance_closed \ M" and "finite M" using assms wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[of M] unfolding finite_SMP_representation_def list_all_iff card_gt_0_iff by blast+ lemma (in typed_model) is_wt_instance_of_condD: assumes t_instance_s: "is_wt_instance_of_cond \ t s" obtains \ where "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" using t_instance_s unfolding is_wt_instance_of_cond_def Let_def by (cases "mgu t s") fastforce+ lemma (in typed_model) is_wt_instance_of_condD': assumes t_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m t" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" proof - obtain \ where s: "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis is_wt_instance_of_condD[OF t_instance_s]) have 0: "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" using s(1) t_wf_trm s_wf_trm by auto note 1 = mgu_wt_if_same_type[OF s(2) 0 s(1)] note 2 = conjunct1[OF mgu_gives_MGU[OF s(2)]] show ?thesis using s(1) inj_var_ran_unifiable_has_subst_match[OF 2 s(3,4)] wt_subst_compose[OF 1 subst_var_inv_wt[OF 1, of "fv t"]] wf_trms_subst_compose[OF mgu_wf_trms[OF s(2) 0] subst_var_inv_wf_trms[of \ "fv t"]] by auto qed lemma (in typed_model) is_wt_instance_of_condD'': assumes s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" and t_var: "t = Var x" shows "\y. s = Var y \ \ (Var y) = \ (Var x)" proof - obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and s: "Var x = s \ \" using is_wt_instance_of_condD'[OF _ s_wf_trm t_instance_s] t_var by auto obtain y where y: "s = Var y" using s by (cases s) auto show ?thesis using wt_subst_trm''[OF \] s y by metis qed lemma (in typed_model) has_all_wt_instances_ofD: assumes N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" obtains s \ where "s \ M" "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis t_in_N N_instance_M is_wt_instance_of_condD has_all_wt_instances_of_def) lemma (in typed_model) has_all_wt_instances_ofD': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t \ M \\<^sub>s\<^sub>e\<^sub>t \" using assms is_wt_instance_of_condD' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_wt_instances_ofD'': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "Var x \ N" shows "\y. Var y \ M \ \ (Var y) = \ (Var x)" using assms is_wt_instance_of_condD'' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_instances_of_if_subset: assumes "N \ M" shows "has_all_wt_instances_of \ N M" unfolding has_all_wt_instances_of_def proof fix t assume t: "t \ N" hence "is_wt_instance_of_cond \ t t" using inj_onI[of "fv t"] mgu_same_empty[of t] unfolding is_wt_instance_of_cond_def by force thus "\s \ M. is_wt_instance_of_cond \ t s" using t assms by blast qed lemma (in typed_model) SMP_I': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "t \ SMP M" using has_all_wt_instances_ofD'[OF N_wf_trms M_wf_trms N_instance_M t_in_N] SMP.Substitution[OF SMP.MP[of _ M]] by blast subsubsection \Lemma: Proving Type-Flaw Resistance\ locale typed_model' = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" + assumes \_Var_fst: "\\ n m. \ (Var (\,n)) = \ (Var (\,m))" and Ana_const: "\c T. arity c = 0 \ Ana (Fun c T) = ([],[])" and Ana_subst'_or_Ana_keys_subterm: "(\f T \ K R. Ana (Fun f T) = (K,R) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)) \ (\t K R k. Ana t = (K,R) \ k \ set K \ k \ t)" begin lemma var_rename_inv_comp: "t \ (var_rename n \\<^sub>s var_rename_inv n) = t" proof (induction t) case (Fun f T) hence "map (\t. t \ var_rename n \\<^sub>s var_rename_inv n) T = T" by (simp add: map_idI) thus ?case by (metis eval_term.simps(2)) qed (simp add: var_rename_def var_rename_inv_def subst_compose) lemma var_rename_fv_disjoint: "fv s \ fv (t \ var_rename (max_var s)) = {}" proof - have 1: "\v \ fv s. snd v \ max_var s" by simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint: assumes "finite M" "s \ M" shows "fv s \ fv (t \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" proof - have 1: "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using assms proof (induction M rule: finite_induct) case (insert t M) thus ?case proof (cases "t = s") case False hence "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using insert by simp moreover have "max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M) \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t (insert t M))" using insert.hyps(1) insert.prems by force ultimately show ?thesis by auto qed simp qed simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint': assumes "finite M" shows "fv\<^sub>s\<^sub>e\<^sub>t M \ fv\<^sub>s\<^sub>e\<^sub>t (N \\<^sub>s\<^sub>e\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" using var_rename_fv_set_disjoint[OF assms] by auto lemma var_rename_is_renaming[simp]: "subst_range (var_rename n) \ range Var" "subst_range (var_rename_inv n) \ range Var" unfolding var_rename_def var_rename_inv_def by auto lemma var_rename_wt[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename n)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n)" by (auto simp add: var_rename_def var_rename_inv_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \_Var_fst) lemma var_rename_wt': assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "s = m \ \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n \\<^sub>s \)" "s = m \ var_rename n \ var_rename_inv n \\<^sub>s \" using assms(2) wt_subst_compose[OF var_rename_wt(2)[of n] assms(1)] var_rename_inv_comp[of m n] by force+ lemma var_rename_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_range[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename n))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename_inv n))" using var_rename_is_renaming by fastforce+ lemma Fun_range_case: "(\f T. Fun f T \ M \ P f T) \ (\u \ M. case u of Fun f T \ P f T | _ \ True)" "(\f T. Fun f T \ M \ P f T) \ (\u \ M. is_Fun u \ P (the_Fun u) (args u))" by (auto split: "term.splits") lemma is_TComp_var_instance_closedD: assumes x: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "\ (Var x) = TComp f T" and closed: "is_TComp_var_instance_closed \ M" shows "\g U. Fun g U \ M \ \ (Fun g U) = \ (Var x) \ (\u \ set U. is_Var u) \ distinct U" using assms unfolding is_TComp_var_instance_closed_def list_all_iff list_ex_iff by fastforce lemma is_TComp_var_instance_closedD': assumes "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "TComp f T \ \ (Var x)" and closed: "is_TComp_var_instance_closed \ M" and wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\g U. Fun g U \ M \ \ (Fun g U) = TComp f T \ (\u \ set U. is_Var u) \ distinct U" using assms(1,2) proof (induction "\ (Var x)" arbitrary: x) case (Fun g U) note IH = Fun.hyps(1) have g: "arity g > 0" using Fun.hyps(2) fun_type_inv[of "Var x"] \_Var_fst by simp_all then obtain V where V: "Fun g V \ M" "\ (Fun g V) = \ (Var x)" "\v \ set V. \x. v = Var x" "distinct V" "length U = length V" using is_TComp_var_instance_closedD[OF Fun.prems(1) Fun.hyps(2)[symmetric] closed(1)] by (metis Fun.hyps(2) fun_type_id_eq fun_type_length_eq is_VarE) hence U: "U = map \ V" using fun_type[OF g(1), of V] Fun.hyps(2) by simp hence 1: "\ v \ set U" when v: "v \ set V" for v using v by simp have 2: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var y)" when z: "Var z \ set V" for z using V(1) fv_subset_subterms Fun_param_in_subterms[OF z] by fastforce show ?case proof (cases "TComp f T = \ (Var x)") case False then obtain u where u: "u \ set U" "TComp f T \ u" - using Fun.prems(2) Fun.hyps(2) by moura + using Fun.prems(2) Fun.hyps(2) by atomize_elim auto then obtain y where y: "Var y \ set V" "\ (Var y) = u" using U V(3) \_Var_fst by auto show ?thesis using IH[OF _ 2[OF y(1)]] u y(2) by metis qed (use V in fastforce) qed simp lemma TComp_var_instance_wt_subst_exists: assumes gT: "\ (Fun g T) = TComp g (map \ U)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g T)" and U: "\u \ set U. \y. u = Var y" "distinct U" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Fun g T = Fun g U \ \" proof - define the_i where "the_i \ \y. THE x. x < length U \ U ! x = Var y" define \ where \: "\ \ \y. if Var y \ set U then T ! the_i y else Var y" have g: "arity g > 0" using gT(1,2) fun_type_inv(1) by blast have UT: "length U = length T" using fun_type_length_eq gT(1) by fastforce have 1: "the_i y < length U \ U ! the_i y = Var y" when y: "Var y \ set U" for y using theI'[OF distinct_Ex1[OF U(2) y]] unfolding the_i_def by simp have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ 1 gT(1) fun_type[OF g] UT unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis (no_types, lifting) nth_map term.inject(2)) have "\i \ = T ! i" using \ 1 U(1) UT distinct_Ex1[OF U(2)] in_set_conv_nth by (metis (no_types, lifting) eval_term.simps(1)) hence "T = map (\t. t \ \) U" by (simp add: UT nth_equalityI) hence 3: "Fun g T = Fun g U \ \" by simp have "subst_range \ \ set T" using \ 1 U(1) UT by (auto simp add: subst_domain_def) hence 4: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using gT(2) wf_trm_param by auto show ?thesis by (metis 2 3 4) qed lemma TComp_var_instance_closed_has_Var: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ t" proof (cases "\ (Var x)") case (Var a) hence "t = \ x" using t wf_\x \_wt by (metis (full_types) const_type_inv_wf fun_if_subterm subtermeq_Var_const(2) wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y_ex wt_subst_trm''[OF \_wt, of "Var x"] by fastforce next case (Fun f T) hence \_\x: "\ (\ x) = TComp f T" using wt_subst_trm''[OF \_wt, of "Var x"] by auto show ?thesis proof (cases "t = \ x") case False hence t_subt_\x: "t \ \ x" using t(1) \_\x by fastforce obtain T' where T': "\ x = Fun f T'" using \_\x t_subt_\x fun_type_id_eq by (cases "\ x") auto obtain g S where gS: "Fun g S \ \ x" "t \ set S" using Fun_ex_if_subterm[OF t_subt_\x] by blast have gS_wf: "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" by (rule wf_trm_subtermeq[OF wf_\x gS(1)]) hence "arity g > 0" using gS(2) by (metis length_pos_if_in_set wf_trm_arity) hence gS_\: "\ (Fun g S) = TComp g (map \ S)" using fun_type by blast obtain h U where hU: "Fun h U \ M" "\ (Fun h U) = Fun g (map \ S)" "\u \ set U. is_Var u" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] subtermeq_imp_subtermtypeeq[OF wf_\x] gS \_\x Fun gS_\ by metis obtain y where y: "Var y \ set U" "\ (Var y) = \ t" using hU(3) fun_type_param_ex[OF hU(2) gS(2)] by fast have "y \ fv\<^sub>s\<^sub>e\<^sub>t M" using hU(1) y(1) by force thus ?thesis using y(2) closed by metis qed (metis y_ex Fun \_\x) qed lemma TComp_var_instance_closed_has_Fun: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and t_\: "\ t = TComp g T" and t_fun: "is_Fun t" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m" proof - obtain T'' where T'': "t = Fun g T''" using t_\ t_fun fun_type_id_eq by blast have g: "arity g > 0" using t_\ fun_type_inv[of t] by simp_all have "TComp g T \ \ (Var x)" using \_wt t t_\ by (metis wf_\x subtermeq_imp_subtermtypeeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) then obtain U where U: "Fun g U \ M" "\ (Fun g U) = TComp g T" "\u \ set U. \y. u = Var y" "distinct U" "length T'' = length U" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] by (metis t_\ T'' fun_type_id_eq fun_type_length_eq is_VarE) hence UT': "T = map \ U" using fun_type[OF g, of U] by simp show ?thesis using TComp_var_instance_wt_subst_exists UT' T'' U(1,3,4) t t_\ wf_\x wf_trm_subtermeq by (metis term.disc(2)) qed lemma TComp_var_and_subterm_instance_closed_has_subterms_instances: assumes M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and t: "t \\<^sub>s\<^sub>e\<^sub>t M" and s: "s \ t \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ s = m \ \" using subterm_subst_unfold[OF s] proof assume "\s'. s' \ t \ s = s' \ \" then obtain s' where s': "s' \ t" "s = s' \ \" by blast then obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s' \ M \\<^sub>s\<^sub>e\<^sub>t \" using t has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] term.order_trans[of s' t] by blast then obtain m where m: "m \ M" "s' = m \ \" by blast have "s = m \ (\ \\<^sub>s \)" using s'(2) m(2) by simp thus ?thesis using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv t. s \ \ x" then obtain x where x: "x \ fv t" "s \ \ x" "s \ \ x" by blast note 0 = TComp_var_instance_closed_has_Var[OF M_var_inst_cl M_wf] note 1 = has_all_wt_instances_ofD''[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" using \(2) wf_trm_subterm[OF _ x(2)] by fastforce+ have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x(1) s fv_subset_subterms[OF t] by auto obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ s" using 0[of \ x s, OF \x_wf x_fv_ex x(3) \(1)] by metis then obtain z where z: "Var z \ M" "\ (Var z) = \ s" using 1[of y] vars_iff_subtermeq_set[of y M] by metis define \ where "\ \ Var(z := s)::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s = Var z \ \" using z(2) s_wf_trm unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using z(1) by blast qed context begin private lemma SMP_D_aux1: assumes "t \ SMP M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\x \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using assms(1) proof (induction t rule: SMP.induct) case (MP t) show ?case proof fix x assume x: "x \ fv t" hence "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t M" using MP.hyps vars_iff_subtermeq by fastforce then obtain \ s where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and s: "s \ M" "Var x = s \ \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF wf_M] wf_M closed(1)] by blast then obtain y where y: "s = Var y" by (cases s) auto thus "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using s wt_subst_trm''[OF \(1), of "Var y"] by force qed next case (Subterm t t') hence "fv t' \ fv t" using subtermeq_vars_subset by auto thus ?case using Subterm.IH by auto next case (Substitution t \) note IH = Substitution.IH show ?case proof fix x assume x: "x \ fv (t \ \)" then obtain y where y: "y \ fv t" "\ (Var x) \ \ (Var y)" using Substitution.hyps(2,3) by (metis subst_apply_img_var subtermeqI' subtermeq_imp_subtermtypeeq vars_iff_subtermeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def wf_trm_subst_rangeD) let ?P = "\x. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" show "?P x" using y IH proof (induction "\ (Var y)" arbitrary: y t) case (Var a) hence "\ (Var x) = \ (Var y)" by auto thus ?case using Var(2,4) by auto next case (Fun f T) obtain z where z: "\w \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var w)" "\ (Var z) = \ (Var y)" using Fun.prems(1,3) by blast show ?case proof (cases "\ (Var x) = \ (Var y)") case True thus ?thesis using Fun.prems by auto next case False then obtain \ where \: "\ \ set T" "\ (Var x) \ \" using Fun.prems(2) Fun.hyps(2) by auto then obtain U where U: "Fun f U \ M" "\ (Fun f U) = \ (Var z)" "\u \ set U. \v. u = Var v" "distinct U" using is_TComp_var_instance_closedD'[OF z(1) _ closed(2) wf_M] Fun.hyps(2) z(2) by (metis fun_type_id_eq subtermeqI' is_VarE) hence 1: "\x \ fv (Fun f U). \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" by force have "arity f > 0" using U(2) z(2) Fun.hyps(2) fun_type_inv(1) by metis hence "\ (Fun f U) = TComp f (map \ U)" using fun_type by auto then obtain u where u: "Var u \ set U" "\ (Var u) = \" using \(1) U(2,3) z(2) Fun.hyps(2) by auto show ?thesis using Fun.hyps(1)[of u "Fun f U"] u \ 1 by force qed qed qed next case (Ana t K T k) have "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto thus ?case using Ana.IH by auto qed private lemma SMP_D_aux2: fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and t_Var: "\x. t = Var x" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof - have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = has_all_wt_instances_ofD'[OF M_Ana_wf M_wf M_Ana_cl] obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using t_Var SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using 0[of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed private lemma SMP_D_aux3: assumes hyps: "t' \ t" and wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" and prems: "is_Fun t'" and IH: "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t' = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \ \ is_Fun m)" proof (cases "\f. t = Fun f [] \ t' = Fun f []") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = TComp_var_instance_closed_has_Fun[OF M_var_inst_cl M_wf] note 2 = TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF M_var_inst_cl M_subterms_cl M_wf] have wf_t': "wf\<^sub>t\<^sub>r\<^sub>m t'" using hyps wf_t wf_trm_subterm by blast - obtain c where "t = Fun c [] \ t' = Fun c []" using True by moura + obtain c where "t = Fun c [] \ t' = Fun c []" using True by atomize_elim auto thus ?thesis proof assume c: "t' = Fun c []" show ?thesis proof (cases "\f. t = Fun f []") case True hence "t = t'" using c hyps by force thus ?thesis using IH by fast next case False note F = this then obtain m \ where m: "m \ M" "t = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using IH by blast show ?thesis using subterm_subst_unfold[OF hyps[unfolded m(2)]] proof assume "\m'. m' \ m \ t' = m' \ \" - then obtain m' where m': "m' \ m" "t' = m' \ \" by moura + then obtain m' where m': "m' \ m" "t' = m' \ \" by atomize_elim auto obtain n \ where n: "n \ M" "m' = n \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using 0[of m'] m(1) m'(1) by blast have "t' = n \ (\ \\<^sub>s \)" using m'(2) n(2) by auto thus ?thesis using c n(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv m. t' \ \ x" - then obtain x where x: "x \ fv m" "t' \ \ x" "t' \ \ x" by moura + then obtain x where x: "x \ fv m" "t' \ \ x" "t' \ \ x" by atomize_elim auto have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \(2) by fastforce have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x m by auto show ?thesis proof (cases "\ t'") case (Var a) show ?thesis using c m 2[OF _ hyps[unfolded m(2)] \] by fast next case (Fun g S) show ?thesis using c 1[of \ x t', OF \x_wf x_fv_ex x(3) \(1) Fun] by blast qed qed qed qed (use IH hyps in simp) next case False note F = False then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" - using IH by moura + using IH by atomize_elim auto obtain f T where fT: "t' = Fun f T" "arity f > 0" "\ t' = TComp f (map \ T)" using F prems fun_type wf_trm_subtermeq[OF wf_t hyps] by (metis is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) have closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" using M_SMP_repr unfolding finite_SMP_representation_def by metis+ have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using finite_SMP_representationD[OF M_SMP_repr] by blast show ?thesis proof (cases "\x \ fv m. t' \ \ x") case True - then obtain x where x: "x \ fv m" "t' \ \ x" by moura + then obtain x where x: "x \ fv m" "t' \ \ x" by atomize_elim auto have 1: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using m(1) x(1) by auto have 2: "is_Fun (\ x)" using prems x(2) by auto have 3: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using m(5) by (simp add: wf_trm_subst_rangeD) have "\(\f. \ x = Fun f [])" using F x(2) by auto hence "\f T. \ (Var x) = TComp f T" using 2 3 m(2) by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "\f T. \ t' = Fun f T" using False prems wf_trm_subtermeq[OF wf_t hyps] by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using TComp_var_instance_closed_has_Fun 1 x(2) m(2) prems closed 3 M_wf by metis next case False then obtain m' where m': "m' \ m" "t' = m' \ \" "is_Fun m'" using hyps m(3) subterm_subst_not_img_subterm by blast then obtain \ m'' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "m'' \ M" "m' = m'' \ \" using m(1) has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf closed(1)] by blast hence t'_m'': "t' = m'' \ \ \\<^sub>s \" using m'(2) by fastforce note \\ = wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(5)] show ?thesis proof (cases "is_Fun m''") case True thus ?thesis using \(3,4) m'(2,3) m(4) fT t'_m'' \\ by blast next case False - then obtain x where x: "m'' = Var x" by moura + then obtain x where x: "m'' = Var x" by atomize_elim auto hence "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "t' \ (\ \\<^sub>s \) x" "\ (Var x) = Fun f (map \ T)" "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" using \\ t'_m'' \(3) fv_subset[OF \(3)] fT(3) eval_term.simps(1)[of _ "\ \\<^sub>s \"] wt_subst_trm''[OF \\(1), of "Var x"] by force+ thus ?thesis using x TComp_var_instance_closed_has_Fun[ of M "\ \\<^sub>s \" x t' f "map \ T", OF closed(2) M_wf _ _ _ \\(1) fT(3) prems] by blast qed qed qed lemma SMP_D: assumes "t \ SMP M" "is_Fun t" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" proof - have wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" "is_TComp_var_instance_closed \ M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show ?thesis using assms(1,2) proof (induction t rule: SMP.induct) case (MP t) moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "t = t \ Var" by simp_all ultimately show ?case by blast next case (Subterm t t') hence t_fun: "is_Fun t" by auto note * = Subterm.hyps(2) SMP_wf_trm[OF Subterm.hyps(1) wf_M(1)] Subterm.prems Subterm.IH[OF t_fun] M_SMP_repr show ?case by (rule SMP_D_aux3[OF *]) next case (Substitution t \) have wf: "wf\<^sub>t\<^sub>r\<^sub>m t" by (metis Substitution.hyps(1) wf_M(1) SMP_wf_trm) hence wf': "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Substitution.hyps(3) wf_trm_subst by blast show ?case proof (cases "\ t") case (Var a) hence 1: "\ (t \ \) = TAtom a" using Substitution.hyps(2) by (metis wt_subst_trm'') then obtain c where c: "t \ \ = Fun c []" using TAtom_term_cases[OF wf' 1] Substitution.prems by fastforce hence "(\x. t = Var x) \ t = t \ \" by (cases t) auto thus ?thesis proof assume t_Var: "\x. t = Var x" then obtain x where x: "t = Var x" "\ x = Fun c []" "\ (Var x) = TAtom a" using c 1 wt_subst_trm''[OF Substitution.hyps(2), of t] by force obtain m \ where m: "m \ M" "t = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" - using SMP_D_aux2[OF Substitution.hyps(1) t_Var M_SMP_repr] by moura + using SMP_D_aux2[OF Substitution.hyps(1) t_Var M_SMP_repr] by atomize_elim auto have "m \ (\ \\<^sub>s \) = Fun c []" using c m(2) by auto thus ?thesis using c m(1) wt_subst_compose[OF \(1) Substitution.hyps(2)] wf_trms_subst_compose[OF \(2) Substitution.hyps(3)] by metis qed (use c Substitution.IH in auto) next case (Fun f T) hence 1: "\ (t \ \) = TComp f T" using Substitution.hyps(2) by (metis wt_subst_trm'') have 2: "\(\f. t = Fun f [])" using Fun TComp_term_cases[OF wf] by auto obtain T'' where T'': "t \ \ = Fun f T''" using 1 2 fun_type_id_eq Fun Substitution.prems by fastforce have f: "arity f > 0" using fun_type_inv[OF 1] by metis show ?thesis proof (cases t) case (Fun g U) then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" - using Substitution.IH Fun 2 by moura + using Substitution.IH Fun 2 by atomize_elim auto have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "t \ \ = m \ (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wt_subst_compose[OF m(2) Substitution.hyps(2)] m(3) wf_trms_subst_compose[OF m(5) Substitution.hyps(3)] by auto thus ?thesis using m(1,4) by metis next case (Var x) then obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF Substitution.hyps(1) closed(1,3) wf_M] Fun - by moura + by atomize_elim auto hence 3: "\ (Var y) = TComp f T" using Var Fun \_Var_fst by simp obtain h V where V: "Fun h V \ M" "\ (Fun h V) = \ (Var y)" "\u \ set V. \z. u = Var z" "distinct V" by (metis is_VarE is_TComp_var_instance_closedD[OF _ 3 closed(3)] y(1)) moreover have "length T'' = length V" using 3 V(2) fun_type_length_eq 1 T'' by metis ultimately have TV: "T = map \ V" by (metis fun_type[OF f(1)] 3 fun_type_id_eq term.inject(2)) obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ \ = Fun h V \ \" using TComp_var_instance_wt_subst_exists 1 3 T'' TV V(2,3,4) wf' by (metis fun_type_id_eq) have 9: "\ (Fun h V) = \ (\ x)" using y(2) Substitution.hyps(2) V(2) 1 3 Var by auto show ?thesis using Var \ 9 V(1) by force qed qed next case (Ana t K T k) have 1: "is_Fun t" using Ana.hyps(2,3) by auto - then obtain f U where U: "t = Fun f U" by moura + then obtain f U where U: "t = Fun f U" by atomize_elim auto have 2: "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto have wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF Ana.hyps(1)] wf\<^sub>t\<^sub>r\<^sub>m_code wf_M by auto hence wf_k: "wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf'[OF Ana.hyps(2)] wf\<^sub>t\<^sub>r\<^sub>m_code Ana.hyps(3) by auto have wf_M_keys: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\((set \ fst \ Ana) ` M))" proof fix t assume "t \ (\((set \ fst \ Ana) ` M))" then obtain s where s: "s \ M" "t \ (set \ fst \ Ana) s" by blast obtain K R where KR: "Ana s = (K,R)" by (metis surj_pair) hence "t \ set K" using s(2) by simp thus "wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF KR] wf_M s(1) by blast qed show ?case using Ana_subst'_or_Ana_keys_subterm proof assume "\t K T k. Ana t = (K, T) \ k \ set K \ k \ t" hence *: "k \ t" using Ana.hyps(2,3) by auto show ?thesis by (rule SMP_D_aux3[OF * wf_t Ana.prems Ana.IH[OF 1] M_SMP_repr]) next assume Ana_subst': "\f T \ K M. Ana (Fun f T) = (K, M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" have "arity f > 0" using Ana_const[of f U] U Ana.hyps(2,3) by fastforce hence "U \ []" using wf_t U unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = m \ \" "is_Fun m" using Ana.IH[OF 1] U by auto hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U Ana.hyps(2) by auto - obtain Km Tm where Ana_m: "Ana m = (Km,Tm)" by moura + obtain Km Tm where Ana_m: "Ana m = (Km,Tm)" by atomize_elim auto hence "Ana (m \ \) = (Km \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,Tm \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U m(4) is_FunE[OF m(5)] Ana.hyps(2) by metis then obtain km where km: "km \ set Km" "k = km \ \" using Ana.hyps(2,3) m(4) by auto then obtain \ km' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and km': "km' \ M" "km = km' \ \" using Ana_m m(1) has_all_wt_instances_ofD'[OF wf_M_keys wf_M closed(2), of km] by force have k\\: "k = km' \ \ \\<^sub>s \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using km(2) km' wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(3)] by auto show ?case proof (cases "is_Fun km'") case True thus ?thesis using k\\ km'(1) by blast next case False note F = False then obtain x where x: "km' = Var x" by auto hence 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using fv_subset[OF km'(1)] by auto obtain kf kT where kf: "k = Fun kf kT" using Ana.prems by auto show ?thesis proof (cases "kT = []") case True thus ?thesis using k\\(1) k\\(2) k\\(3) kf km'(1) by blast next case False hence 4: "arity kf > 0" using wf_k kf TAtom_term_cases const_type by fastforce then obtain kT' where kT': "\ k = TComp kf kT'" by (simp add: fun_type kf) then obtain V where V: "Fun kf V \ M" "\ (Fun kf V) = \ (Var x)" "\u \ set V. \v. u = Var v" "distinct V" "is_Fun (Fun kf V)" using is_TComp_var_instance_closedD[OF _ _ closed(3), of x] x m(2) k\\(1) 3 wt_subst_trm''[OF k\\(2)] by (metis fun_type_id_eq term.disc(2) is_VarE) have 5: "kT' = map \ V" using fun_type[OF 4] x kT' k\\ m(2) V(2) by (metis term.inject(2) wt_subst_trm'') thus ?thesis using TComp_var_instance_wt_subst_exists wf_k kf 4 V(3,4) kT' V(1,5) by metis qed qed qed qed qed lemma SMP_D': fixes M defines "\ \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" assumes M_SMP_repr: "finite_SMP_representation arity Ana \ M" and s: "s \ SMP M" "is_Fun s" "\f. s = Fun f []" and t: "t \ SMP M" "is_Fun t" "\f. t = Fun f []" obtains \ s0 \ t0 where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "is_Fun t0" "t = t0 \ \ \ \" "\ t = \ t0" proof - obtain \ s0 where s0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "s = s0 \ \" "is_Fun s0" using s(3) SMP_D[OF s(1,2) M_SMP_repr] unfolding \_def by metis obtain \ t0 where t0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "t = t0 \ \ \ \" "is_Fun t0" using t(3) SMP_D[OF t(1,2) M_SMP_repr] var_rename_wt'[of _ t] wf_trms_subst_compose_Var_range(1)[OF _ var_rename_is_renaming(2)] unfolding \_def by metis have "\ s = \ s0" "\ t = \ (t0 \ \)" "\ (t0 \ \) = \ t0" using s0 t0 wt_subst_trm'' by (metis, metis, metis \_def var_rename_wt(1)) thus ?thesis using s0 t0 that by simp qed lemma SMP_D'': fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof (cases "(\x. t = Var x) \ (\c. t = Fun c [])") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed show ?thesis using True proof assume "\x. t = Var x" then obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl, of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) end lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" shows "tfr\<^sub>s\<^sub>e\<^sub>t M" proof - let ?\ = "var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" have M_SMP_repr: "finite_SMP_representation arity Ana \ M" by (metis comp_tfr\<^sub>s\<^sub>e\<^sub>t_def assms) have M_finite: "finite M" using assms card_gt_0_iff unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def by blast show ?thesis proof (unfold tfr\<^sub>s\<^sub>e\<^sub>t_def; intro ballI impI) fix s t assume "s \ SMP M - Var`\" "t \ SMP M - Var`\" hence st: "s \ SMP M" "is_Fun s" "t \ SMP M" "is_Fun t" by auto have "\(\\. Unifier \ s t)" when st_type_neq: "\ s \ \ t" proof (cases "\f. s = Fun f [] \ t = Fun f []") case False then obtain \ s0 \ t0 where s0: "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and t0: "t0 \ M" "is_Fun t0" "t = t0 \ ?\ \ \" "\ t = \ t0" using SMP_D'[OF M_SMP_repr st(1,2) _ st(3,4)] by metis hence "\(\\. Unifier \ s0 (t0 \ ?\))" using assms mgu_None_is_subst_neq st_type_neq wt_subst_trm''[OF var_rename_wt(1)] unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def Let_def by metis thus ?thesis using vars_term_disjoint_imp_unifier[OF var_rename_fv_set_disjoint[OF M_finite]] s0(1) t0(1) unfolding s0(3) t0(3) by (metis (no_types, opaque_lifting) subst_subst_compose) qed (use st_type_neq st(2,4) in auto) thus "\ s = \ t" when "\\. Unifier \ s t" by (metis that) qed qed lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t': assumes "let N = SMP0 Ana \ M in set M \ set N \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (set M)" by (rule tfr_subset(2)[ OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF conjunct2[OF assms[unfolded Let_def]]] conjunct1[OF assms[unfolded Let_def]]]) lemma tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p: "tfr\<^sub>s\<^sub>t\<^sub>p a = comp_tfr\<^sub>s\<^sub>t\<^sub>p \ a" proof (cases a) case (Equality ac t t') thus ?thesis using mgu_always_unifies[of t _ t'] mgu_gives_MGU[of t t'] by auto next case (Inequality X F) thus ?thesis using tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of X F] comp_tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of \ X F] Fun_range_case(2)[of "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)"] unfolding is_Var_def by auto qed auto lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S" shows "tfr\<^sub>s\<^sub>t S" unfolding tfr\<^sub>s\<^sub>t_def proof have comp_tfr\<^sub>s\<^sub>e\<^sub>t_M: "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" using assms unfolding comp_tfr\<^sub>s\<^sub>t_def by blast have wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and S_trms_instance_M: "has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M" using assms wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t unfolding comp_tfr\<^sub>s\<^sub>t_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def list_all_iff by blast+ show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" using tfr_subset(3)[OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF comp_tfr\<^sub>s\<^sub>e\<^sub>t_M] SMP_SMP_subset] SMP_I'[OF wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M S_trms_instance_M] by blast have "list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S" by (metis assms comp_tfr\<^sub>s\<^sub>t_def) thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by (induct S) (simp_all add: tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p) qed lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t': assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ (set (SMP0 Ana \ (trms_list\<^sub>s\<^sub>t S))) S" shows "tfr\<^sub>s\<^sub>t S" by (rule tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t[OF assms]) subsubsection \Lemmata for Checking Ground SMP (GSMP) Disjointness\ context begin private lemma ground_SMP_disjointI_aux1: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and g_def: "g \ \M. {t \ M. fv t = {}}" shows "f (SMP M) = g (SMP M)" proof have "t \ f (SMP M)" when t: "t \ SMP M" "fv t = {}" for t proof - define \ where "\ \ Var::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t \ \" using subst_apply_term_empty[of t] that(2) wt_subst_Var wf_trm_subst_range_Var unfolding \_def by auto thus ?thesis using SMP.Substitution[OF t(1), of \] t(2) unfolding f_def by fastforce qed thus "g (SMP M) \ f (SMP M)" unfolding g_def by blast qed (use f_def g_def in blast) private lemma ground_SMP_disjointI_aux2: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "f M = f (SMP M)" proof have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show "f (SMP M) \ f M" proof fix t assume "t \ f (SMP M)" then obtain s \ where s: "t = s \ \" "s \ SMP M" "fv (s \ \) = {}" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF s(2) M_wf] s(1) wf_trm_subst[OF \(2)] by blast obtain m \ where m: "m \ M" "s = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using SMP_D''[OF s(2) M_SMP_repr] by blast have "t = m \ (\ \\<^sub>s \)" "fv (m \ (\ \\<^sub>s \)) = {}" using s(1,3) m(2) by simp_all thus "t \ f M" using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] unfolding f_def by blast qed qed (auto simp add: f_def) private lemma ground_SMP_disjointI_aux3: fixes A B C::"('fun, ('fun, 'atom) term \ nat) term set" defines "P \ \t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and Q_def: "Q \ \t. intruder_synth' public arity {} t" and R_def: "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" and AB: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and C: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. P t s \ Q t \ R t" shows "f A \ f B \ f C \ {m. {} \\<^sub>c m}" proof fix t assume "t \ f A \ f B" then obtain ta tb \a \b where ta: "t = ta \ \a" "ta \ A" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" "fv (ta \ \a) = {}" and tb: "t = tb \ \b" "tb \ B" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \b" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \b)" "fv (tb \ \b) = {}" unfolding f_def by blast have ta_tb_wf: "wf\<^sub>t\<^sub>r\<^sub>m ta" "wf\<^sub>t\<^sub>r\<^sub>m tb" "fv ta \ fv tb = {}" "\ ta = \ tb" using ta(1,2) tb(1,2) AB fv_subset_subterms wt_subst_trm''[OF ta(3), of ta] wt_subst_trm''[OF tb(3), of tb] by (fast, fast, blast, simp) obtain \ where \: "Unifier \ ta tb" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using vars_term_disjoint_imp_unifier[OF ta_tb_wf(3), of \a \b] ta(1) tb(1) wt_Unifier_if_Unifier[OF ta_tb_wf(1,2,4)] by blast hence "Q ta \ R ta" using ABC ta(2) tb(2) unfolding P_def by blast+ thus "t \ f C \ {m. {} \\<^sub>c m}" proof show "Q ta \ ?thesis" using ta(1) pgwt_ground[of ta] pgwt_is_empty_synth[of ta] subst_ground_ident[of ta \a] unfolding Q_def f_def intruder_synth_code[symmetric] by simp next assume "R ta" then obtain ua \a where ua: "ta = ua \ \a" "ua \ C" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" using \ ABC ta_tb_wf(1,2) ta(2) tb(2) C is_wt_instance_of_condD' unfolding P_def R_def by metis have "t = ua \ (\a \\<^sub>s \a)" "fv t = {}" using ua(1) ta(1,5) tb(1,5) by auto thus ?thesis using ua(2) wt_subst_compose[OF ua(3) ta(3)] wf_trms_subst_compose[OF ua(4) ta(4)] unfolding f_def by blast qed qed lemma ground_SMP_disjointI: fixes A B::"('fun, ('fun, 'atom) term \ nat) term set" and C defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "g \ \M. {t \ M. fv t = {}}" and "Q \ \t. intruder_synth' public arity {} t" and "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" assumes AB_fv_disj: "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and A_SMP_repr: "finite_SMP_representation arity Ana \ A" and B_SMP_repr: "finite_SMP_representation arity Ana \ B" and C_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. \ t = \ s \ mgu t s \ None \ Q t \ R t" shows "g (SMP A) \ g (SMP B) \ f C \ {m. {} \\<^sub>c m}" proof - have AB_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" using A_SMP_repr B_SMP_repr finite_SMP_representationD(1) by (blast, blast) let ?P = "\t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" have ABC': "\t \ A. \s \ B. ?P t s \ Q t \ R t" by (metis (no_types) ABC mgu_None_is_subst_neq wt_subst_trm'') show ?thesis using ground_SMP_disjointI_aux1[OF f_def g_def, of A] ground_SMP_disjointI_aux1[OF f_def g_def, of B] ground_SMP_disjointI_aux2[OF f_def A_SMP_repr] ground_SMP_disjointI_aux2[OF f_def B_SMP_repr] ground_SMP_disjointI_aux3[OF f_def Q_def R_def AB_wf AB_fv_disj C_wf ABC'] by argo qed end end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy @@ -1,3665 +1,3665 @@ (* Title: Typing_Result.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Typing Result\ text\\label{sec:Typing-Result}\ theory Typing_Result imports Typed_Model begin subsection \Locale Setup\ locale typing_result = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" + assumes infinite_typed_consts: "\a. infinite {c. \ (Fun c []) = TAtom a \ public c}" and no_private_funs[simp]: "\f. arity f > 0 \ public f" begin subsubsection \Minor Lemmata\ lemma fun_type_inv': assumes "\ t = TComp f T" shows "arity f > 0" "public f" using assms fun_type_inv by simp_all lemma infinite_public_consts[simp]: "infinite {c. public c \ arity c = 0}" proof - fix a::'atom define A where "A \ {c. \ (Fun c []) = TAtom a \ public c}" define B where "B \ {c. public c \ arity c = 0}" have "arity c = 0" when c: "c \ A" for c using c const_type_inv unfolding A_def by blast hence "A \ B" unfolding A_def B_def by blast hence "infinite B" using infinite_typed_consts[of a, unfolded A_def[symmetric]] by (metis infinite_super) thus ?thesis unfolding B_def by blast qed lemma infinite_fun_syms[simp]: "infinite {c. public c \ arity c > 0} \ infinite \\<^sub>f" "infinite \" "infinite \\<^sub>p\<^sub>u\<^sub>b" "infinite (UNIV::'fun set)" by (metis \\<^sub>f_unfold finite_Collect_conjI, metis infinite_public_consts finite_Collect_conjI, use infinite_public_consts \pub_unfold in \force simp add: Collect_conj_eq\, metis UNIV_I finite_subset subsetI infinite_public_consts(1)) lemma id_univ_proper_subset[simp]: "\\<^sub>f \ UNIV" "(\f. arity f > 0) \ \ \ UNIV" by (metis finite.emptyI inf_top.right_neutral top.not_eq_extremum disjoint_fun_syms infinite_fun_syms(2) inf_commute) (metis top.not_eq_extremum UNIV_I const_arity_eq_zero less_irrefl) lemma exists_fun_notin_funs_term: "\f::'fun. f \ funs_term t" by (metis UNIV_eq_I finite_fun_symbols infinite_fun_syms(4)) lemma exists_fun_notin_funs_terms: assumes "finite M" shows "\f::'fun. f \ \(funs_term ` M)" by (metis assms finite_fun_symbols infinite_fun_syms(4) ex_new_if_finite finite_UN) lemma exists_notin_funs\<^sub>s\<^sub>t: "\f. f \ funs\<^sub>s\<^sub>t (S::('fun,'var) strand)" by (metis UNIV_eq_I finite_funs\<^sub>s\<^sub>t infinite_fun_syms(4)) lemma infinite_typed_consts': "infinite {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" proof - { fix c assume "\ (Fun c []) = TAtom a" "public c" hence "arity c = 0" using const_type[of c] fun_type[of c "[]"] by auto } hence "{c. \ (Fun c []) = TAtom a \ public c \ arity c = 0} = {c. \ (Fun c []) = TAtom a \ public c}" by auto thus "infinite {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" using infinite_typed_consts[of a] by metis qed lemma atypes_inhabited: "\c. \ (Fun c []) = TAtom a \ wf\<^sub>t\<^sub>r\<^sub>m (Fun c []) \ public c \ arity c = 0" proof - obtain c where "\ (Fun c []) = TAtom a" "public c" "arity c = 0" using infinite_typed_consts'(1)[of a] not_finite_existsD by blast thus ?thesis using const_type_inv[OF \\ (Fun c []) = TAtom a\] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed lemma atype_ground_term_ex: "\t. fv t = {} \ \ t = TAtom a \ wf\<^sub>t\<^sub>r\<^sub>m t" using atypes_inhabited[of a] by force lemma type_ground_inhabited: "\t'. fv t' = {} \ \ t = \ t'" proof - { fix \::"('fun, 'atom) term_type" assume "\f T. Fun f T \ \ \ 0 < arity f" hence "\t'. fv t' = {} \ \ = \ t'" proof (induction \) case (Fun f T) hence "arity f > 0" by auto from Fun.IH Fun.prems(1) have "\Y. map \ Y = T \ (\x \ set Y. fv x = {})" proof (induction T) case (Cons x X) hence "\g Y. Fun g Y \ Fun f X \ 0 < arity g" by auto hence "\Y. map \ Y = X \ (\x\set Y. fv x = {})" using Cons by auto moreover have "\t'. fv t' = {} \ x = \ t'" using Cons by auto ultimately obtain y Y where "fv y = {}" "\ y = x" "map \ Y = X" "\x\set Y. fv x = {}" - using Cons by moura + using Cons by atomize_elim auto hence "map \ (y#Y) = x#X \ (\x\set (y#Y). fv x = {})" by auto thus ?case by meson qed simp then obtain Y where "map \ Y = T" "\x \ set Y. fv x = {}" by metis hence "fv (Fun f Y) = {}" "\ (Fun f Y) = TComp f T" using fun_type[OF \arity f > 0\] by auto thus ?case by (metis exI[of "\t. fv t = {} \ \ t = TComp f T" "Fun f Y"]) qed (metis atype_ground_term_ex) } thus ?thesis by (metis \_wf'') qed lemma type_wfttype_inhabited: assumes "\f T. Fun f T \ \ \ 0 < arity f" "wf\<^sub>t\<^sub>r\<^sub>m \" shows "\t. \ t = \ \ wf\<^sub>t\<^sub>r\<^sub>m t" using assms proof (induction \) case (Fun f Y) have IH: "\t. \ t = y \ wf\<^sub>t\<^sub>r\<^sub>m t" when y: "y \ set Y " for y proof - have "wf\<^sub>t\<^sub>r\<^sub>m y" using Fun y unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (metis Fun_param_is_subterm term.le_less_trans) moreover have "Fun g Z \ y \ 0 < arity g" for g Z using Fun y by auto ultimately show ?thesis using Fun.IH[OF y] by auto qed from Fun have "arity f = length Y" "arity f > 0" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force+ moreover from IH have "\X. map \ X = Y \ (\x \ set X. wf\<^sub>t\<^sub>r\<^sub>m x)" by (induct Y, simp_all, metis list.simps(9) set_ConsD) ultimately show ?case by (metis fun_type length_map wf_trmI) qed (use atypes_inhabited wf\<^sub>t\<^sub>r\<^sub>m_def in blast) lemma type_pgwt_inhabited: "wf\<^sub>t\<^sub>r\<^sub>m t \ \t'. \ t = \ t' \ public_ground_wf_term t'" proof - assume "wf\<^sub>t\<^sub>r\<^sub>m t" { fix \ assume "\ t = \" hence "\t'. \ t = \ t' \ public_ground_wf_term t'" using \wf\<^sub>t\<^sub>r\<^sub>m t\ proof (induction \ arbitrary: t) case (Var a t) then obtain c where "\ t = \ (Fun c [])" "arity c = 0" "public c" using const_type_inv[of _ "[]" a] infinite_typed_consts(1)[of a] not_finite_existsD by force thus ?case using PGWT[OF \public c\, of "[]"] by auto next case (Fun f Y t) have *: "arity f > 0" "public f" "arity f = length Y" using fun_type_inv[OF \\ t = TComp f Y\] fun_type_inv_wf[OF \\ t = TComp f Y\ \wf\<^sub>t\<^sub>r\<^sub>m t\] by auto have "\y. y \ set Y \ \t'. y = \ t' \ public_ground_wf_term t'" using Fun.prems(1) Fun.IH \_wf''[of _ _ t] \_wf'[OF \wf\<^sub>t\<^sub>r\<^sub>m t\] type_wfttype_inhabited by (metis Fun_param_is_subterm term.order_trans wf_trm_subtermeq) hence "\X. map \ X = Y \ (\x \ set X. public_ground_wf_term x)" by (induct Y, simp_all, metis list.simps(9) set_ConsD) - then obtain X where X: "map \ X = Y" "\x. x \ set X \ public_ground_wf_term x" by moura + then obtain X where X: "map \ X = Y" "\x. x \ set X \ public_ground_wf_term x" by atomize_elim auto hence "arity f = length X" using *(3) by auto have "\ t = \ (Fun f X)" "public_ground_wf_term (Fun f X)" using fun_type[OF *(1), of X] Fun.prems(1) X(1) apply simp using PGWT[OF *(2) \arity f = length X\ X(2)] by metis thus ?case by metis qed } thus ?thesis using \wf\<^sub>t\<^sub>r\<^sub>m t\ by auto qed end subsection \The Typing Result for the Composition-Only Intruder\ context typing_result begin subsubsection \Well-typedness and Type-Flaw Resistance Preservation\ context begin private lemma LI_preserves_tfr_stp_all_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" by simp_all moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (map Send1 X)" by (induct X) auto ultimately show ?case by simp next case (Unify S f Y \ X S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by simp have "fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S') \ bvars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) moreover have "fv (Fun f X) \ fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" by auto moreover have "fv (Fun f Y) \ fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" using Unify.hyps(2) fv_subset_if_in_strand_ik'[of "Fun f Y" S] by force ultimately have bvars_disj: "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f X) = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f Y) = {}" by blast+ have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) by force+ thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using SMP_append[of S "Send1 (Fun f X)#S'"] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by (metis wf_trm_subst_range_iff) moreover have "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] bvars_disj by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis next case (Equality S \ t t' a S' \) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" "\ t = \ t'" using tfr_stp_all_same_type[of S a t t' S'] tfr_stp_all_split(5)[of S _ S'] MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] Equality.prems(3) by blast+ moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by (metis wf_trm_subst_range_iff) moreover have "fv\<^sub>s\<^sub>t (S@Equality a t t'#S') \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "bvars\<^sub>s\<^sub>t (S@S') \ fv t = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv t' = {}" by auto hence "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis qed private lemma LI_in_SMP_subset_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "trms\<^sub>s\<^sub>t S \ SMP M" shows "trms\<^sub>s\<^sub>t S' \ SMP M" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)]) \ SMP M" proof - have "SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)]) \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using trms\<^sub>s\<^sub>t_append SMP_mono by auto thus ?thesis using SMP_union[of "trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" M] SMP_subset_union_eq[OF Compose.prems(6)] by auto qed thus ?case using Compose.prems(6) by auto next case (Unify S f Y \ X S' \) have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" by auto moreover have "MGU \ (Fun f X) (Fun f Y)" by (metis mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]) moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(4) by force+ moreover have "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" by (meson SMP_ikI Unify.hyps(2) contra_subsetD ik_append_subset(1)) ultimately have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" "\ (Fun f X) = \ (Fun f Y)" using ik\<^sub>s\<^sub>t_subterm_exD[OF \Fun f Y \ ik\<^sub>s\<^sub>t S\] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))\ unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis (full_types) SMP_wf_trm Unify.prems(4), blast) hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Send1 (Fun f X)#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Unify.prems(6) wt_subst_SMP_subset by metis thus ?case by auto next case (Equality S \ t t' a S' \) hence "\ t = \ t'" using tfr_stp_all_same_type MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis moreover have "t \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" "t' \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using Equality.prems(1) by auto moreover have "MGU \ t t'" using mgu_gives_MGU[OF Equality.hyps(2)[symmetric]] by metis moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(4) by force+ ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Equality.prems wt_subst_SMP_subset by metis thus ?case by auto qed private lemma LI_preserves_tfr_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S') \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) let ?SMPmap = "SMP (trms\<^sub>s\<^sub>t (S@map Send1 X@S')) - (Var`\)" have "?SMPmap \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')) - (Var`\)" using SMP_fun_map_snd_subset[of X f] SMP_append[of "map Send1 X" S'] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_append[of S "Send1 (Fun f X)#S'"] SMP_append[of S "map Send1 X@S'"] by auto hence "\s \ ?SMPmap. \t \ ?SMPmap. (\\. Unifier \ s t) \ \ s = \ t" using Compose unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Compose[OF Compose.hyps]], of S'] Compose.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Unify S f Y \ X S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Unify[OF Unify.hyps] Unify.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Send1 (Fun f X)#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Unify.prems(4) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Unify[OF Unify.hyps]], of S'] Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Equality S \ t t' a S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Equality[OF Equality.hyps] Equality.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Equality a t t'#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Equality[OF Equality.hyps]], of _ S'] Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed private lemma LI_preserves_welltypedness_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using assms proof (induction rule: LI_rel.induct) case (Unify S f Y \ X S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using SMP_append[of S "Send1 (Fun f X)#S'"] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) next case (Equality S \ t t' a S' \) have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by simp_all moreover have "\ t = \ t'" using \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S')\ MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) qed metis lemma LI_preserves_welltypedness: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" (is "?A \'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" (is "?B \'") proof - have "?A \' \ ?B \'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) hence "?A \2 \ ?B \2" using LI_preserves_welltypedness_single by presburger moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" by (fact LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)]) moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" using LI_preserves_tfr_single[OF step.hyps(1)] step.prems by presburger+ moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_tfr_stp_all_single[OF step.hyps(1)] step.prems by fastforce ultimately show ?case using step.IH by presburger qed simp thus "?A \'" "?B \'" by simp_all qed lemma LI_preserves_tfr: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S')" (is "?A S'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" (is "?B S'") and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" (is "?C S'") proof - have "?A S' \ ?B S' \ ?C S'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)] LI_preserves_tfr_single[OF step.hyps(1) step.prems(1,2)] LI_preserves_tfr_stp_all_single[OF step.hyps(1) step.prems(1,2)] step.prems(3,4,5,6) by metis+ moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \2)" using LI_preserves_welltypedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems] by simp_all ultimately show ?case using step.IH by presburger qed blast thus "?A S'" "?B S'" "?C S'" by simp_all qed lemma LI_preproc_preserves_tfr: assumes "tfr\<^sub>s\<^sub>t S" shows "tfr\<^sub>s\<^sub>t (LI_preproc S)" unfolding tfr\<^sub>s\<^sub>t_def proof have S: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms unfolding tfr\<^sub>s\<^sub>t_def by metis+ show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (LI_preproc S))" by (metis S(1) LI_preproc_trms_eq) show "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using S(2) proof (induction S) case (Cons x S) have IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using Cons by simp have x: "tfr\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems by simp show ?case using x IH unfolding list_all_iff by (cases x) auto qed simp qed end subsubsection \Simple Constraints are Well-typed Satisfiable\ text \Proving the existence of a well-typed interpretation\ context begin lemma wt_interpretation_exists: obtains \::"('fun,'var) subst" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" proof define \ where "\ = (\x. (SOME t. \ (Var x) = \ t \ public_ground_wf_term t))" { fix x t assume "\ x = t" hence "\ (Var x) = \ t \ public_ground_wf_term t" using someI_ex[of "\t. \ (Var x) = \ t \ public_ground_wf_term t", OF type_pgwt_inhabited[of "Var x"]] unfolding \_def wf\<^sub>t\<^sub>r\<^sub>m_def by simp } hence props: "\ v = t \ \ (Var v) = \ t \ public_ground_wf_term t" for v t by metis have "\ v \ Var v" for v using props pgwt_ground by fastforce hence "subst_domain \ = UNIV" by auto moreover have "ground (subst_range \)" by (simp add: props pgwt_ground) ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by metis show "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using props by simp show "subst_range \ \ public_ground_wf_terms" by (auto simp add: props) qed lemma wt_grounding_subst_exists: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" proof - obtain \ where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast show ?thesis using pgwt_wellformed interpretation_grounds[OF \(1)] \(2,3) by blast qed private fun fresh_pgwt::"'fun set \ ('fun,'atom) term_type \ ('fun,'var) term" where "fresh_pgwt S (TAtom a) = Fun (SOME c. c \ S \ \ (Fun c []) = TAtom a \ public c) []" | "fresh_pgwt S (TComp f T) = Fun f (map (fresh_pgwt S) T)" private lemma fresh_pgwt_same_type: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "\ (fresh_pgwt S (\ t)) = \ t" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\ (fresh_pgwt S \) = \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" using Fun.prems fun_type_inv by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \ (fresh_pgwt S t) = t" using Fun.prems Fun.IH by auto hence "map \ (map (fresh_pgwt S) T) = T" by (induct T) auto thus ?case using fun_type[OF f] by simp qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma fresh_pgwt_empty_synth: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "{} \\<^sub>c fresh_pgwt S (\ t)" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "{} \\<^sub>c fresh_pgwt S \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] intruder_synth.ComposeC[of "[]" _ "{}"] const_type_inv by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ {} \\<^sub>c fresh_pgwt S t" using Fun.prems Fun.IH by auto moreover have "length (map (fresh_pgwt S) T) = arity f" using f(2) by auto ultimately show ?case using intruder_synth.ComposeC[of "map (fresh_pgwt S) T" f] f by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma fresh_pgwt_has_fresh_const: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" obtains c where "Fun c [] \ fresh_pgwt S (\ t)" "c \ S" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\c. Fun c [] \ fresh_pgwt S \ \ c \ S" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" "T \ []" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto obtain t' where t': "t' \ set T" by (meson all_not_in_conv f(4) set_empty) have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \c. Fun c [] \ fresh_pgwt S t \ c \ S" using Fun.prems Fun.IH by auto then obtain c where c: "Fun c [] \ fresh_pgwt S t'" "c \ S" using t' by metis thus ?case using t' by auto qed } thus ?thesis using that assms \_wf'[OF assms(2)] \_wf'' by blast qed private lemma fresh_pgwt_subterm_fresh: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" "funs_term s \ S" shows "s \ subterms (fresh_pgwt S (\ t))" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "s \ subterms (fresh_pgwt S \)" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] assms(4) by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ s \ subterms (fresh_pgwt S t)" using Fun.prems Fun.IH by auto moreover have "s \ fresh_pgwt S (Fun f T)" proof - obtain c where c: "Fun c [] \ fresh_pgwt S (Fun f T)" "c \ S" using fresh_pgwt_has_fresh_const[OF assms(1)] type_wfttype_inhabited Fun.prems by metis hence "\Fun c [] \ s" using assms(4) subtermeq_imp_funs_term_subset by force thus ?thesis using c(1) by auto qed ultimately show ?case by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma wt_fresh_pgwt_term_exists: assumes "finite T" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" obtains t where "\ t = \ s" "{} \\<^sub>c t" "\s \ T. \u \ subterms s. u \ subterms t" proof - have finite_S: "finite (\(funs_term ` T))" using assms(1) by auto have 1: "\ (fresh_pgwt (\(funs_term ` T)) (\ s)) = \ s" using fresh_pgwt_same_type[OF finite_S assms(2)] by auto have 2: "{} \\<^sub>c fresh_pgwt (\(funs_term ` T)) (\ s)" using fresh_pgwt_empty_synth[OF finite_S assms(2)] by auto have 3: "\v \ T. \u \ subterms v. u \ subterms (fresh_pgwt (\(funs_term ` T)) (\ s))" using fresh_pgwt_subterm_fresh[OF finite_S assms(2)] assms(3) wf_trm_subtermeq subtermeq_imp_funs_term_subset by force show ?thesis by (rule that[OF 1 2 3]) qed lemma wt_bij_finite_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range Var) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range Var. \u \ subst_range Var. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (force simp add: subst_domain_def) next case (insert x S) then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp del: subst_range.simps) have *: "finite (T \ subst_range \)" using insert.prems(1) insert.hyps(1) \(1) by simp have **: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" by simp have ***: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (T \ subst_range \)" using assms(3) \(6) by blast obtain t where t: "\ t = \ (Var x)" "{} \\<^sub>c t" "\s \ T \ subst_range \. \u \ subterms s. u \ subterms t" using wt_fresh_pgwt_term_exists[OF * ** ***] by auto obtain \ where \: "\ \ \y. if x = y then t else \ y" by simp have t_ground: "fv t = {}" using t(2) pgwt_ground[of t] pgwt_is_empty_synth[of t] by auto hence x_dom: "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \ by auto moreover have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence ground_imgs: "ground (subst_range \)" using \(3) pgwt_ground pgwt_is_empty_synth by force ultimately have x_img: "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert t (subst_range \))" using ground_imgs x_dom t_ground by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \ t_ground by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert t (subst_range \)" proof show "subst_range \ \ insert t (subst_range \)" proof fix t' assume "t' \ subst_range \" then obtain y where "y \ subst_domain \" "t' = \ y" by auto thus "t' \ insert t (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert t (subst_range \) \ subst_range \" proof fix t' assume t': "t' \ insert t (subst_range \)" hence "fv t' = {}" using ground_imgs x_img t_ground by auto hence "t' \ Var x" by auto show "t' \ subst_range \" proof (cases "t' = t") case False hence "t' \ subst_range \" using t' by auto then obtain y where "\ y \ subst_range \" "t' = \ y" by auto hence "y \ subst_domain \" "t' \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t' \ Var y\ \t' = \ y\ subst_imgI[of \ y] by auto qed (metis subst_imgI \ \t' \ Var x\) qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs t_ground by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" using \ \_dom True * t(3) by (metis Un_iff term.order_refl insertE)+ hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" proof - { fix s assume "s \ t" hence "s \ {t. {} \\<^sub>c t} - T" using t(2,3) by (metis Diff_eq_empty_iff Diff_iff Un_upper1 term.order_refl deduct_synth_subterm mem_Collect_eq) } thus ?thesis using \(3) \ \_img by auto qed moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ t(1) \(5) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(6) t(2) pgwt_is_empty_synth pgwt_wellformed wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by metis moreover have "\s\subst_range \. \u\subst_range \. (\v. v \ s \ v \ u) \ s = u" using \(4) \_img t(3) by (auto simp del: subst_range.simps) ultimately show ?case by blast qed private lemma wt_bij_finite_tatom_subst_exists_single: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - let ?U = "{c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` ?U) - T" using bij_finite_const_subst_exists'[OF assms(1,2) infinite_typed_consts'[of a]] by auto { fix x assume "x \ subst_domain \" hence "\ (Var x) = \ (\ x)" by auto } moreover { fix x assume "x \ subst_domain \" hence "\c \ ?U. \ x = Fun c [] \ arity c = 0" using \ by auto hence "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) const_type wf_trmI[of "[]"] by auto hence "\ (Var x) = \ (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) \(1) by force+ } ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wf_trm_subst_range_iff[of \] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using \ by auto qed lemma wt_bij_finite_tatom_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \a. \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subst_range Var \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (auto simp add: subst_domain_def) next case (insert x S) then obtain a where a: "\ (Var x) = TAtom a" by fastforce from insert obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by auto let ?S' = "{y \ S. \ (Var y) = TAtom a}" let ?T' = "T \ subst_range \" have *: "finite (insert x ?S')" using insert by simp have **: "finite ?T'" using insert.prems(1) insert.hyps(1) \(1) by simp have ***: "\y. y \ insert x ?S' \ \ (Var y) = TAtom a" using a by auto obtain \ where \: "subst_domain \ = insert x ?S'" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - ?T'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_tatom_subst_exists_single[OF * ** ***] const_type_inv[of _ "[]" a] by blast obtain \ where \: "\ \ \y. if x = y then \ y else \ y" by simp have x_dom: "x \ subst_domain \" "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \(1) \ by (auto simp add: subst_domain_def) moreover have ground_imgs: "ground (subst_range \)" "ground (subst_range \)" using pgwt_ground \(3) \(3) by auto ultimately have x_img: "\ x \ subst_range \" "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert (\ x) (subst_range \))" using ground_imgs x_dom by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \(1) \ by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert (\ x) (subst_range \)" proof show "subst_range \ \ insert (\ x) (subst_range \)" proof fix t assume "t \ subst_range \" then obtain y where "y \ subst_domain \" "t = \ y" by auto thus "t \ insert (\ x) (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert (\ x) (subst_range \) \ subst_range \" proof fix t assume t: "t \ insert (\ x) (subst_range \)" hence "fv t = {}" using ground_imgs x_img(2) by auto hence "t \ Var x" by auto show "t \ subst_range \" proof (cases "t = \ x") case True thus ?thesis using subst_imgI \ \t \ Var x\ by metis next case False hence "t \ subst_range \" using t by auto then obtain y where "\ y \ subst_range \" "t = \ y" by auto hence "y \ subst_domain \" "t \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t \ Var y\ \t = \ y\ subst_imgI[of \ y] by auto qed qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs x_img by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" using \ \_dom x_img(2) \(3) True by (metis (no_types) *(1) DiffE Un_upper2 insertE subsetCE, metis (no_types) *(2) DiffE Un_upper2 insertE subsetCE) hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subst_range \ \ (\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b - T" using \(3) \(3) \ by (auto simp add: subst_domain_def) moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \(4) \(4) \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(5) \(5) wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by presburger ultimately show ?case by blast qed theorem wt_sat_if_simple: assumes "simple S" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and \': "\X F. Inequality X F \ set S \ ineq_model \' X F" "ground (subst_range \')" "subst_domain \' = {x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X}" and tfr_stp_all: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - from \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ have "wf\<^sub>s\<^sub>t {} S" "subst_idem \" and S_\_disj: "\v \ vars\<^sub>s\<^sub>t S. \ v = Var v" using subst_idemI[of \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ obtain \::"('fun,'var) subst" where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast hence \_deduct: "\x M. M \\<^sub>c \ x" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using pgwt_deducible pgwt_wellformed by fastforce+ let ?P = "\\ X. subst_domain \ = set X \ ground (subst_range \)" let ?Sineqsvars = "{x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" let ?Strms = "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" have finite_vars: "finite ?Sineqsvars" "finite ?Strms" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ?Strms" using wf_trm_subtermeq assms(5) by fastforce+ define Q1 where "Q1 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" define Q1' where "Q1' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \x \ (fv t \ fv t') - set X. \a. \ (Var x) = TAtom a)" define Q2' where "Q2' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \f T. Fun f T \ subterms t \ subterms t' \ T = [] \ (\s \ set T. s \ Var ` set X))" have ex_P: "\X. \\. ?P \ X" using interpretation_subst_exists' by blast have tfr_ineq: "\X F. Inequality X F \ set S \ Q1 F X \ Q2 F X" using tfr_stp_all Q1_def Q2_def tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of S] by blast have S_fv_bvars_disj: "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by metis hence ineqs_vars_not_bound: "\X F x. Inequality X F \ set S \ x \ ?Sineqsvars \ x \ set X" using strand_fv_bvars_disjoint_unfold by blast have \_vars_S_bvars_disj: "(subst_domain \ \ range_vars \) \ set X = {}" when "Inequality X F \ set S" for F X using wf_constr_bvars_disj[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\] strand_fv_bvars_disjointD(1)[OF S_fv_bvars_disj that] by blast obtain \::"('fun,'var) subst" where \_fv_dom: "subst_domain \ = ?Sineqsvars" and \_subterm_inj: "subterm_inj_on \ (subst_domain \)" and \_fresh_pub_img: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - ?Strms" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_subst_exists[OF finite_vars] subst_inj_on_is_bij_betw subterm_inj_on_alt_def' - by moura + by atomize_elim auto have \_bij_dom_img: "bij_betw \ (subst_domain \) (subst_range \)" by (metis \_subterm_inj subst_inj_on_is_bij_betw subterm_inj_on_alt_def) have "finite (subst_domain \)" by(metis \_fv_dom finite_vars(1)) hence \_finite_img: "finite (subst_range \)" using \_bij_dom_img bij_betw_finite by blast have \_img_subterms: "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" by (metis \_subterm_inj subterm_inj_on_alt_def') have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence "subst_range \ \ public_ground_wf_terms - ?Strms" and \_pgwt_img: "subst_range \ \ public_ground_wf_terms" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ public_ground_wf_terms" using \_fresh_pub_img pgwt_is_empty_synth by blast+ have \_img_ground: "ground (subst_range \)" using \_pgwt_img pgwt_ground by auto hence \_inj: "inj \" using \_bij_dom_img subst_inj_is_bij_betw_dom_img_if_ground_img by auto have \_ineqs_fv_dom: "\X F. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using \_fv_dom by fastforce have \_dom_bvars_disj: "\X F. Inequality X F \ set S \ subst_domain \ \ set X = {}" using ineqs_vars_not_bound \_fv_dom by fastforce have \'1: "\X F \. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \'" using \'(3) ineqs_vars_not_bound by fastforce have \'2: "\X F. Inequality X F \ set S \ subst_domain \' \ set X = {}" using \'(3) ineqs_vars_not_bound by blast have doms_eq: "subst_domain \' = subst_domain \" using \'(3) \_fv_dom by simp have \_ineqs_neq: "ineq_model \ X F" when "Inequality X F \ set S" for X F proof - obtain a::"'fun" where a: "a \ \(funs_term ` subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \))" using exists_fun_notin_funs_terms[OF subterms_union_finite[OF \_finite_img]] - by moura + by atomize_elim auto hence a': "\T. Fun a T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" "\S. Fun a [] \ set (Fun a []#S)" "Fun a [] \ Var ` set X" by (meson a UN_I term.set_intros(1), auto) define t where "t \ Fun a (Fun a []#map fst F)" define t' where "t' \ Fun a (Fun a []#map snd F)" note F_in = that have t_fv: "fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" unfolding t_def t'_def by force have t_subterms: "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ {t, t', Fun a []}" unfolding t_def t'_def by force have "t \ \ \ \ \ t' \ \ \ \" when "?P \ X" for \ proof - have tfr_assms: "Q1 F X \ Q2 F X" using tfr_ineq F_in by metis have "Q1 F X \ \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \c. \ x = Fun c []" proof fix x assume "Q1 F X" and x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" - then obtain a where "\ (Var x) = TAtom a" unfolding Q1_def by moura + then obtain a where "\ (Var x) = TAtom a" unfolding Q1_def by atomize_elim auto hence a: "\ (\ x) = TAtom a" using \_wt unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp have "x \ subst_domain \" using \_ineqs_fv_dom x F_in by auto then obtain f T where fT: "\ x = Fun f T" by (meson \_img_ground ground_img_obtain_fun) hence "T = []" using \_wf_trm a TAtom_term_cases by fastforce thus "\c. \ x = Fun c []" using fT by metis qed hence 1: "Q1 F X \ \x \ (fv t \ fv t') - set X. \c. \ x = Fun c []" using t_fv by auto have 2: "\Q1 F X \ Q2 F X" by (metis tfr_assms) have 3: "subst_domain \ \ set X = {}" using \_dom_bvars_disj F_in by auto have 4: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms t \ subterms t') = {}" proof - define M1 where "M1 \ {t, t', Fun a []}" define M2 where "M2 \ ?Strms" have "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M2" using F_in unfolding M2_def by force moreover have "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M1" using t_subterms unfolding M1_def by blast ultimately have *: "subterms t \ subterms t' \ M2 \ M1" by auto have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M1 = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M2 = {}" using a' \_fresh_pub_img unfolding t_def t'_def M1_def M2_def by blast+ thus ?thesis using * by blast qed have 5: "(fv t \ fv t') - subst_domain \ \ set X" using \_ineqs_fv_dom[OF F_in] t_fv by auto have 6: "\\. ?P \ X \ t \ \ \ \' \ t' \ \ \ \'" by (metis t_def t'_def \'(1) F_in ineq_model_singleE ineq_model_single_iff) have 7: "fv t \ fv t' - set X \ subst_domain \'" using \'1 F_in t_fv by force have 8: "subst_domain \' \ set X = {}" using \'2 F_in by auto have 9: "Q1' t t' X" when "Q1 F X" using that t_fv unfolding Q1_def Q1'_def t_def t'_def by blast have 10: "Q2' t t' X" when "Q2 F X" unfolding Q2'_def proof (intro allI impI) fix f T assume "Fun f T \ subterms t \ subterms t'" moreover { assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" hence "T = [] \ (\s\set T. s \ Var ` set X)" by (metis Q2_def that) } moreover { assume "Fun f T = t" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t_def using a'(2,3) by simp } moreover { assume "Fun f T = t'" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t'_def using a'(2,3) by simp } moreover { assume "Fun f T = Fun a []" hence "T = [] \ (\s\set T. s \ Var ` set X)" by simp } ultimately show "T = [] \ (\s\set T. s \ Var ` set X)" using t_subterms by blast qed note 11 = \_subterm_inj \_img_ground 3 4 5 note 12 = 6 7 8 \'(2) doms_eq show "t \ \ \ \ \ t' \ \ \ \" using 1 2 9 10 that sat_ineq_subterm_inj_subst[OF 11 _ 12] unfolding Q1'_def Q2'_def by metis qed thus ?thesis by (metis t_def t'_def ineq_model_singleI ineq_model_single_iff) qed have \_ineqs_fv_dom': "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom[OF that(1)] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv (t \ \) \ fv (t' \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (simp_all add: subst_apply_pairs_def) moreover have "fv (t \ \) = fv t - subst_domain \" "fv (t' \ \) = fv t' - subst_domain \" using g that(2) by (simp_all add: subst_fv_unfold_ground_img range_vars_alt_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" using Cons by auto ultimately show ?case using Cons.prems that(2) by auto qed (simp add: subst_apply_pairs_def) have \_ineqs_ground: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = {}" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom'[OF that] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv (t \ \) \ subst_domain \" "fv (t' \ \) \ subst_domain \" using Cons.prems by (auto simp add: subst_apply_pairs_def) hence "fv (t \ \ \ \) = {}" "fv (t' \ \ \ \) = {}" using subst_fv_dom_ground_if_ground_img[OF _ \_img_ground] by metis+ thus ?case using g Cons by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) from \_pgwt_img \_ineqs_neq have \_deduct: "M \\<^sub>c \ x" when "x \ subst_domain \" for x M using that pgwt_deducible by fastforce { fix M::"('fun,'var) terms" have "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" using \wf\<^sub>s\<^sub>t {} S\ \simple S\ S_\_disj \_ineqs_neq \_ineqs_fv_dom' \_vars_S_bvars_disj proof (induction S arbitrary: M rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) hence S_sat: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" and "\ v = Var v" by auto hence *: "\M. M \\<^sub>c Var v \ (\ \\<^sub>s \ \\<^sub>s \)" using \_deduct \_deduct by (metis ideduct_synth_subst_apply eval_term.simps(1) subst_subst_compose trm_subst_ident') define M' where "M' \ M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have "\t \ set [Var v]. M' \\<^sub>c t \ (\ \\<^sub>s \ \\<^sub>s \)" using *[of M'] by simp thus ?case using strand_sem_append(1)[OF S_sat, of "[Send1 (Var v)]", unfolded M'_def[symmetric]] strand_sem_c.simps(1)[of M'] strand_sem_c.simps(2)[of M' "[Var v]" "[]"] by presburger next case (ConsIneq X F S) have dom_disj: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {}" using ConsIneq.prems(1) subst_dom_vars_in_subst by force hence *: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by blast have **: "ineq_model \ X F" by (meson ConsIneq.prems(2) in_set_conv_decomp) have "\x. x \ vars\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t (S@[Inequality X F])" "\x. x \ set S \ x \ set (S@[Inequality X F])" by auto hence IH: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" by (metis ConsIneq.IH ConsIneq.prems(1,2,3,4)) have "ineq_model (\ \\<^sub>s \) X F" proof - have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "?P \ X" for \ using ConsIneq.prems(3)[OF _ that] by simp hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset ex_P by (metis Diff_subset_conv Un_commute) thus ?thesis by (metis ineq_model_ground_subst[OF _ \_img_ground **]) qed hence "ineq_model (\ \\<^sub>s \ \\<^sub>s \) X F" using * ineq_model_subst' subst_compose_assoc ConsIneq.prems(4) by (metis UnCI list.set_intros(1) set_append) thus ?case using IH by (auto simp add: ineq_model_def) qed auto } moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" by (metis wt_subst_compose \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, metis assms(4) \_wf_trm \_wf_trm wf_trm_subst subst_img_comp_subset') ultimately show ?thesis using interpretation_comp(1)[OF \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, of "\ \\<^sub>s \"] subst_idem_support[OF \subst_idem \\, of "\ \\<^sub>s \"] subst_compose_assoc unfolding constr_sem_c_def by metis qed end subsubsection \Theorem: Type-flaw resistant constraints are well-typed satisfiable (composition-only)\ text \ There exists well-typed models of satisfiable type-flaw resistant constraints in the semantics where the intruder is limited to composition only (i.e., he cannot perform decomposition/analysis of deducible messages). \ theorem wt_attack_if_tfr_attack: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ \\<^sub>c \S, \\" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" obtains \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "\\<^sub>\ \\<^sub>c \S, \\" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have tfr: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (LI_preproc S))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (LI_preproc S))" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using assms(5,6) LI_preproc_preserves_tfr unfolding tfr\<^sub>s\<^sub>t_def by (metis, metis LI_preproc_trms_eq, metis) have wf_constr: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (LI_preproc S) \" by (metis LI_preproc_preserves_wellformedness assms(3)) obtain S' \' where *: "simple S'" "(LI_preproc S,\) \\<^sup>* (S',\')" "\{}; S'\\<^sub>c \" using LI_completeness[OF assms(3,2)] unfolding constr_sem_c_def by (meson term.order_refl) have **: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S' \'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using LI_preserves_welltypedness[OF *(2) wf_constr assms(4,7) tfr] LI_preserves_wellformedness[OF *(2) wf_constr] LI_preserves_tfr[OF *(2) wf_constr assms(4,7) tfr] by metis+ define A where "A \ {x \ vars\<^sub>s\<^sub>t S'. \X F. Inequality X F \ set S' \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" define B where "B \ UNIV - A" let ?\ = "rm_vars B \" have gr\: "ground (subst_range \)" "ground (subst_range ?\)" using assms(1) rm_vars_img_subset[of B \] by (auto simp add: subst_domain_def) { fix X F assume "Inequality X F \ set S'" hence *: "ineq_model \ X F" using strand_sem_c_imp_ineq_model[OF *(3)] by (auto simp del: subst_range.simps) hence "ineq_model ?\ X F" proof - { fix \ assume 1: "subst_domain \ = set X" "ground (subst_range \)" and 2: "list_ex (\f. fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \) F" have "list_ex (\f. fst f \ \ \\<^sub>s rm_vars B \ \ snd f \ \ \\<^sub>s rm_vars B \) F" using 2 proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) thus ?case using Cons Unifier_ground_rm_vars[OF gr\(1), of "t \ \" B "t' \ \"] by auto qed simp } thus ?thesis using * unfolding ineq_model_def list_ex_iff case_prod_unfold by simp qed } moreover have "subst_domain \ = UNIV" using assms(1) by metis hence "subst_domain ?\ = A" using rm_vars_dom[of B \] B_def by blast ultimately obtain \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \S', \'\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_sat_if_simple[OF *(1) **(1,2,5,4) _ gr\(2) _ **(3)] A_def by (auto simp del: subst_range.simps) thus ?thesis using that LI_soundness[OF assms(3) *(2)] by metis qed text \ Contra-positive version: if a type-flaw resistant constraint does not have a well-typed model then it is unsatisfiable \ corollary secure_if_wt_secure: assumes "\(\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\)" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\(\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\))" using wt_attack_if_tfr_attack[OF _ _ assms(2,3,4,5,6)] assms(1) by metis end subsection \Lifting the Composition-Only Typing Result to the Full Intruder Model\ context typing_result begin subsubsection \Analysis Invariance\ definition (in typed_model) Ana_invar_subst where "Ana_invar_subst \ \ (\f T K M \. Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t \) \ Ana (Fun f T) = (K, M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))" lemma (in typed_model) Ana_invar_subst_subset: assumes "Ana_invar_subst M" "N \ M" shows "Ana_invar_subst N" using assms unfolding Ana_invar_subst_def by blast lemma (in typed_model) Ana_invar_substD: assumes "Ana_invar_subst \" and "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t \" "Ana (Fun f T) = (K, M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using assms Ana_invar_subst_def by blast end subsubsection \Preliminary Definitions\ text \Strands extended with "decomposition steps"\ datatype (funs\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'b) extstrand_step = Step "('a,'b) strand_step" | Decomp "('a,'b) term" context typing_result begin context begin private fun trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p where "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Step x) = trms\<^sub>s\<^sub>t\<^sub>p x" | "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Decomp t) = {t}" private abbreviation trms\<^sub>e\<^sub>s\<^sub>t where "trms\<^sub>e\<^sub>s\<^sub>t S \ \(trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p ` set S)" private type_synonym ('a,'b) extstrand = "('a,'b) extstrand_step list" private type_synonym ('a,'b) extstrands = "('a,'b) extstrand set" private definition decomp::"('fun,'var) term \ ('fun,'var) strand" where "decomp t \ (case (Ana t) of (K,T) \ [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\T\\<^sub>s\<^sub>t])" private fun to_st where "to_st [] = []" | "to_st (Step x#S) = x#(to_st S)" | "to_st (Decomp t#S) = (decomp t)@(to_st S)" private fun to_est where "to_est [] = []" | "to_est (x#S) = Step x#to_est S" private abbreviation "ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>t (to_st A)" private abbreviation "wf\<^sub>e\<^sub>s\<^sub>t V A \ wf\<^sub>s\<^sub>t V (to_st A)" private abbreviation "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t (to_st A)" private abbreviation "vars\<^sub>e\<^sub>s\<^sub>t A \ vars\<^sub>s\<^sub>t (to_st A)" private abbreviation "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A \ wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t (to_st A)" private abbreviation "funs\<^sub>e\<^sub>s\<^sub>t A \ funs\<^sub>s\<^sub>t (to_st A)" private definition wf\<^sub>s\<^sub>t\<^sub>s'::"('fun,'var) strands \ ('fun,'var) extstrand \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}) \ (\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}) \ (\S \ \. fv\<^sub>s\<^sub>t (to_st \) \ bvars\<^sub>s\<^sub>t S = {})" private definition wf\<^sub>s\<^sub>t\<^sub>s::"('fun,'var) strands \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s \ \ (\S \ \. wf\<^sub>s\<^sub>t {} (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {})" private inductive well_analyzed::"('fun,'var) extstrand \ bool" where Nil[simp]: "well_analyzed []" | Step: "well_analyzed A \ well_analyzed (A@[Step x])" | Decomp: "\well_analyzed A; t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)\ \ well_analyzed (A@[Decomp t])" private fun subst_apply_extstrandstep (infix "\\<^sub>e\<^sub>s\<^sub>t\<^sub>p" 51) where "subst_apply_extstrandstep (Step x) \ = Step (x \\<^sub>s\<^sub>t\<^sub>p \)" | "subst_apply_extstrandstep (Decomp t) \ = Decomp (t \ \)" private lemma subst_apply_extstrandstep'_simps[simp]: "(Step (send\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t)" "(Step (receive\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t)" "(Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\a: (t \ \) \ (t' \ \)\\<^sub>s\<^sub>t)" "(Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\\<^sub>s\<^sub>t)" by simp_all private lemma vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p_subst_apply_simps[simp]: "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (send\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (receive\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = set X \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" by auto private definition subst_apply_extstrand (infix "\\<^sub>e\<^sub>s\<^sub>t" 51) where "S \\<^sub>e\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) S" private abbreviation update\<^sub>s\<^sub>t::"('fun,'var) strands \ ('fun,'var) strand \ ('fun,'var) strands" where "update\<^sub>s\<^sub>t \ S \ (case S of Nil \ \ - {S} | Cons _ S' \ insert S' (\ - {S}))" private inductive_set decomps\<^sub>e\<^sub>s\<^sub>t:: "('fun,'var) terms \ ('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrands" (* \: intruder knowledge \: additional messages *) for \ and \ and \ where Nil: "[] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" | Decomp: "\\ \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \; Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ \ \); Ana (Fun f T) = (K,M); M \ []; (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \; \k. k \ set K \ (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \\ \ \@[Decomp (Fun f T)] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" private fun decomp_rm\<^sub>e\<^sub>s\<^sub>t::"('fun,'var) extstrand \ ('fun,'var) extstrand" where "decomp_rm\<^sub>e\<^sub>s\<^sub>t [] = []" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Decomp t#S) = decomp_rm\<^sub>e\<^sub>s\<^sub>t S" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Step x#S) = Step x#(decomp_rm\<^sub>e\<^sub>s\<^sub>t S)" private inductive sem\<^sub>e\<^sub>s\<^sub>t_d::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ \t \ set ts. (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (send\ts\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Decomp t])" private inductive sem\<^sub>e\<^sub>s\<^sub>t_c::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ \t \ set ts. (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (send\ts\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Decomp t])" subsubsection \Preliminary Lemmata\ private lemma wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>s\<^sub>t\<^sub>s \ = wf\<^sub>s\<^sub>t\<^sub>s' \ []" by (simp add: wf\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>s\<^sub>t\<^sub>s'_def) private lemma decomp_ik: assumes "Ana t = (K,M)" shows "ik\<^sub>s\<^sub>t (decomp t) = set M" using ik_rcv_map ik_rcv_map' by (auto simp add: decomp_def inv_def assms) private lemma decomp_assignment_rhs_empty: assumes "Ana t = (K,M)" shows "assignment_rhs\<^sub>s\<^sub>t (decomp t) = {}" by (auto simp add: decomp_def inv_def assms) private lemma decomp_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (decomp t)" by (auto simp add: decomp_def list_all_def) private lemma trms\<^sub>e\<^sub>s\<^sub>t_ikI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case by (cases x) auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case proof (cases x) case (Equality ac t t') thus ?thesis using 2 by (cases ac) auto qed auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] decomp_assignment_rhs_empty[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof - obtain t' where "t' \ ik\<^sub>e\<^sub>s\<^sub>t A" "t \ t'" using trms\<^sub>e\<^sub>s\<^sub>t_ikI assms by auto thus ?thesis by (meson contra_subsetD in_subterms_subset_Union trms\<^sub>e\<^sub>s\<^sub>t_ikI) qed private lemma trms\<^sub>e\<^sub>s\<^sub>tD: assumes "t \ trms\<^sub>e\<^sub>s\<^sub>t A" shows "t \ trms\<^sub>s\<^sub>t (to_st A)" using assms proof (induction A) case (Cons a A) obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "t \ trms\<^sub>s\<^sub>t (decomp t)" unfolding decomp_def by force thus ?case using Cons.IH Cons.prems by (cases a) auto qed simp private lemma subst_apply_extstrand_nil[simp]: "[] \\<^sub>e\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_extstrand_def) private lemma subst_apply_extstrand_singleton[simp]: "[Step (receive\ts\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))]" "[Step (send\ts\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Send (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))]" "[Step (\a: t \ t'\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Equality a (t \ \) (t' \ \))]" "[Decomp t] \\<^sub>e\<^sub>s\<^sub>t \ = [Decomp (t \ \)]" unfolding subst_apply_extstrand_def by auto private lemma extstrand_subst_hom: "(S@S') \\<^sub>e\<^sub>s\<^sub>t \ = (S \\<^sub>e\<^sub>s\<^sub>t \)@(S' \\<^sub>e\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>e\<^sub>s\<^sub>t \ = (x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>e\<^sub>s\<^sub>t \)" unfolding subst_apply_extstrand_def by auto private lemma decomp_vars: "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" proof - obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "decomp t = [send\[t]\\<^sub>s\<^sub>t,Send K,Receive M]" unfolding decomp_def by simp moreover have "\(set (map fv K)) = fv\<^sub>s\<^sub>e\<^sub>t (set K)" "\(set (map fv M)) = fv\<^sub>s\<^sub>e\<^sub>t (set M)" by auto moreover have "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" using Ana_subterm[OF Ana(1)] Ana_keys_fv[OF Ana(1)] by (simp_all add: UN_least psubsetD subtermeq_vars_subset) ultimately show "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" by auto qed private lemma bvars\<^sub>e\<^sub>s\<^sub>t_cons: "bvars\<^sub>e\<^sub>s\<^sub>t (x#X) = bvars\<^sub>e\<^sub>s\<^sub>t [x] \ bvars\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma bvars\<^sub>e\<^sub>s\<^sub>t_append: "bvars\<^sub>e\<^sub>s\<^sub>t (A@B) = bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by force qed simp private lemma fv\<^sub>e\<^sub>s\<^sub>t_cons: "fv\<^sub>e\<^sub>s\<^sub>t (x#X) = fv\<^sub>e\<^sub>s\<^sub>t [x] \ fv\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma fv\<^sub>e\<^sub>s\<^sub>t_append: "fv\<^sub>e\<^sub>s\<^sub>t (A@B) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using fv\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] fv\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by auto qed simp private lemma bvars_decomp: "bvars\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = bvars\<^sub>e\<^sub>s\<^sub>t A" "bvars\<^sub>e\<^sub>s\<^sub>t (Decomp t#A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars\<^sub>e\<^sub>s\<^sub>t_append decomp_vars(3) by fastforce+ private lemma bvars_decomp_rm: "bvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars_decomp by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) simp_all+ private lemma fv_decomp_rm: "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma ik_assignment_rhs_decomp_fv: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" shows "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A" proof - have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv t" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_vars by simp moreover have "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by force moreover have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using fv_subset_subterms[OF assms(1)] by simp ultimately show ?thesis by blast qed private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto+ private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A = wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" by simp private lemma decomp_set_unfold: assumes "Ana t = (K, M)" shows "set (decomp t) = {send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t}" using assms unfolding decomp_def by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_finite: "finite (ik\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_ik\<^sub>s\<^sub>t) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite: "finite (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_assignment_rhs\<^sub>s\<^sub>t) private lemma to_est_append: "to_est (A@B) = to_est A@to_est B" by (induct A rule: to_est.induct) auto private lemma to_st_to_est_inv: "to_st (to_est A) = A" by (induct A rule: to_est.induct) auto private lemma to_st_append: "to_st (A@B) = (to_st A)@(to_st B)" by (induct A rule: to_st.induct) auto private lemma to_st_cons: "to_st (a#B) = (to_st [a])@(to_st B)" using to_st_append[of "[a]" B] by simp private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (x#S) = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t [x] \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S'" using to_st_cons[of x S] to_st_append[of S S'] by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_append: "ik\<^sub>e\<^sub>s\<^sub>t (A@B) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t B" by (metis ik_append to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t B" by (metis assignment_rhs_append to_st_append) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons: "ik\<^sub>e\<^sub>s\<^sub>t (a#A) = ik\<^sub>e\<^sub>s\<^sub>t [a] \ ik\<^sub>e\<^sub>s\<^sub>t A" by (metis ik_append to_st_cons) private lemma ik\<^sub>e\<^sub>s\<^sub>t_append_subst: "ik\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), simp add: image_Un to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), use assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append in blast) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons_subst: "ik\<^sub>e\<^sub>s\<^sub>t (a#A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t ([a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \]) \ ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_cons extstrand_subst_hom(2), metis image_Un ik\<^sub>e\<^sub>s\<^sub>t_cons) private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_append: "decomp_rm\<^sub>e\<^sub>s\<^sub>t (S@S') = (decomp_rm\<^sub>e\<^sub>s\<^sub>t S)@(decomp_rm\<^sub>e\<^sub>s\<^sub>t S')" by (induct S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_single[simp]: "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (send\ts\\<^sub>s\<^sub>t)] = [Step (send\ts\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (receive\ts\\<^sub>s\<^sub>t)] = [Step (receive\ts\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Decomp t] = []" by auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t S) \ ik\<^sub>e\<^sub>s\<^sub>t S" proof (induction S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) case (3 x S) thus ?case by (cases x) auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ ik\<^sub>e\<^sub>s\<^sub>t D \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M') have "ik\<^sub>s\<^sub>t (decomp (Fun f T)) \ subterms (Fun f T)" "ik\<^sub>s\<^sub>t (decomp (Fun f T)) = ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" using decomp_ik[OF Decomp.hyps(3)] Ana_subterm[OF Decomp.hyps(3)] by auto hence "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f T)]) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[OF Decomp.hyps(2)] by blast thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] using Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ decomp_rm\<^sub>e\<^sub>s\<^sub>t D = []" by (induct D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) (auto simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "A \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" "B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" shows "A@B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" using assms(2) proof (induction B rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil show ?case using assms(1) by simp next case (Decomp B f X K T) hence "S \ ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \ \ S \ ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF Decomp.IH(1) Decomp.hyps(2,3,4)] ideduct_synth_mono[OF Decomp.hyps(5)] ideduct_synth_mono[OF Decomp.hyps(6)] by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_subterms: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A') \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using assms proof (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f X K T) hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" by auto hence "subterms\<^sub>s\<^sub>e\<^sub>t (set X) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[of "Fun f X" "M \ N"] params_subterms_Union[of X f] by blast moreover have "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)]) = set T" using Decomp.hyps(3) decomp_ik by simp hence "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)])) \ subterms\<^sub>s\<^sub>e\<^sub>t (set X)" using Ana_fun_subterm[OF Decomp.hyps(3)] by auto ultimately show ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f X)]"] Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A' = {}" using assms by (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) (simp_all add: decomp_assignment_rhs_empty assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append: assumes "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D = (\m \ M. ik\<^sub>e\<^sub>s\<^sub>t m)" using assms proof (induction M rule: finite_induct) case empty moreover have "[] \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>s\<^sub>t (to_st []) = {}" using decomps\<^sub>e\<^sub>s\<^sub>t.Nil by auto ultimately show ?case by blast next case (insert m M) - then obtain D where "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>e\<^sub>s\<^sub>t D = (\m\M. ik\<^sub>s\<^sub>t (to_st m))" by moura + then obtain D where "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>e\<^sub>s\<^sub>t D = (\m\M. ik\<^sub>s\<^sub>t (to_st m))" by atomize_elim auto moreover have "m \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" using insert.prems(1) by blast ultimately show ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[of D A N \ m] ik\<^sub>e\<^sub>s\<^sub>t_append[of D m] by blast qed private lemma decomp_snd_exists[simp]: "\D. decomp t = send\[t]\\<^sub>s\<^sub>t#D" by (metis (mono_tags, lifting) decomp_def prod.case surj_pair) private lemma decomp_nonnil[simp]: "decomp t \ []" using decomp_snd_exists[of t] by fastforce private lemma to_st_nil_inv[dest]: "to_st A = [] \ A = []" by (induct A rule: to_st.induct) auto private lemma well_analyzedD: assumes "well_analyzed A" "Decomp t \ set A" shows "\f T. t = Fun f T" using assms proof (induction A rule: well_analyzed.induct) case (Decomp A t') hence "\f T. t' = Fun f T" by (cases t') auto moreover have "Decomp t \ set A \ t = t'" using Decomp by auto ultimately show ?case using Decomp.IH by auto qed auto private lemma well_analyzed_inv: assumes "well_analyzed (A@[Decomp t])" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)" using assms well_analyzed.cases[of "A@[Decomp t]"] by fastforce private lemma well_analyzed_split_left_single: "well_analyzed (A@[a]) \ well_analyzed A" by (induction "A@[a]" rule: well_analyzed.induct) auto private lemma well_analyzed_split_left: "well_analyzed (A@B) \ well_analyzed A" proof (induction B rule: List.rev_induct) case (snoc b B) thus ?case using well_analyzed_split_left_single[of "A@B" b] by simp qed simp private lemma well_analyzed_append: assumes "well_analyzed A" "well_analyzed B" shows "well_analyzed (A@B)" using assms(2,1) proof (induction B rule: well_analyzed.induct) case (Step B x) show ?case using well_analyzed.Step[OF Step.IH[OF Step.prems]] by simp next case (Decomp B t) thus ?case using well_analyzed.Decomp[OF Decomp.IH[OF Decomp.prems]] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto qed simp_all private lemma well_analyzed_singleton: "well_analyzed [Step (send\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" "\well_analyzed [Decomp t]" proof - show "well_analyzed [Step (send\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" using well_analyzed.Step[OF well_analyzed.Nil] by simp_all show "\well_analyzed [Decomp t]" using well_analyzed.cases[of "[Decomp t]"] by auto qed private lemma well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv: "well_analyzed A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = fv\<^sub>e\<^sub>s\<^sub>t A" proof assume "well_analyzed A" thus "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: well_analyzed.induct) case Decomp thus ?case using ik_assignment_rhs_decomp_fv decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Step A x) have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Step x]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t\<^sub>p x" "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t (A@[Step x])) = fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>s\<^sub>t\<^sub>p x" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using Step by auto qed simp qed (rule fv_decomp_rm) private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_split_left: assumes "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (\@\')" shows "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" using assms sem\<^sub>e\<^sub>s\<^sub>t_d.cases by (induction \' rule: List.rev_induct) fastforce+ private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>d' \" proof show "\M\<^sub>0; to_st \\\<^sub>d' \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>d' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Receive[OF IH] Step by auto next case (Equality a t t') thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF IH] * Step by auto next case (Inequality X F) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF IH] * Step by auto qed next case (Decomp t) - obtain K M where Ana: "Ana t = (K,M)" by moura + obtain K M where Ana: "Ana t = (K,M)" by atomize_elim auto have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = [send\[t]\\<^sub>s\<^sub>t,Send K,Receive M]" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; [Send K]\\<^sub>d' \" using * by auto hence "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using * strand_sem_Send_split(2) strand_sem_d.simps(2) unfolding strand_sem_eq_defs(2) list_all_iff by meson thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_d.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>d' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_d.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>d' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t]\\<^sub>d' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send1 k]\\<^sub>d' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send K]\\<^sub>d' \" using strand_sem_Send_map(4)[of _ "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_Send_map(6) unfolding strand_sem_eq_defs(2) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Receive M]\\<^sub>d' \" by (metis strand_sem_Receive_map(6) strand_sem_eq_defs(2)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]\\<^sub>d' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>d' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>c' \" proof show "\M\<^sub>0; to_st \\\<^sub>c' \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>c' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF IH] Step by auto next case (Equality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF IH] * Step by auto next case (Inequality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF IH] * Step by auto qed next case (Decomp t) - obtain K M where Ana: "Ana t = (K,M)" by moura + obtain K M where Ana: "Ana t = (K,M)" by atomize_elim auto have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; [send\K\\<^sub>s\<^sub>t]\\<^sub>c' \" using * by auto hence "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" when k: "k \ set K" for k using * strand_sem_Send_split(5)[OF _ k] strand_sem_Send_map(5) unfolding strand_sem_eq_defs(1) by auto thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>c' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (auto simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>c' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t]\\<^sub>c' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send1 k]\\<^sub>c' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send K]\\<^sub>c' \" using strand_sem_Send_map(3)[of K, of "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_Send_map(5) unfolding strand_sem_eq_defs(1) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Receive M]\\<^sub>c' \" by (metis strand_sem_Receive_map(5) strand_sem_eq_defs(1)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]\\<^sub>c' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>c' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms proof (induction M\<^sub>0 \ A arbitrary: t rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Receive M\<^sub>0 \ A t') hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence IH: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using Receive.IH by auto show ?case using ideduct_mono[OF IH] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by (metis Un_subset_iff Un_upper1 Un_upper2 image_mono) next case (Equality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Inequality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Decompose M\<^sub>0 \ A t' K M t) have *: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ \" using Decompose.hyps(2) proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp { fix k assume "k \ set K" hence "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose.hyps by auto hence "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp } hence **: "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k" by auto show ?case proof (cases "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \") case True thus ?thesis using Decompose.IH Decompose.prems(2) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case False hence "t \ ik\<^sub>s\<^sub>t (decomp t') \\<^sub>s\<^sub>e\<^sub>t \" using Decompose.prems(1) ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence ***: "t \ set (M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Decompose.hyps(3) decomp_ik by auto hence "M \ []" by auto hence ****: "Ana (t' \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(3)] by auto have "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" by (rule intruder_deduct.Decompose[OF * **** ** ***]) thus ?thesis using ideduct_mono decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms(2) proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \" by auto moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using ideduct_mono[OF intruder_deduct.Axiom] by auto } moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux[OF assms(1)] by auto } ultimately show ?case by auto qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction M\<^sub>0 \ A rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF Send.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct unfolding list_all_iff by auto next case (Receive t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Receive by auto next case (Equality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF Equality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case (Inequality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF Inequality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case Decompose thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed auto private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \" shows "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A@D)" using assms(2,1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) hence *: "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A @ D)" "ik\<^sub>e\<^sub>s\<^sub>t (A@D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t (A @ D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF *(1,2) Decomp.hyps(3) *(3)] by simp qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "wf\<^sub>e\<^sub>s\<^sub>t V A" shows "wf\<^sub>e\<^sub>s\<^sub>t V (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) have "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomp_vars fv_subset_subterms[OF Decomp.hyps(2)] by fast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset[of "to_st A"] by blast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st (A@D)) \ V" using to_st_append[of A D] strand_vars_split(2)[of "to_st A" "to_st D"] by (metis le_supI1) thus ?case using wf_append_suffix[OF Decomp.IH[OF Decomp.prems], of "decomp (Fun f T)"] to_st_append[of "A@D" "[Decomp (Fun f T)]"] by auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" shows "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF Decomp.IH[OF Decomp.prems] _ Decomp.hyps(3)] Decomp.hyps(5,6) ideduct_synth_mono ik\<^sub>e\<^sub>s\<^sub>t_append by (metis (mono_tags, lifting) List.append_assoc image_Un sup_ge1) qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(M \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t)" obtains D' where "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" proof - have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" using assms(2) proof (induction t rule: intruder_deduct_induct) case (Compose X f) from Compose.IH have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c x" proof (induction X) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" - by moura + by atomize_elim force hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by (metis set_ConsD) qed (auto intro: decomps\<^sub>e\<^sub>s\<^sub>t.Nil) thus ?case using intruder_synth.ComposeC[OF Compose.hyps(1,2)] by metis next case (Decompose t K T t\<^sub>i) have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \k \ set K. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c k" using Decompose.IH proof (induction K) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" - using assms(1) by moura + using assms(1) by atomize_elim force hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by auto qed auto then obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c k" by metis obtain D'' where D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c t" by (metis Decompose.IH(1)) obtain f X where fX: "t = Fun f X" "t\<^sub>i \ set X" using Decompose.hyps(2,4) by (cases t) (auto dest: Ana_fun_subterm) from decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] D'(2) D''(2) have *: "D'@D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c k" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) hence **: "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ideduct_synth_subst by auto have "t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)" using Decompose.hyps(2,4) ik_rcv_map unfolding decomp_def by auto with *(3) fX(1) Decompose.hyps(2) show ?case proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence t_in_subterms: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF *(1)] subset_subterms_Union by auto have "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using ideduct_synth_subst[OF intruder_synth.AxiomC[OF AxiomC.hyps(1)]] by metis moreover have "T \ []" using decomp_ik[OF \Ana t = (K,T)\] \t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)\ by auto ultimately have "D'@D''@[Decomp (Fun f X)] \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" using AxiomC decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF *(1) _ _ _ _ **] subset_subterms_Union t_in_subterms by (simp add: subset_eq) moreover have "decomp t = to_st [Decomp (Fun f X)]" using AxiomC.prems(1,2) by auto ultimately show ?case by (metis AxiomC.prems(3) UnCI intruder_synth.AxiomC ik\<^sub>e\<^sub>s\<^sub>t_append to_st_append) qed (auto intro!: fX(2) *(1)) qed (fastforce intro: intruder_synth.AxiomC assms(1)) hence "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) - thus thesis using that[OF decomps\<^sub>e\<^sub>s\<^sub>t_append[OF assms(1)]] assms ik\<^sub>e\<^sub>s\<^sub>t_append by moura + thus thesis using that[OF decomps\<^sub>e\<^sub>s\<^sub>t_append[OF assms(1)]] assms ik\<^sub>e\<^sub>s\<^sub>t_append by atomize_elim auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" proof - let ?IK = "\M. \D \ M. ik\<^sub>e\<^sub>s\<^sub>t D" have "?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \) \ (\t \ A \ N. subterms t)" by (auto dest!: decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset) hence "finite (?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \))" using subterms_union_finite[OF assms(1)] subterms_union_finite[OF assms(2)] infinite_super by auto then obtain M where M: "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "?IK M = ?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \)" - using finite_subset_Union by moura + using finite_subset_Union by atomize_elim auto show ?thesis using decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append[OF M(1,2)] M(3) by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t" proof (rule ccontr) assume neg: "\(\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" - using decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist[OF assms] by moura + using decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist[OF assms] by atomize_elim force then obtain t where t: "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" using neg by (fastforce intro: ideduct_mono) obtain D' where D': "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" by (metis decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux t D(1)) hence "ik\<^sub>e\<^sub>s\<^sub>t D \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "ik\<^sub>e\<^sub>s\<^sub>t (D@D') \ ik\<^sub>e\<^sub>s\<^sub>t D" using D(2) D'(1) by auto ultimately show False by simp qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst: assumes "ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "wf\<^sub>e\<^sub>s\<^sub>t {} A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" and "well_analyzed A" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have ik_eq: "ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp with snoc have IH: "ik\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" using well_analyzed_split_left[OF snoc.prems(2)] by (auto simp add: to_st_append ik\<^sub>e\<^sub>s\<^sub>t_append_subst) have "ik\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_ik) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp moreover have assignment_rhs_eq: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp hence "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using snoc.IH well_analyzed_split_left[OF snoc.prems(2)] by simp hence IH: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(1), metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) have "assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_assignment_rhs_empty) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) Var" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist[OF ik\<^sub>e\<^sub>s\<^sub>t_finite assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite, of "A \\<^sub>e\<^sub>s\<^sub>t \" "A \\<^sub>e\<^sub>s\<^sub>t \"] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append assms(1) by force let ?P = "\D D'. \t. (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t" have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ?P D D'" using D(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil have "ik\<^sub>e\<^sub>s\<^sub>t [] = ik\<^sub>e\<^sub>s\<^sub>t [] \\<^sub>s\<^sub>e\<^sub>t \" by auto thus ?case by (metis decomps\<^sub>e\<^sub>s\<^sub>t.Nil) next case (Decomp D f T K M) obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "?P D D'" using Decomp.IH by auto hence IH: "\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c k" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c Fun f T" using Decomp.hyps(5,6) by auto have D'_ik: "ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "ik\<^sub>e\<^sub>s\<^sub>t D' \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF D'(1)] by (metis subst_all_mono, metis) show ?case using IH(2,1) Decomp.hyps(2,3,4) proof (induction "Fun f T" arbitrary: f T K M rule: intruder_synth_induct) case (AxiomC f T) then obtain s where s: "s \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D'" "Fun f T = s \ \" using AxiomC.prems by blast hence fT_s_in: "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using AxiomC D'_ik subset_subterms_Union[of "ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A"] subst_all_mono[OF subset_subterms_Union, of \] by (metis (no_types) Un_iff image_eqI subset_Un_eq, metis (no_types) Un_iff subset_Un_eq) - obtain Ks Ms where Ana_s: "Ana s = (Ks,Ms)" by moura + obtain Ks Ms where Ana_s: "Ana s = (Ks,Ms)" by atomize_elim auto have AD'_props: "wf\<^sub>e\<^sub>s\<^sub>t {} (A@D')" "\{}; to_st (A@D')\\<^sub>c \" using decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c[OF D'(1) assms(2)] decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf[OF D'(1) assms(3)] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st strand_sem_eq_defs(1) by auto show ?case proof (cases s) case (Var x) \ \In this case \\ x\ (is a subterm of something that) was derived from an "earlier intruder knowledge" because \A\ is well-formed and has \\\ as a model. So either the intruder composed \Fun f T\ himself (making \Decomp (Fun f T)\ unnecessary) or \Fun f T\ is an instance of something else in the intruder knowledge (in which case the "something" can be used in place of \Fun f T\)\ hence "Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')" "\ x = Fun f T" using s ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?thesis proof (cases "\m \ set M. ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c m") case True \ \All terms acquired by decomposing \Fun f T\ are already derivable. Hence there is no need to consider decomposition of \Fun f T\ at all.\ have *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K,M)\] ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') thus ?case proof assume "t' \ set M" moreover have "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using True by auto qed (metis D'(2) intruder_synth.AxiomC) qed auto } thus ?thesis using D'(1) * by metis next case False \ \Some term acquired by decomposition of \Fun f T\ cannot be derived in \\\<^sub>c\. \Fun f T\ must therefore be an instance of something else in the intruder knowledge, because of well-formedness.\ then obtain t\<^sub>i where t\<^sub>i: "t\<^sub>i \ set T" "\ik\<^sub>e\<^sub>s\<^sub>t (A@D') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using Ana_fun_subterm[OF \Ana (Fun f T) = (K,M)\] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) obtain S where fS: "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@D')) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@D'))" "\ x = Fun f S \ \" using strand_sem_wf_ik_or_assignment_rhs_fun_subterm[ OF AD'_props \Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')\ _ t\<^sub>i \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\] \\ x = Fun f T\ - by moura + by atomize_elim metis hence fS_in: "Fun f S \ \ \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using imageI[OF s(1), of "\x. x \ \"] Var ik\<^sub>e\<^sub>s\<^sub>t_append[of A D'] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A D'] decomps\<^sub>e\<^sub>s\<^sub>t_subterms[OF D'(1)] decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty[OF D'(1)] by auto - obtain KS MS where Ana_fS: "Ana (Fun f S) = (KS, MS)" by moura + obtain KS MS where Ana_fS: "Ana (Fun f S) = (KS, MS)" by atomize_elim auto hence "K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Ana_invar_substD[OF assms(5) fS_in(2)] s(2) fS(2) \s = Var x\ \Ana (Fun f T) = (K,M)\ by simp_all hence "MS \ []" using \M \ []\ by simp have "\k. k \ set KS \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using AxiomC.prems(1) \K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by (simp add: image_Un) hence D'': "D'@[Decomp (Fun f S)] \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) fS_in(2) Ana_fS \MS \ []\] AxiomC.prems(1) intruder_synth.AxiomC[OF fS_in(1)] by simp moreover { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" by (simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \" using decomp_ik \Ana (Fun f T) = (K,M)\ \Ana (Fun f S) = (KS,MS)\ \M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by simp thus ?case using ideduct_synth_mono[ OF intruder_synth.AxiomC[of t' "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \"], of "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \)"] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) next assume "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (metis D'(2) intruder_synth.AxiomC) hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (simp add: ideduct_synth_mono) thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D' "[Decomp (Fun f S)]"] image_Un[of "\x. x \ \" "ik\<^sub>e\<^sub>s\<^sub>t D'" "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)]"] by (simp add: sup_aci(2)) qed qed auto } ultimately show ?thesis using D'' by auto qed next case (Fun g S) \ \Hence \Decomp (Fun f T)\ can be substituted for \Decomp (Fun g S)\\ hence KM: "K = Ks \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = Ms \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "set K = set Ks \\<^sub>s\<^sub>e\<^sub>t \" "set M = set Ms \\<^sub>s\<^sub>e\<^sub>t \" using fT_s_in(2) \Ana (Fun f T) = (K,M)\ Ana_s s(2) Ana_invar_substD[OF assms(5), of g S] by auto hence Ms_nonempty: "Ms \ []" using \M \ []\ by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun g S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using AxiomC proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ set M" by (simp add: decomp_ik ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof (elim disjE) assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t D" hence *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using D'(2) by simp show ?case by (auto intro: ideduct_synth_mono[OF *] simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) next assume "t' \ set M" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun g S)] \\<^sub>s\<^sub>e\<^sub>t \" using KM(2) Fun decomp_ik[OF Ana_s] by auto thus ?case by (simp add: image_Un ik\<^sub>e\<^sub>s\<^sub>t_append) qed (simp add: ideduct_synth_mono[OF intruder_synth.AxiomC]) qed auto } thus ?thesis using s Fun Ana_s AxiomC.prems(1) KM(3) fT_s_in decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) _ _ Ms_nonempty, of g S Ks] by (metis AxiomC.hyps image_Un image_eqI intruder_synth.AxiomC) qed next case (ComposeC T f) have *: "\m. m \ set M \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c m" using Ana_fun_subterm[OF \Ana (Fun f T) = (K, M)\] ComposeC.hyps(3) by auto have **: "ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K, M)\] ik\<^sub>e\<^sub>s\<^sub>t_append by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (induct rule: intruder_synth_induct) (auto simp add: D'(2) * **) } thus ?case using D'(1) by auto qed qed thus ?thesis using D(2) assms(1) by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst_list: assumes "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "wf\<^sub>e\<^sub>s\<^sub>t {} A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" and "well_analyzed A" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. \t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" (is "\D \ ?A. ?B D ts") proof - note 0 = decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst[OF _ assms(2-6)] show ?thesis using assms(1) proof (induction ts) case (Cons t ts) have 1: "ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and 2: "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using Cons.prems by auto obtain D where D: "D \ ?A" "ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using 0[OF 1] by blast obtain D' where D': "D' \ ?A" "?B D' ts" using Cons.IH[OF 2] by auto have "ik\<^sub>e\<^sub>s\<^sub>t (A@D@D') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using ideduct_synth_mono[OF D(2)] ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)[of \ "A@D" D'] by fastforce hence "?B (D@D') (t#ts)" using D'(2) ideduct_synth_mono[of "ik\<^sub>e\<^sub>s\<^sub>t (A@D') \\<^sub>s\<^sub>e\<^sub>t \" _ "ik\<^sub>e\<^sub>s\<^sub>t (A@D@D') \\<^sub>s\<^sub>e\<^sub>t \"] ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)[of \] by auto thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D(1) D'(1)] by blast qed (fastforce intro: decomps\<^sub>e\<^sub>s\<^sub>t.Nil) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ []) \" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "send\ts\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)) (\@[Step (receive\ts\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "send\ts\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (receive\ts\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "receive\ts\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (receive\ts\\<^sub>s\<^sub>t#S)) (\@[Step (send\ts\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "receive\ts\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (send\ts\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)) (\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\a: t \ t'\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "a = Assign \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "a = Check \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (cases a) (metis wf_vars_mono, metis) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by (cases a) auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv t \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv t' \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\X\\\: F\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)) (\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\X\\\: F\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by metis have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t # S)" by auto ultimately have 5: "\S' \ \. (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ bvars\<^sub>s\<^sub>t S' = {}" "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = set X \ bvars\<^sub>e\<^sub>s\<^sub>t \" "\S \ \. fv\<^sub>s\<^sub>t S \ set X = {}" using to_st_append by (blast, force, force, force) have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5(3,4) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by blast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5(1,2) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq: assumes "x#S \ \" shows "\(trms\<^sub>s\<^sub>t ` update\<^sub>s\<^sub>t \ (x#S)) \ trms\<^sub>s\<^sub>t\<^sub>p x = \(trms\<^sub>s\<^sub>t ` \)" (is "?A = ?B") proof show "?B \ ?A" proof have "trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t (x#S)" by auto hence "\t'. t' \ ?B \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ t' \ ?A" by simp moreover { fix t' assume t': "t' \ ?B" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where S': "t' \ trms\<^sub>s\<^sub>t S'" "S' \ \" by auto hence "S' = x#S \ S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto moreover { assume "S' = x#S" hence "t' \ trms\<^sub>s\<^sub>t S" using S' t' by simp hence "t' \ ?A" by auto } ultimately have "t' \ ?A" using t' S' by auto } ultimately show "\t'. t' \ ?B \ t' \ ?A" by metis qed show "?A \ ?B" proof have "\t'. t' \ ?A \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t\<^sub>p x \ ?B" using assms by force+ moreover { fix t' assume t': "t' \ ?A" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where "t' \ trms\<^sub>s\<^sub>t S'" "S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto hence "S' = S \ S' \ \" by auto moreover have "trms\<^sub>s\<^sub>t S \ ?B" using assms trms\<^sub>s\<^sub>t_cons[of x S] by blast ultimately have "t' \ ?B" using t' by fastforce } ultimately show "\t'. t' \ ?A \ t' \ ?B" by blast qed qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd: assumes "send\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\ts\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ set ts" "\(trms\<^sub>s\<^sub>t ` \') \ set ts = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_commute Un_left_commute) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv: assumes "receive\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\ts\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ set ts" "\(trms\<^sub>s\<^sub>t ` \') \ set ts = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_commute Un_left_commute) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ {t,t'}" "\(trms\<^sub>s\<^sub>t ` \') \ {t,t'} = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "\(trms\<^sub>s\<^sub>t ` \') \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (simp add: Un_commute sup_left_commute) qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset: assumes "x#S \ \" shows "\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" (is ?A) "\(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" (is ?B) proof - { fix t assume "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S')" by auto have *: "ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t (x#S))" using ik_append[of "dual\<^sub>s\<^sub>t [x]" "dual\<^sub>s\<^sub>t S"] dual\<^sub>s\<^sub>t_append[of "[x]" S] by auto hence "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" proof (cases "S' = S") case True thus ?thesis using * assms S' by auto next case False thus ?thesis using S' by auto qed } moreover { fix t assume "t \ \(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ assignment_rhs\<^sub>s\<^sub>t S'" by auto have "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (x#S)" using assignment_rhs_append[of "[x]" S] by simp hence "t \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms S' by (cases "S' = S") auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd: assumes "send\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\ts\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \) \ set ts" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "set ts \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" using assms(1) by force ultimately have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv: assumes "receive\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\t\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - have 1: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" when "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" for t' proof - have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using that assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto qed have 2: "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" when "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" "a = Assign" for t'' proof - have "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ {t'}" using that assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "t' \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms(1) that by force ultimately show ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) that by auto qed have 3: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \' = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \" (is ?C) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \))" (is ?D) when "a = Check" proof - show ?C using that assms(2,3) by (simp add: assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) show ?D using assms(1,2,3) ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset(2) by auto qed show ?A using 1 2 by (metis subsetI) show ?B using 1 2 3 by (cases a) blast+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed subsubsection \Transition Systems Definitions\ inductive pts_symbolic:: "(('fun,'var) strands \ ('fun,'var) strand) \ (('fun,'var) strands \ ('fun,'var) strand) \ bool" (infix "\\<^sup>\" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[receive\t\\<^sub>s\<^sub>t])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[send\t\\<^sub>s\<^sub>t])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[\a: t \ t'\\<^sub>s\<^sub>t])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[\X\\\: F\\<^sub>s\<^sub>t])" private inductive pts_symbolic_c:: "(('fun,'var) strands \ ('fun,'var) extstrand) \ (('fun,'var) strands \ ('fun,'var) extstrand) \ bool" (infix "\\<^sup>\\<^sub>c" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[Step (receive\t\\<^sub>s\<^sub>t)])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[Step (send\t\\<^sub>s\<^sub>t)])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose[simp]: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ (\,\) \\<^sup>\\<^sub>c (\,\@[Decomp (Fun f T)])" abbreviation pts_symbolic_rtrancl (infix "\\<^sup>\\<^sup>*" 50) where "a \\<^sup>\\<^sup>* b \ pts_symbolic\<^sup>*\<^sup>* a b" private abbreviation pts_symbolic_c_rtrancl (infix "\\<^sup>\\<^sub>c\<^sup>*" 50) where "a \\<^sup>\\<^sub>c\<^sup>* b \ pts_symbolic_c\<^sup>*\<^sup>* a b" lemma pts_symbolic_induct[consumes 1, case_names Nil Send Receive Equality Inequality]: assumes "(\,\) \\<^sup>\ (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[receive\t\\<^sub>s\<^sub>t]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[send\t\\<^sub>s\<^sub>t]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[\a: t \ t'\\<^sub>s\<^sub>t]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[\X\\\: F\\<^sub>s\<^sub>t]\ \ P" shows "P" apply (rule pts_symbolic.cases[OF assms(1)]) using assms(2,3,4,5,6) by simp_all private lemma pts_symbolic_c_induct[consumes 1, case_names Nil Send Receive Equality Inequality Decompose]: assumes "(\,\) \\<^sup>\\<^sub>c (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[Step (receive\t\\<^sub>s\<^sub>t)]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[Step (send\t\\<^sub>s\<^sub>t)]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]\ \ P" and "\f T. \Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \); \' = \; \' = \@[Decomp (Fun f T)]\ \ P" shows "P" apply (rule pts_symbolic_c.cases[OF assms(1)]) using assms(2,3,4,5,6,7) by simp_all private lemma pts_symbolic_c_preserves_wf_prot: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' \' \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Decompose hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" using bvars_decomp ik_assignment_rhs_decomp_fv by metis+ thus ?case using Decompose unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by (metis wf_vars_mono wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2)) qed (metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq) qed metis private lemma pts_symbolic_c_preserves_wf_is: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "wf\<^sub>s\<^sub>t V (to_st \)" shows "wf\<^sub>s\<^sub>t V (to_st \')" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) hence "(\, \) \\<^sup>\\<^sub>c\<^sup>* (\2, \2)" by auto hence *: "wf\<^sub>s\<^sub>t\<^sub>s' \1 \1" "wf\<^sub>s\<^sub>t\<^sub>s' \2 \2" using pts_symbolic_c_preserves_wf_prot[OF _ step.prems(1)] step.hyps(1) by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto next case (Send ts S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (receive\ts\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t by auto thus ?case using Send wf_rcv_append''' to_st_append by simp next case (Receive ts) thus ?case using wf_snd_append to_st_append by simp next case (Equality a t t' S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (\a: t \ t'\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" when "a = Assign" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t that by auto thus ?case using Equality wf_eq_append''' to_st_append by (cases a) auto next case (Inequality t t' S) thus ?case using wf_ineq_append'' to_st_append by simp next case (Decompose f T) hence "fv (Fun f T) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1" by (metis fv_subterms_set fv_subset subset_trans ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset) hence "vars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using decomp_vars[of "Fun f T"] wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t[of \1] by auto thus ?case using to_st_append[of \1 "[Decomp (Fun f T)]"] wf_append_suffix[OF Decompose.prems] Decompose.hyps(3) by (metis append_Nil2 decomp_vars(1,2) to_st.simps(1,3)) qed qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" shows "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \'))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(trms\<^sub>s\<^sub>t ` \1) = \(trms\<^sub>s\<^sub>t ` \2)" by force thus ?case using Nil by metis next case (Decompose f T) obtain t where t: "t \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" "Fun f T \ t" using Decompose.hyps(1) by auto have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems wf_trm_subterm[of _ t] trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI[OF t(1)] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis UN_E Un_iff) have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \1)" using trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI t by auto hence "Fun f T \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" by (metis (no_types) SMP.MP SMP.Subterm UN_E t(2)) hence "{Fun f T} \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" using SMP.Subterm[of "Fun f T"] by auto moreover have "trms\<^sub>e\<^sub>s\<^sub>t \2 = insert (Fun f T) (trms\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose.hyps(3) by auto ultimately have *: "SMP (trms\<^sub>e\<^sub>s\<^sub>t \1) = SMP (trms\<^sub>e\<^sub>s\<^sub>t \2)" using SMP_subset_union_eq[of "{Fun f T}"] by (simp add: Un_commute) hence "SMP ((\(trms\<^sub>s\<^sub>t ` \1)) \ (trms\<^sub>e\<^sub>s\<^sub>t \1)) = SMP ((\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2))" using Decompose.hyps(2) SMP_union by auto moreover have "\t \ trms\<^sub>e\<^sub>s\<^sub>t \1. wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using Decompose.prems wf_trm_subterm t(2) t_wf unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto hence "\t \ trms\<^sub>e\<^sub>s\<^sub>t \2. wf\<^sub>t\<^sub>r\<^sub>m t" by (metis * SMP.MP SMP_wf_trm) hence "\t \ (\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2). wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems Decompose.hyps(2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force ultimately show ?thesis using Decompose.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed (metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq) qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "\S \ \ \ {to_st \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\S \ \' \ {to_st \'}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp have 2: "\2 = \1 - {[]}" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp_all have "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S assume "S \ \2" hence "S \ \1" using 2(1) by simp thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using 2(2) by simp qed thus ?case using 1 by auto next case (Send t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {send\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed thus ?case using 1 by auto next case (Receive t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {receive\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed show ?case using 1 3 by auto next case (Equality a t t' S) have 1: "to_st \2 = to_st \1@[\a: t \ t'\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Equality by (simp_all add: to_st_append) have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\a: t \ t'\\<^sub>s\<^sub>t]" using Equality by fastforce have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\a: t \ t'\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp have 5: "\2 = insert S (\1 - {\a: t \ t'\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp_all have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Inequality X F S) have 1: "to_st \2 = to_st \1@[\X\\\: F\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Inequality by (simp_all add: to_st_append) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: F\\<^sub>s\<^sub>t#S)" using Inequality(1,4) by blast hence 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\X\\\: F\\<^sub>s\<^sub>t]" by simp have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\X\\\: F\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp have 5: "\2 = insert S (\1 - {\X\\\: F\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp_all have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Decompose f T) hence 1: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by blast have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st [Decomp (Fun f T)])" using Decompose.prems decomp_tfr\<^sub>s\<^sub>t\<^sub>p by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1@to_st [Decomp (Fun f T)])" by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using Decompose.hyps(3) to_st_append[of \1 "[Decomp (Fun f T)]"] by auto thus ?case using 1 by blast qed qed private lemma pts_symbolic_c_preserves_well_analyzed: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "well_analyzed \" shows "well_analyzed \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Receive thus ?case by (metis well_analyzed_singleton(1) well_analyzed_append) next case Send thus ?case by (metis well_analyzed_singleton(2) well_analyzed_append) next case Equality thus ?case by (metis well_analyzed_singleton(3) well_analyzed_append) next case Inequality thus ?case by (metis well_analyzed_singleton(4) well_analyzed_append) next case (Decompose f T) hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1) - (Var`\)" by auto thus ?case by (metis well_analyzed.Decomp Decompose.prems Decompose.hyps(3)) qed simp qed metis private lemma pts_symbolic_c_preserves_Ana_invar_subst: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \)) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)))" shows "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \') \ (ik\<^sub>e\<^sub>s\<^sub>t \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \') \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \1) = \(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \2)" "\(assignment_rhs\<^sub>s\<^sub>t ` \1) = \(assignment_rhs\<^sub>s\<^sub>t ` \2)" by force+ thus ?case using Nil by metis next case Send show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd[OF Send.hyps] Ana_invar_subst_subset[OF Send.prems] by (metis Un_mono) next case Receive show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv[OF Receive.hyps] Ana_invar_subst_subset[OF Receive.prems] by (metis Un_mono) next case Equality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq[OF Equality.hyps] Ana_invar_subst_subset[OF Equality.prems] by (metis Un_mono) next case Inequality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq[OF Inequality.hyps] Ana_invar_subst_subset[OF Inequality.prems] by (metis Un_mono) next case (Decompose f T) let ?X = "\(assignment_rhs\<^sub>s\<^sub>t`\2) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2" let ?Y = "\(assignment_rhs\<^sub>s\<^sub>t`\1) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" - obtain K M where Ana: "Ana (Fun f T) = (K,M)" by moura + obtain K M where Ana: "Ana (Fun f T) = (K,M)" by atomize_elim auto hence *: "ik\<^sub>e\<^sub>s\<^sub>t \2 = ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2 = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" using ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append decomp_ik decomp_assignment_rhs_empty Decompose.hyps(3) by auto { fix g S assume "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\2) \ ik\<^sub>e\<^sub>s\<^sub>t \2 \ ?X)" hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M \ ?X)" using * Decompose.hyps(2) by auto hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(assignment_rhs\<^sub>s\<^sub>t`\1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose * Ana_fun_subterm[OF Ana] by auto moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" using trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI Decompose.hyps(1) by auto hence "subterms (Fun f T) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (metis in_subterms_subset_Union) hence "subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (meson Un_upper2 Ana_subterm[OF Ana] subterms_subset_set psubsetE subset_trans) ultimately have "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ ?Y)" by auto } thus ?case using Decompose unfolding Ana_invar_subst_def by metis qed qed private lemma pts_symbolic_c_preserves_constr_disj_vars: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "fv\<^sub>e\<^sub>s\<^sub>t \ \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}" shows "fv\<^sub>e\<^sub>s\<^sub>t \' \ bvars\<^sub>e\<^sub>s\<^sub>t \' = {}" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) have *: "\S. S \ \1 \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \1 = {}" "\S. S \ \1 \ fv\<^sub>e\<^sub>s\<^sub>t \1 \ bvars\<^sub>s\<^sub>t S = {}" using pts_symbolic_c_preserves_wf_prot[OF step.hyps(1) step.prems(1)] unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto next case (Send ts S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (send\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Send(1)] Send(4) by auto next case (Receive ts S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (receive\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Receive(1)] Receive(4) by auto next case (Equality a t t' S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (\a: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by fastforce+ thus ?case using *(1)[OF Equality(1)] Equality(4) by auto next case (Inequality X F S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1 \ set X" "fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t#S) = (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append strand_vars_split(3)[of "[\X\\\: F\\<^sub>s\<^sub>t]" S] by auto+ moreover have "fv\<^sub>e\<^sub>s\<^sub>t \1 \ set X = {}" using *(2)[OF Inequality(1)] by auto ultimately show ?case using *(1)[OF Inequality(1)] Inequality(4) by auto next case (Decompose f T) thus ?case using Decompose(3,4) bvars_decomp ik_assignment_rhs_decomp_fv[OF Decompose(1)] by auto qed qed subsubsection \Theorem: The Typing Result Lifted to the Transition System Level\ private lemma wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm: assumes "well_analyzed A" "wf\<^sub>s\<^sub>t\<^sub>s' S (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>s\<^sub>t\<^sub>s' S A" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) show "\S\S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A) (dual\<^sub>s\<^sub>t S)" by (metis (no_types) assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset wf_vars_mono le_iff_sup) show "\Sa\S. \S'\S. fv\<^sub>s\<^sub>t Sa \ bvars\<^sub>s\<^sub>t S' = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def) show "\S\S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t A = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def bvars_decomp_rm) show "\S\S. fv\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t S = {}" by (metis assms wf\<^sub>s\<^sub>t\<^sub>s'_def well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv) qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" shows "(S,A) \\<^sup>\\<^sub>c\<^sup>* (S,A@D)" using assms(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp B f X K T) have "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using ik\<^sub>e\<^sub>s\<^sub>t_append[of A B] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A B] by auto hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using Decomp.hyps by auto hence "(S,A@B) \\<^sup>\\<^sub>c (S,A@B@[Decomp (Fun f X)])" using pts_symbolic_c.Decompose[of f X "A@B"] by simp thus ?case using Decomp.IH rtrancl_into_rtrancl rtranclp_rtrancl_eq[of pts_symbolic_c "(S,A)" "(S,A@B)"] by auto qed simp private lemma pts_symbolic_to_pts_symbolic_c: assumes "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \\<^sup>\\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d" and wf: "wf\<^sub>s\<^sub>t\<^sub>s' \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" and tar: "Ana_invar_subst ((\(ik\<^sub>s\<^sub>t` dual\<^sub>s\<^sub>t` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \ (\(assignment_rhs\<^sub>s\<^sub>t` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)))" and wa: "well_analyzed \\<^sub>d" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d'. \' = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d') \ (\,\\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\',\\<^sub>d') \ sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d'" using assms(1,2) proof (induction rule: rtranclp_induct2) case refl thus ?case using assms by auto next case (step \1 \1 \2 \2) have "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \1)" using step.hyps(2) step.prems by (induct rule: pts_symbolic_induct, metis, (metis sem\<^sub>e\<^sub>s\<^sub>t_d_split_left to_est_append)+) then obtain \1d where \1d: "\1 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1d)" "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d)" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \1d" - using step.IH by moura + using step.IH by atomize_elim auto show ?case using step.hyps(2) proof (induction rule: pts_symbolic_induct) case Nil hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d)" using \1d pts_symbolic_c.Nil[OF Nil.hyps(1), of \1d] by simp thus ?case using \1d Nil by auto next case (Send t S) hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using Send.hyps(2) pts_symbolic_c.Send[OF Send.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])) = \2" using Send.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Equality a t t' S) hence "t \ \ = t' \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using Equality.hyps(2) pts_symbolic_c.Equality[OF Equality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])) = \2" using Equality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Inequality X F S) hence "ineq_model \ X F" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using Inequality.hyps(2) pts_symbolic_c.Inequality[OF Inequality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])) = \2" using Inequality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Receive ts S) hence "\t \ set ts. ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] strand_sem_split(4)[of "{}" \1 "[send\ts\\<^sub>s\<^sub>t]" \] to_st_append to_est_append to_st_to_est_inv by auto moreover have "ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \" using \1d(1) decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset by auto ultimately have *: "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using ideduct_mono by auto have "wf\<^sub>s\<^sub>t\<^sub>s' \ \\<^sub>d" by (rule wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm[OF wa assms(4)]) hence **: "wf\<^sub>e\<^sub>s\<^sub>t {} \1d" by (rule pts_symbolic_c_preserves_wf_is[OF \1d(2) _ assms(5)]) have "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\1) \ (ik\<^sub>e\<^sub>s\<^sub>t \1d) \ (\(assignment_rhs\<^sub>s\<^sub>t`\1) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)))" using tar \1d(2) pts_symbolic_c_preserves_Ana_invar_subst by metis hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t \1d)" "Ana_invar_subst (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)" using Ana_invar_subst_subset by blast+ moreover have "well_analyzed \1d" using pts_symbolic_c_preserves_well_analyzed[OF \1d(2) wa] by metis ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1d) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d) \" "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t (\1d@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst_list[OF * \1d(3) ** assms(8)] unfolding Ana_invar_subst_def by auto have "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d@D)" using \1d(2) decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c[OF D(1), of \1] by auto hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d@D@[Step (send\ts\\<^sub>s\<^sub>t)])" using Receive(2) pts_symbolic_c.Receive[OF Receive.hyps(1), of "\1d@D"] by auto moreover have "\2 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@D@[Step (send\ts\\<^sub>s\<^sub>t)]))" using Receive.hyps(3) \1d(1) decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty[OF D(1)] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append by auto moreover have "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@D@[Step (send\ts\\<^sub>s\<^sub>t)])" using D(2) sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append[OF \1d(3) D(1)]] by simp ultimately show ?case by auto qed qed private lemma pts_symbolic_c_to_pts_symbolic: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \'" shows "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \')" proof - show "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" using assms(1) proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) show ?case using step.hyps(2,1) step.IH proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case using pts_symbolic.Nil[OF Nil.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by simp next case (Send t S) thus ?case using pts_symbolic.Send[OF Send.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Receive t S) thus ?case using pts_symbolic.Receive[OF Receive.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Equality a t t' S) thus ?case using pts_symbolic.Equality[OF Equality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Inequality t t' S) thus ?case using pts_symbolic.Inequality[OF Inequality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Decompose t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by simp qed qed simp qed (rule sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c[OF assms(2)]) private lemma pts_symbolic_to_pts_symbolic_c_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "\ \ \\\" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d. \ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d) \ (\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d) \ (\ \\<^sub>c \to_st \\<^sub>d\)" using assms pts_symbolic_to_pts_symbolic_c[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] to_st_to_est_inv[of \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def simp del: subst_range.simps) private lemma pts_symbolic_c_to_pts_symbolic_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\)" "\ \\<^sub>c \to_st \\" shows "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \))" "\ \ \to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)\" using assms pts_symbolic_c_to_pts_symbolic[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def) private lemma to_st_trms_wf: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st A))" using assms proof (induction A) case (Cons x A) hence IH: "\t \ trms\<^sub>s\<^sub>t (to_st A). wf\<^sub>t\<^sub>r\<^sub>m t" by auto with Cons show ?case proof (cases x) case (Decomp t) hence "wf\<^sub>t\<^sub>r\<^sub>m t" using Cons.prems by auto - obtain K T where Ana_t: "Ana t = (K,T)" by moura + obtain K T where Ana_t: "Ana t = (K,T)" by atomize_elim auto hence "trms\<^sub>s\<^sub>t (decomp t) \ {t} \ set K \ set T" using decomp_set_unfold[OF Ana_t] by force moreover have "\t \ set T. wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_subterm[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ wf_trm_subterm by auto ultimately have "\t \ trms\<^sub>s\<^sub>t (decomp t). wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ by auto thus ?thesis using IH Decomp by auto qed auto qed simp private lemma to_st_trms_SMP_subset: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof fix t assume "t \ trms\<^sub>s\<^sub>t (to_st A)" thus "t \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A) case (Cons x A) hence *: "t \ trms\<^sub>s\<^sub>t (to_st [x]) \ trms\<^sub>s\<^sub>t (to_st A)" using to_st_append[of "[x]" A] by auto have **: "trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>s\<^sub>t (to_st (x#A))" "trms\<^sub>e\<^sub>s\<^sub>t A \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using to_st_append[of "[x]" A] by auto show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (to_st A)") case True thus ?thesis using Cons.IH SMP_mono[OF **(2)] by auto next case False hence ***: "t \ trms\<^sub>s\<^sub>t (to_st [x])" using * by auto thus ?thesis proof (cases x) case (Decomp t') hence ****: "t \ trms\<^sub>s\<^sub>t (decomp t')" "t' \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using *** by auto - obtain K T where Ana_t': "Ana t' = (K,T)" by moura + obtain K T where Ana_t': "Ana t' = (K,T)" by atomize_elim auto hence "t \ {t'} \ set K \ set T" using decomp_set_unfold[OF Ana_t'] ****(1) by force moreover { assume "t = t'" hence ?thesis using SMP.MP[OF ****(2)] by simp } moreover { assume "t \ set K" hence ?thesis using SMP.Ana[OF SMP.MP[OF ****(2)] Ana_t'] by auto } moreover { assume "t \ set T" "t \ t'" hence "t \ t'" using Ana_subterm[OF Ana_t'] by blast hence ?thesis using SMP.Subterm[OF SMP.MP[OF ****(2)]] by auto } ultimately show ?thesis using Decomp by auto qed auto qed qed simp qed private lemma to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st A))" proof - have *: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" using to_st_trms_wf to_st_trms_SMP_subset assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto have "trms\<^sub>s\<^sub>t (to_st A) = trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>e\<^sub>s\<^sub>t A" by (blast dest!: trms\<^sub>e\<^sub>s\<^sub>tD) hence "SMP (trms\<^sub>e\<^sub>s\<^sub>t A) = SMP (trms\<^sub>s\<^sub>t (to_st A))" using SMP_subset_union_eq[OF *] by auto thus ?thesis using * assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed theorem wt_attack_if_tfr_attack_pts: assumes "wf\<^sub>s\<^sub>t\<^sub>s \\<^sub>0" "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "\S \ \\<^sub>0. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" and "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\, Var\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\, Var\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have "(\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []) = \(trms\<^sub>s\<^sub>t ` \\<^sub>0)" "to_st [] = []" "list_all tfr\<^sub>s\<^sub>t\<^sub>p []" using assms by simp_all hence *: "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" "\S \ \\<^sub>0 \ {to_st []}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by (metis, metis, metis, simp) obtain \\<^sub>d where \\<^sub>d: "\ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d)" "\ \\<^sub>c \to_st \\<^sub>d\" using pts_symbolic_to_pts_symbolic_c_from_initial assms *(3) by metis hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t[OF _ *(1,2)] by blast+ hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis DiffE DiffI SMP_union UnCI, metis UnCI) hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" by (metis to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t, metis to_st_trms_wf) moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (to_st \\<^sub>d) Var" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "subst_domain Var \ vars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" "range_vars Var \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" by (simp_all add: range_vars_alt_def) moreover have "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" using pts_symbolic_c_preserves_wf_is[OF \\<^sub>d(2) *(3), of "{}"] by auto moreover have "fv\<^sub>s\<^sub>t (to_st \\<^sub>d) \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" using pts_symbolic_c_preserves_constr_disj_vars[OF \\<^sub>d(2)] assms(1) wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by fastforce ultimately show ?thesis unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp qed moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p[OF \\<^sub>d(2) *(4)] by blast moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp_all ultimately obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \to_st \\<^sub>d, Var\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_attack_if_tfr_attack[OF assms(7) \\<^sub>d(3)] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))\ \list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)\ unfolding tfr\<^sub>s\<^sub>t_def by metis hence "\\<^sub>\ \ \\, Var\" using pts_symbolic_c_to_pts_symbolic_from_initial \\<^sub>d by metis thus ?thesis using \\<^sub>\(1,3,4) by metis qed subsubsection \Corollary: The Typing Result on the Level of Constraints\ text \There exists well-typed models of satisfiable type-flaw resistant constraints\ corollary wt_attack_if_tfr_attack_d: assumes "wf\<^sub>s\<^sub>t {} \" "fv\<^sub>s\<^sub>t \ \ bvars\<^sub>s\<^sub>t \ = {}" "tfr\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t \)" and "Ana_invar_subst (ik\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>s\<^sub>t \)" and "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - { fix S A have "({S},A) \\<^sup>\\<^sup>* ({},A@dual\<^sub>s\<^sub>t S)" proof (induction S arbitrary: A) case Nil thus ?case using pts_symbolic.Nil[of "{[]}"] by auto next case (Cons x S) hence "({S}, A@dual\<^sub>s\<^sub>t [x]) \\<^sup>\\<^sup>* ({}, A@dual\<^sub>s\<^sub>t (x#S))" by (metis dual\<^sub>s\<^sub>t_append List.append_assoc List.append_Nil List.append_Cons) moreover have "({x#S}, A) \\<^sup>\ ({S}, A@dual\<^sub>s\<^sub>t [x])" using pts_symbolic.Send[of _ S "{x#S}"] pts_symbolic.Receive[of _ S "{x#S}"] pts_symbolic.Equality[of _ _ _ S "{x#S}"] pts_symbolic.Inequality[of _ _ S "{x#S}"] by (cases x) auto ultimately show ?case by simp qed } hence 0: "({dual\<^sub>s\<^sub>t \},[]) \\<^sup>\\<^sup>* ({},\)" using dual\<^sub>s\<^sub>t_self_inverse by (metis List.append_Nil) have "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) \ bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) = {}" using assms(2) dual\<^sub>s\<^sub>t_fv dual\<^sub>s\<^sub>t_bvars by metis+ hence 1: "wf\<^sub>s\<^sub>t\<^sub>s {dual\<^sub>s\<^sub>t \}" using assms(1,2) dual\<^sub>s\<^sub>t_self_inverse[of \] unfolding wf\<^sub>s\<^sub>t\<^sub>s_def by auto have "\(trms\<^sub>s\<^sub>t ` {\}) = trms\<^sub>s\<^sub>t \" "\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}) = trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by auto hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {\}))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {\}))" "(\(trms\<^sub>s\<^sub>t ` {\})) = \(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \})" using assms(3,4) unfolding tfr\<^sub>s\<^sub>t_def by (metis, metis, metis dual\<^sub>s\<^sub>t_trms_eq) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" by metis+ have 4: "\S \ {dual\<^sub>s\<^sub>t \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p assms(3) unfolding tfr\<^sub>s\<^sub>t_def by blast have "assignment_rhs\<^sub>s\<^sub>t \ = assignment_rhs\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by (induct \ rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto hence 5: "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}) \ \(assignment_rhs\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}))" using assms(5) dual\<^sub>s\<^sub>t_self_inverse[of \] by auto show ?thesis by (rule wt_attack_if_tfr_attack_pts[OF 1 2 3 4 5 0 assms(6,7)]) qed end end end