diff --git a/thys/ResiduatedTransitionSystem/ResiduatedTransitionSystem.thy b/thys/ResiduatedTransitionSystem/ResiduatedTransitionSystem.thy --- a/thys/ResiduatedTransitionSystem/ResiduatedTransitionSystem.thy +++ b/thys/ResiduatedTransitionSystem/ResiduatedTransitionSystem.thy @@ -1,8849 +1,8849 @@ chapter "Residuated Transition Systems" theory ResiduatedTransitionSystem imports Main begin section "Basic Definitions and Properties" subsection "Partial Magmas" text \ A \emph{partial magma} consists simply of a partial binary operation. We represent the partiality by assuming the existence of a unique value \null\ that behaves as a zero for the operation. \ (* TODO: Possibly unify with Category3.partial_magma? *) locale partial_magma = fixes OP :: "'a \ 'a \ 'a" assumes ex_un_null: "\!n. \t. OP n t = n \ OP t n = n" begin definition null :: 'a where "null = (THE n. \t. OP n t = n \ OP t n = n)" lemma null_eqI: assumes "\t. OP n t = n \ OP t n = n" shows "n = null" using assms null_def ex_un_null the1_equality [of "\n. \t. OP n t = n \ OP t n = n"] by auto lemma null_is_zero [simp]: shows "OP null t = null" and "OP t null = null" using null_def ex_un_null theI' [of "\n. \t. OP n t = n \ OP t n = n"] by auto end subsection "Residuation" text \ A \emph{residuation} is a partial binary operation subject to three axioms. The first, \con_sym_ax\, states that the domain of a residuation is symmetric. The second, \con_imp_arr_resid\, constrains the results of residuation either to be \null\, which indicates inconsistency, or something that is self-consistent, which we will define below to be an ``arrow''. The ``cube axiom'', \cube_ax\, states that if \v\ can be transported by residuation around one side of the ``commuting square'' formed by \t\ and \u \ t\, then it can also be transported around the other side, formed by \u\ and \t \ u\, with the same result. \ type_synonym 'a resid = "'a \ 'a \ 'a" locale residuation = partial_magma resid for resid :: "'a resid" (infix "\\" 70) + assumes con_sym_ax: "t \\ u \ null \ u \\ t \ null" and con_imp_arr_resid: "t \\ u \ null \ (t \\ u) \\ (t \\ u) \ null" and cube_ax: "(v \\ t) \\ (u \\ t) \ null \ (v \\ t) \\ (u \\ t) = (v \\ u) \\ (t \\ u)" begin text \ The axiom \cube_ax\ is equivalent to the following unconditional form. The locale assumptions use the weaker form to avoid having to treat the case \(v \ t) \ (u \ t) = null\ specially for every interpretation. \ lemma cube: shows "(v \\ t) \\ (u \\ t) = (v \\ u) \\ (t \\ u)" using cube_ax by metis text \ We regard \t\ and \u\ as \emph{consistent} if the residuation \t \ u\ is defined. It is convenient to make this a definition, with associated notation. \ definition con (infix "\" 50) where "t \ u \ t \\ u \ null" lemma conI [intro]: assumes "t \\ u \ null" shows "t \ u" using assms con_def by blast lemma conE [elim]: assumes "t \ u" and "t \\ u \ null \ T" shows T using assms con_def by simp lemma con_sym: assumes "t \ u" shows "u \ t" using assms con_def con_sym_ax by blast text \ We call \t\ an \emph{arrow} if it is self-consistent. \ definition arr where "arr t \ t \ t" lemma arrI [intro]: assumes "t \ t" shows "arr t" using assms arr_def by simp lemma arrE [elim]: assumes "arr t" and "t \ t \ T" shows T using assms arr_def by simp lemma not_arr_null [simp]: shows "\ arr null" by (auto simp add: con_def) lemma con_implies_arr: assumes "t \ u" shows "arr t" and "arr u" using assms by (metis arrI con_def con_imp_arr_resid cube null_is_zero(2))+ lemma arr_resid [simp]: assumes "t \ u" shows "arr (t \\ u)" using assms con_imp_arr_resid by blast lemma arr_resid_iff_con: shows "arr (t \\ u) \ t \ u" by auto text \ The residuation of an arrow along itself is the \emph{canonical target} of the arrow. \ definition trg where "trg t \ t \\ t" lemma resid_arr_self: shows "t \\ t = trg t" using trg_def by auto text \ An \emph{identity} is an arrow that is its own target. \ definition ide where "ide a \ a \ a \ a \\ a = a" lemma ideI [intro]: assumes "a \ a" and "a \\ a = a" shows "ide a" using assms ide_def by auto lemma ideE [elim]: assumes "ide a" and "\a \ a; a \\ a = a\ \ T" shows T using assms ide_def by blast lemma ide_implies_arr [simp]: assumes "ide a" shows "arr a" using assms by blast end subsection "Residuated Transition System" text \ A \emph{residuated transition system} consists of a residuation subject to additional axioms that concern the relationship between identities and residuation. These axioms make it possible to sensibly associate with each arrow certain nonempty sets of identities called the \emph{sources} and \emph{targets} of the arrow. Axiom \ide_trg\ states that the canonical target \trg t\ of an arrow \t\ is an identity. Axiom \resid_arr_ide\ states that identities are right units for residuation, when it is defined. Axiom \resid_ide_arr\ states that the residuation of an identity along an arrow is again an identity, assuming that the residuation is defined. Axiom \con_imp_coinitial_ax\ states that if arrows \t\ and \u\ are consistent, then there is an identity that is consistent with both of them (\emph{i.e.}~they have a common source). Axiom \con_target\ states that an identity of the form \t \ u\ (which may be regarded as a ``target'' of \u\) is consistent with any other arrow \v \ u\ obtained by residuation along \u\. We note that replacing the premise \ide (t \ u)\ in this axiom by either \arr (t \ u)\ or \t \ u\ would result in a strictly stronger statement. \ locale rts = residuation + assumes ide_trg [simp]: "arr t \ ide (trg t)" and resid_arr_ide: "\ide a; t \ a\ \ t \\ a = t" and resid_ide_arr [simp]: "\ide a; a \ t\ \ ide (a \\ t)" and con_imp_coinitial_ax: "t \ u \ \a. ide a \ a \ t \ a \ u" and con_target: "\ide (t \\ u); u \ v\ \ t \\ u \ v \\ u" begin text \ We define the \emph{sources} of an arrow \t\ to be the identities that are consistent with \t\. \ definition sources where "sources t = {a. ide a \ t \ a}" text \ We define the \emph{targets} of an arrow \t\ to be the identities that are consistent with the canonical target \trg t\. \ definition targets where "targets t = {b. ide b \ trg t \ b}" lemma in_sourcesI [intro, simp]: assumes "ide a" and "t \ a" shows "a \ sources t" using assms sources_def by simp lemma in_sourcesE [elim]: assumes "a \ sources t" and "\ide a; t \ a\ \ T" shows T using assms sources_def by auto lemma in_targetsI [intro, simp]: assumes "ide b" and "trg t \ b" shows "b \ targets t" using assms targets_def resid_arr_self by simp lemma in_targetsE [elim]: assumes "b \ targets t" and "\ide b; trg t \ b\ \ T" shows T using assms targets_def resid_arr_self by force lemma trg_in_targets: assumes "arr t" shows "trg t \ targets t" using assms by (meson ideE ide_trg in_targetsI) lemma source_is_ide: assumes "a \ sources t" shows "ide a" using assms by blast lemma target_is_ide: assumes "a \ targets t" shows "ide a" using assms by blast text \ Consistent arrows have a common source. \ lemma con_imp_common_source: assumes "t \ u" shows "sources t \ sources u \ {}" using assms by (meson disjoint_iff in_sourcesI con_imp_coinitial_ax con_sym) text \ Arrows are characterized by the property of having a nonempty set of sources, or equivalently, by that of having a nonempty set of targets. \ lemma arr_iff_has_source: shows "arr t \ sources t \ {}" using con_imp_common_source con_implies_arr(1) sources_def by blast lemma arr_iff_has_target: shows "arr t \ targets t \ {}" using trg_def trg_in_targets by fastforce text \ The residuation of a source of an arrow along that arrow gives a target of the same arrow. However, it is \emph{not} true that every target of an arrow \t\ is of the form \u \ t\ for some \u\ with \t \ u\. \ lemma resid_source_in_targets: assumes "a \ sources t" shows "a \\ t \ targets t" by (metis arr_resid assms con_target con_sym resid_arr_ide ide_trg in_sourcesE resid_ide_arr in_targetsI resid_arr_self) text \ Residuation along an identity reflects identities. \ lemma ide_backward_stable: assumes "ide a" and "ide (t \\ a)" shows "ide t" by (metis assms ideE resid_arr_ide arr_resid_iff_con) lemma resid_reflects_con: assumes "t \ v" and "u \ v" and "t \\ v \ u \\ v" shows "t \ u" using assms cube by (elim conE) auto lemma con_transitive_on_ide: assumes "ide a" and "ide b" and "ide c" shows "\a \ b; b \ c\ \ a \ c" using assms by (metis resid_arr_ide con_target con_sym) lemma sources_are_con: assumes "a \ sources t" and "a' \ sources t" shows "a \ a'" using assms by (metis (no_types, lifting) CollectD con_target con_sym resid_ide_arr sources_def resid_reflects_con) lemma sources_con_closed: assumes "a \ sources t" and "ide a'" and "a \ a'" shows "a' \ sources t" using assms by (metis (no_types, lifting) con_target con_sym resid_arr_ide mem_Collect_eq sources_def) lemma sources_eqI: assumes "sources t \ sources t' \ {}" shows "sources t = sources t'" using assms sources_def sources_are_con sources_con_closed by blast lemma targets_are_con: assumes "b \ targets t" and "b' \ targets t" shows "b \ b'" using assms sources_are_con sources_def targets_def by blast lemma targets_con_closed: assumes "b \ targets t" and "ide b'" and "b \ b'" shows "b' \ targets t" using assms sources_con_closed sources_def targets_def by blast lemma targets_eqI: assumes "targets t \ targets t' \ {}" shows "targets t = targets t'" using assms targets_def targets_are_con targets_con_closed by blast text \ Arrows are \emph{coinitial} if they have a common source, and \emph{coterminal} if they have a common target. \ definition coinitial where "coinitial t u \ sources t \ sources u \ {}" definition coterminal where "coterminal t u \ targets t \ targets u \ {}" lemma coinitialI [intro]: assumes "arr t" and "sources t = sources u" shows "coinitial t u" using assms coinitial_def arr_iff_has_source by simp lemma coinitialE [elim]: assumes "coinitial t u" and "\arr t; arr u; sources t = sources u\ \ T" shows T using assms coinitial_def sources_eqI arr_iff_has_source by auto lemma con_imp_coinitial: assumes "t \ u" shows "coinitial t u" using assms by (simp add: coinitial_def con_imp_common_source) lemma coinitial_iff: shows "coinitial t t' \ arr t \ arr t' \ sources t = sources t'" by (metis arr_iff_has_source coinitial_def inf_idem sources_eqI) lemma coterminal_iff: shows "coterminal t t' \ arr t \ arr t' \ targets t = targets t'" by (metis arr_iff_has_target coterminal_def inf_idem targets_eqI) lemma coterminal_iff_con_trg: shows "coterminal t u \ trg t \ trg u" by (metis coinitial_iff con_imp_coinitial coterminal_iff in_targetsE trg_in_targets resid_arr_self arr_resid_iff_con sources_def targets_def) lemma coterminalI [intro]: assumes "arr t" and "targets t = targets u" shows "coterminal t u" using assms coterminal_iff arr_iff_has_target by auto lemma coterminalE [elim]: assumes "coterminal t u" and "\arr t; arr u; targets t = targets u\ \ T" shows T using assms coterminal_iff by auto lemma sources_resid [simp]: assumes "t \ u" shows "sources (t \\ u) = targets u" unfolding targets_def trg_def using assms conI conE by (metis con_imp_arr_resid assms coinitial_iff con_imp_coinitial cube ex_un_null sources_def) lemma targets_resid_sym: assumes "t \ u" shows "targets (t \\ u) = targets (u \\ t)" using assms apply (intro targets_eqI) by (metis (no_types, opaque_lifting) assms cube inf_idem arr_iff_has_target arr_def arr_resid_iff_con sources_resid) text \ Arrows \t\ and \u\ are \emph{sequential} if the set of targets of \t\ equals the set of sources of \u\. \ definition seq where "seq t u \ arr t \ arr u \ targets t = sources u" lemma seqI [intro]: assumes "arr t" and "arr u" and "targets t = sources u" shows "seq t u" using assms seq_def by auto lemma seqE [elim]: assumes "seq t u" and "\arr t; arr u; targets t = sources u\ \ T" shows T using assms seq_def by blast subsubsection "Congruence of Transitions" text \ Residuation induces a preorder \\\ on transitions, defined by \t \ u\ if and only if \t \ u\ is an identity. \ abbreviation prfx (infix "\" 50) where "t \ u \ ide (t \\ u)" lemma prfx_implies_con: assumes "t \ u" shows "t \ u" using assms arr_resid_iff_con by blast lemma prfx_reflexive: assumes "arr t" shows "t \ t" by (simp add: assms resid_arr_self) lemma prfx_transitive [trans]: assumes "t \ u" and "u \ v" shows "t \ v" using assms con_target resid_ide_arr ide_backward_stable cube conI by metis text \ The equivalence \\\ associated with \\\ is substitutive with respect to residuation. \ abbreviation cong (infix "\" 50) where "t \ u \ t \ u \ u \ t" lemma cong_reflexive: assumes "arr t" shows "t \ t" using assms prfx_reflexive by simp lemma cong_symmetric: assumes "t \ u" shows "u \ t" using assms by simp lemma cong_transitive [trans]: assumes "t \ u" and "u \ v" shows "t \ v" using assms prfx_transitive by auto lemma cong_subst_left: assumes "t \ t'" and "t \ u" shows "t' \ u" and "t \\ u \ t' \\ u" apply (meson assms con_sym con_target prfx_implies_con resid_reflects_con) by (metis assms con_sym con_target cube prfx_implies_con resid_ide_arr resid_reflects_con) lemma cong_subst_right: assumes "u \ u'" and "t \ u" shows "t \ u'" and "t \\ u \ t \\ u'" proof - have 1: "t \ u' \ t \\ u' \ u \\ u' \ (t \\ u) \\ (u' \\ u) = (t \\ u') \\ (u \\ u')" using assms cube con_sym con_target cong_subst_left(1) by meson show "t \ u'" using 1 by simp show "t \\ u \ t \\ u'" by (metis 1 arr_resid_iff_con assms(1) cong_reflexive resid_arr_ide) qed lemma cong_implies_coinitial: assumes "u \ u'" shows "coinitial u u'" using assms con_imp_coinitial prfx_implies_con by simp lemma cong_implies_coterminal: assumes "u \ u'" shows "coterminal u u'" using assms by (metis con_implies_arr(1) coterminalI ideE prfx_implies_con sources_resid targets_resid_sym) lemma ide_imp_con_iff_cong: assumes "ide t" and "ide u" shows "t \ u \ t \ u" using assms by (metis con_sym resid_ide_arr prfx_implies_con) lemma sources_are_cong: assumes "a \ sources t" and "a' \ sources t" shows "a \ a'" using assms sources_are_con by (metis CollectD ide_imp_con_iff_cong sources_def) lemma sources_cong_closed: assumes "a \ sources t" and "a \ a'" shows "a' \ sources t" using assms sources_def by (meson in_sourcesE in_sourcesI cong_subst_right(1) ide_backward_stable) lemma targets_are_cong: assumes "b \ targets t" and "b' \ targets t" shows "b \ b'" using assms(1-2) sources_are_cong sources_def targets_def by blast lemma targets_cong_closed: assumes "b \ targets t" and "b \ b'" shows "b' \ targets t" using assms targets_def sources_cong_closed sources_def by blast lemma targets_char: shows "targets t = {b. arr t \ t \\ t \ b}" unfolding targets_def by (metis (no_types, lifting) con_def con_implies_arr(2) con_sym cong_reflexive ide_def resid_arr_ide trg_def) lemma coinitial_ide_are_cong: assumes "ide a" and "ide a'" and "coinitial a a'" shows "a \ a'" using assms coinitial_def by (metis ideE in_sourcesI coinitialE sources_are_cong) lemma cong_respects_seq: assumes "seq t u" and "cong t t'" and "cong u u'" shows "seq t' u'" by (metis assms coterminalE rts.coinitialE rts.cong_implies_coinitial rts.cong_implies_coterminal rts_axioms seqE seqI) end subsection "Weakly Extensional RTS" text \ A \emph{weakly extensional} RTS is an RTS that satisfies the additional condition that identity arrows have trivial congruence classes. This axiom has a number of useful consequences, including that each arrow has a unique source and target. \ locale weakly_extensional_rts = rts + assumes weak_extensionality: "\t \ u; ide t; ide u\ \ t = u" begin lemma con_ide_are_eq: assumes "ide a" and "ide a'" and "a \ a'" shows "a = a'" using assms ide_imp_con_iff_cong weak_extensionality by blast lemma coinitial_ide_are_eq: assumes "ide a" and "ide a'" and "coinitial a a'" shows "a = a'" using assms coinitial_def con_ide_are_eq by blast lemma arr_has_un_source: assumes "arr t" shows "\!a. a \ sources t" using assms by (meson arr_iff_has_source con_ide_are_eq ex_in_conv in_sourcesE sources_are_con) lemma arr_has_un_target: assumes "arr t" shows "\!b. b \ targets t" using assms by (metis arrE arr_has_un_source arr_resid sources_resid) definition src where "src t \ if arr t then THE a. a \ sources t else null" lemma src_in_sources: assumes "arr t" shows "src t \ sources t" using assms src_def arr_has_un_source the1I2 [of "\a. a \ sources t" "\a. a \ sources t"] by simp lemma src_eqI: assumes "ide a" and "a \ t" shows "src t = a" using assms src_in_sources by (metis arr_has_un_source resid_arr_ide in_sourcesI arr_resid_iff_con con_sym) lemma sources_char: shows "sources t = {a. arr t \ src t = a}" using src_in_sources arr_has_un_source arr_iff_has_source by auto lemma targets_char\<^sub>W\<^sub>E: shows "targets t = {b. arr t \ trg t = b}" using trg_in_targets arr_has_un_target arr_iff_has_target by auto lemma arr_src_iff_arr [iff]: shows "arr (src t) \ arr t" by (metis arrI conE null_is_zero(2) sources_are_con arrE src_def src_in_sources) lemma arr_trg_iff_arr [iff]: shows "arr (trg t) \ arr t" by (metis arrI arrE arr_resid_iff_con resid_arr_self) lemma con_imp_eq_src: assumes "t \ u" shows "src t = src u" using assms by (metis con_imp_coinitial_ax src_eqI) lemma src_resid [simp]: assumes "t \ u" shows "src (t \\ u) = trg u" using assms by (metis arr_resid_iff_con con_implies_arr(2) arr_has_un_source trg_in_targets sources_resid src_in_sources) lemma trg_resid_sym: assumes "t \ u" shows "trg (t \\ u) = trg (u \\ t)" using assms by (metis arr_has_un_target arr_resid con_sym targets_resid_sym trg_in_targets) lemma apex_sym: shows "trg (t \\ u) = trg (u \\ t)" using trg_resid_sym con_def by metis lemma seqI\<^sub>W\<^sub>E [intro, simp]: assumes "arr u" and "arr t" and "trg t = src u" shows "seq t u" using assms by (metis (mono_tags, lifting) arrE in_sourcesE resid_arr_ide sources_resid resid_arr_self seqI sources_are_con src_in_sources) lemma seqE\<^sub>W\<^sub>E [elim]: assumes "seq t u" and "\arr u; arr t; trg t = src u\ \ T" shows T using assms by (metis arr_has_un_source seq_def src_in_sources trg_in_targets) lemma coinitial_iff\<^sub>W\<^sub>E: shows "coinitial t u \ arr t \ arr u \ src t = src u" by (metis arr_has_un_source coinitial_def coinitial_iff disjoint_iff_not_equal src_in_sources) lemma coterminal_iff\<^sub>W\<^sub>E: shows "coterminal t u \ arr t \ arr u \ trg t = trg u" by (metis arr_has_un_target coterminal_iff_con_trg coterminal_iff trg_in_targets) lemma coinitialI\<^sub>W\<^sub>E [intro]: assumes "arr t" and "src t = src u" shows "coinitial t u" using assms coinitial_iff\<^sub>W\<^sub>E by (metis arr_src_iff_arr) lemma coinitialE\<^sub>W\<^sub>E [elim]: assumes "coinitial t u" and "\arr t; arr u; src t = src u\ \ T" shows T using assms coinitial_iff\<^sub>W\<^sub>E by blast lemma coterminalI\<^sub>W\<^sub>E [intro]: assumes "arr t" and "trg t = trg u" shows "coterminal t u" using assms coterminal_iff\<^sub>W\<^sub>E by (metis arr_trg_iff_arr) lemma coterminalE\<^sub>W\<^sub>E [elim]: assumes "coterminal t u" and "\arr t; arr u; trg t = trg u\ \ T" shows T using assms coterminal_iff\<^sub>W\<^sub>E by blast lemma ide_src [simp]: assumes "arr t" shows "ide (src t)" using assms by (metis arrE con_imp_coinitial_ax src_eqI) lemma src_ide [simp]: assumes "ide a" shows "src a = a" using arrI assms src_eqI by blast lemma trg_ide [simp]: assumes "ide a" shows "trg a = a" using assms resid_arr_self by force lemma ide_iff_src_self: assumes "arr a" shows "ide a \ src a = a" using assms by (metis ide_src src_ide) lemma ide_iff_trg_self: assumes "arr a" shows "ide a \ trg a = a" using assms ide_def resid_arr_self by auto lemma src_src [simp]: shows "src (src t) = src t" using ide_src src_def src_ide by auto lemma trg_trg [simp]: shows "trg (trg t) = trg t" by (metis con_def cong_reflexive ide_def null_is_zero(2) resid_arr_self residuation.con_implies_arr(1) residuation_axioms) lemma src_trg [simp]: shows "src (trg t) = trg t" by (metis con_def not_arr_null src_def src_resid trg_def) lemma trg_src [simp]: shows "trg (src t) = src t" by (metis ide_src null_is_zero(2) resid_arr_self src_def trg_ide) lemma resid_ide: assumes "ide a" and "coinitial a t" shows (* [simp]: *) "t \\ a = t" and "a \\ t = trg t" using assms resid_arr_ide apply blast using assms by (metis con_def con_sym_ax ideE in_sourcesE in_sourcesI resid_ide_arr coinitialE src_ide src_resid) end subsection "Extensional RTS" text \ An \emph{extensional} RTS is an RTS in which all arrows have trivial congruence classes; that is, congruent arrows are equal. \ locale extensional_rts = rts + assumes extensional: "t \ u \ t = u" begin sublocale weakly_extensional_rts using extensional by unfold_locales auto lemma cong_char: shows "t \ u \ arr t \ t = u" by (metis arrI cong_reflexive prfx_implies_con extensional) end subsection "Composites of Transitions" text \ Residuation can be used to define a notion of composite of transitions. Composites are not unique, but they are unique up to congruence. \ context rts begin definition composite_of where "composite_of u t v \ u \ v \ v \\ u \ t" lemma composite_ofI [intro]: assumes "u \ v" and "v \\ u \ t" shows "composite_of u t v" using assms composite_of_def by blast lemma composite_ofE [elim]: assumes "composite_of u t v" and "\u \ v; v \\ u \ t\ \ T" shows T using assms composite_of_def by auto lemma arr_composite_of: assumes "composite_of u t v" shows "arr v" using assms by (meson composite_of_def con_implies_arr(2) prfx_implies_con) lemma composite_of_unq_upto_cong: assumes "composite_of u t v" and "composite_of u t v'" shows "v \ v'" using assms cube ide_backward_stable prfx_transitive by (elim composite_ofE) metis lemma composite_of_ide_arr: assumes "ide a" shows "composite_of a t t \ t \ a" using assms by (metis composite_of_def con_implies_arr(1) con_sym resid_arr_ide resid_ide_arr prfx_implies_con prfx_reflexive) lemma composite_of_arr_ide: assumes "ide b" shows "composite_of t b t \ t \\ t \ b" using assms by (metis arr_resid_iff_con composite_of_def ide_imp_con_iff_cong con_implies_arr(1) prfx_implies_con prfx_reflexive) lemma composite_of_source_arr: assumes "arr t" and "a \ sources t" shows "composite_of a t t" using assms composite_of_ide_arr sources_def by auto lemma composite_of_arr_target: assumes "arr t" and "b \ targets t" shows "composite_of t b t" by (metis arrE assms composite_of_arr_ide in_sourcesE sources_resid) lemma composite_of_ide_self: assumes "ide a" shows "composite_of a a a" using assms composite_of_ide_arr by blast lemma con_prfx_composite_of: assumes "composite_of t u w" shows "t \ w" and "w \ v \ t \ v" using assms apply force using assms composite_of_def con_target prfx_implies_con resid_reflects_con con_sym by meson lemma sources_composite_of: assumes "composite_of u t v" shows "sources v = sources u" using assms by (meson arr_resid_iff_con composite_of_def con_imp_coinitial cong_implies_coinitial coinitial_iff) lemma targets_composite_of: assumes "composite_of u t v" shows "targets v = targets t" proof - have "targets t = targets (v \\ u)" using assms composite_of_def by (meson cong_implies_coterminal coterminal_iff) also have "... = targets (u \\ v)" using assms targets_resid_sym con_prfx_composite_of by metis also have "... = targets v" using assms composite_of_def by (metis prfx_implies_con sources_resid ideE) finally show ?thesis by auto qed lemma resid_composite_of: assumes "composite_of t u w" and "w \ v" shows "v \\ t \ w \\ t" and "v \\ t \ u" and "v \\ w \ (v \\ t) \\ u" and "composite_of (t \\ v) (u \\ (v \\ t)) (w \\ v)" proof - show 0: "v \\ t \ w \\ t" using assms con_def by (metis con_target composite_ofE conE con_sym cube) show 1: "v \\ w \ (v \\ t) \\ u" proof - have "v \\ w = (v \\ w) \\ (t \\ w)" using assms composite_of_def by (metis (no_types, opaque_lifting) con_target con_sym resid_arr_ide) also have "... = (v \\ t) \\ (w \\ t)" using assms cube by metis also have "... \ (v \\ t) \\ u" using assms 0 cong_subst_right(2) [of "w \\ t" u "v \\ t"] by blast finally show ?thesis by blast qed show 2: "v \\ t \ u" using assms 1 by force show "composite_of (t \\ v) (u \\ (v \\ t)) (w \\ v)" proof (unfold composite_of_def, intro conjI) show "t \\ v \ w \\ v" using assms cube con_target composite_of_def resid_ide_arr by metis show "(w \\ v) \\ (t \\ v) \ u \\ (v \\ t)" by (metis assms(1) 2 composite_ofE con_sym cong_subst_left(2) cube) thus "u \\ (v \\ t) \ (w \\ v) \\ (t \\ v)" using assms by (metis composite_of_def con_implies_arr(2) cong_subst_left(2) prfx_implies_con arr_resid_iff_con cube) qed qed lemma con_composite_of_iff: assumes "composite_of t u v" shows "w \ v \ w \\ t \ u" by (meson arr_resid_iff_con assms composite_ofE con_def con_implies_arr(1) con_sym_ax cong_subst_right(1) resid_composite_of(2) resid_reflects_con) definition composable where "composable t u \ \v. composite_of t u v" lemma composableD [dest]: assumes "composable t u" shows "arr t" and "arr u" and "targets t = sources u" using assms arr_composite_of arr_iff_has_source composable_def sources_composite_of arr_composite_of arr_iff_has_target composable_def targets_composite_of apply auto[2] by (metis assms composable_def composite_ofE con_prfx_composite_of(1) con_sym cong_implies_coinitial coinitial_iff sources_resid) lemma composable_imp_seq: assumes "composable t u" shows "seq t u" using assms by blast lemma bounded_imp_con: assumes "composite_of t u v" and "composite_of t' u' v" shows "con t t'" by (meson assms composite_of_def con_prfx_composite_of prfx_implies_con arr_resid_iff_con con_implies_arr(2)) lemma composite_of_cancel_left: assumes "composite_of t u v" and "composite_of t u' v" shows "u \ u'" using assms composite_of_def cong_transitive by blast end subsubsection "RTS with Composites" locale rts_with_composites = rts + assumes has_composites: "seq t u \ composable t u" begin lemma composable_iff_seq: shows "composable g f \ seq g f" using composable_imp_seq has_composites by blast lemma obtains_composite_of: assumes "seq g f" obtains h where "composite_of g f h" using assms has_composites composable_def by blast lemma diamond_commutes_upto_cong: assumes "composite_of t (u \\ t) v" and "composite_of u (t \\ u) v'" shows "v \ v'" using assms cube ide_backward_stable prfx_transitive by (elim composite_ofE) metis end subsection "Joins of Transitions" context rts begin text \ Transition \v\ is a \emph{join} of \u\ and \v\ when \v\ is the diagonal of the square formed by \u\, \v\, and their residuals. As was the case for composites, joins in an RTS are not unique, but they are unique up to congruence. \ definition join_of where "join_of t u v \ composite_of t (u \\ t) v \ composite_of u (t \\ u) v" lemma join_ofI [intro]: assumes "composite_of t (u \\ t) v" and "composite_of u (t \\ u) v" shows "join_of t u v" using assms join_of_def by simp lemma join_ofE [elim]: assumes "join_of t u v" and "\composite_of t (u \\ t) v; composite_of u (t \\ u) v\ \ T" shows T using assms join_of_def by simp definition joinable where "joinable t u \ \v. join_of t u v" lemma joinable_implies_con: assumes "joinable t u" shows "t \ u" by (meson assms bounded_imp_con join_of_def joinable_def) lemma joinable_implies_coinitial: assumes "joinable t u" shows "coinitial t u" using assms by (simp add: con_imp_coinitial joinable_implies_con) lemma join_of_un_upto_cong: assumes "join_of t u v" and "join_of t u v'" shows "v \ v'" using assms join_of_def composite_of_unq_upto_cong by auto lemma join_of_symmetric: assumes "join_of t u v" shows "join_of u t v" using assms join_of_def by simp lemma join_of_arr_self: assumes "arr t" shows "join_of t t t" by (meson assms composite_of_arr_ide ideE join_of_def prfx_reflexive) lemma join_of_arr_src: assumes "arr t" and "a \ sources t" shows "join_of a t t" and "join_of t a t" proof - show "join_of a t t" by (meson assms composite_of_arr_target composite_of_def composite_of_source_arr join_of_def prfx_transitive resid_source_in_targets) thus "join_of t a t" using join_of_symmetric by blast qed lemma sources_join_of: assumes "join_of t u v" shows "sources t = sources v" and "sources u = sources v" using assms join_of_def sources_composite_of by blast+ lemma targets_join_of: assumes "join_of t u v" shows "targets (t \\ u) = targets v" and "targets (u \\ t) = targets v" using assms join_of_def targets_composite_of by blast+ lemma join_of_resid: assumes "join_of t u w" and "con v w" shows "join_of (t \\ v) (u \\ v) (w \\ v)" using assms con_sym cube join_of_def resid_composite_of(4) by fastforce lemma con_with_join_of_iff: assumes "join_of t u w" shows "u \ v \ v \\ u \ t \\ u \ w \ v" and "w \ v \ t \ v \ v \\ t \ u \\ t" proof - have *: "t \ v \ v \\ t \ u \\ t \ u \ v \ v \\ u \ t \\ u" by (metis arr_resid_iff_con con_implies_arr(1) con_sym cube) show "u \ v \ v \\ u \ t \\ u \ w \ v" by (meson assms con_composite_of_iff con_sym join_of_def) show "w \ v \ t \ v \ v \\ t \ u \\ t" by (meson assms con_prfx_composite_of join_of_def resid_composite_of(2)) qed end subsubsection "RTS with Joins" locale rts_with_joins = rts + assumes has_joins: "t \ u \ joinable t u" subsection "Joins and Composites in a Weakly Extensional RTS" context weakly_extensional_rts begin lemma src_composite_of: assumes "composite_of u t v" shows "src v = src u" using assms by (metis con_imp_eq_src con_prfx_composite_of(1)) lemma trg_composite_of: assumes "composite_of u t v" shows "trg v = trg t" by (metis arr_composite_of arr_has_un_target arr_iff_has_target assms targets_composite_of trg_in_targets) lemma src_join_of: assumes "join_of t u v" shows "src t = src v" and "src u = src v" by (metis assms join_ofE src_composite_of)+ lemma trg_join_of: assumes "join_of t u v" shows "trg (t \\ u) = trg v" and "trg (u \\ t) = trg v" by (metis assms join_of_def trg_composite_of)+ end subsection "Joins and Composites in an Extensional RTS" context extensional_rts begin lemma composite_of_unique: assumes "composite_of t u v" and "composite_of t u v'" shows "v = v'" using assms composite_of_unq_upto_cong extensional by fastforce text \ Here we define composition of transitions. Note that we compose transitions in diagram order, rather than in the order used for function composition. This may eventually lead to confusion, but here (unlike in the case of a category) transitions are typically not functions, so we don't have the constraint of having to conform to the order of function application and composition, and diagram order seems more natural. \ definition comp (infixl "\" 55) where "t \ u \ if composable t u then THE v. composite_of t u v else null" lemma comp_is_composite_of: assumes "composite_of t u v" shows "composite_of t u (t \ u)" and "t \ u = v" proof - show "composite_of t u (t \ u)" using assms comp_def composite_of_unique the1I2 [of "composite_of t u" "composite_of t u"] composable_def by metis thus "t \ u = v" using assms composite_of_unique by simp qed lemma comp_null [simp]: shows "null \ t = null" and "t \ null = null" by (meson composableD not_arr_null comp_def)+ lemma composable_iff_arr_comp: shows "composable t u \ arr (t \ u)" by (metis arr_composite_of comp_is_composite_of(2) composable_def comp_def not_arr_null) lemma composable_iff_comp_not_null: shows "composable t u \ t \ u \ null" by (metis composable_iff_arr_comp comp_def not_arr_null) lemma comp_src_arr [simp]: assumes "arr t" and "src t = a" shows "a \ t = t" using assms comp_is_composite_of(2) composite_of_source_arr src_in_sources by blast lemma comp_arr_trg [simp]: assumes "arr t" and "trg t = b" shows "t \ b = t" using assms comp_is_composite_of(2) composite_of_arr_target trg_in_targets by blast lemma comp_ide_self: assumes "ide a" shows "a \ a = a" using assms comp_is_composite_of(2) composite_of_ide_self by fastforce lemma arr_comp [intro, simp]: assumes "composable t u" shows "arr (t \ u)" using assms composable_iff_arr_comp by blast lemma trg_comp [simp]: assumes "composable t u" shows "trg (t \ u) = trg u" by (metis arr_has_un_target assms comp_is_composite_of(2) composable_def composable_imp_seq arr_iff_has_target seq_def targets_composite_of trg_in_targets) lemma src_comp [simp]: assumes "composable t u" shows "src (t \ u) = src t" using assms comp_is_composite_of arr_iff_has_source sources_composite_of src_def composable_def by auto lemma con_comp_iff: shows "w \ t \ u \ composable t u \ w \\ t \ u" by (meson comp_is_composite_of(1) con_composite_of_iff con_sym con_implies_arr(2) composable_def composable_iff_arr_comp) lemma con_compI [intro]: assumes "composable t u" and "w \\ t \ u" shows "w \ t \ u" and "t \ u \ w" using assms con_comp_iff con_sym by blast+ lemma resid_comp: assumes "t \ u \ w" shows "w \\ (t \ u) = (w \\ t) \\ u" and "(t \ u) \\ w = (t \\ w) \ (u \\ (w \\ t))" proof - have 1: "composable t u" using assms composable_iff_comp_not_null by force show "w \\ (t \ u) = (w \\ t) \\ u" using 1 by (meson assms cong_char composable_def resid_composite_of(3) comp_is_composite_of(1)) show "(t \ u) \\ w = (t \\ w) \ (u \\ (w \\ t))" using assms 1 composable_def comp_is_composite_of(2) resid_composite_of by metis qed lemma prfx_decomp: assumes "t \ u" shows "t \ (u \\ t) = u" by (meson assms arr_resid_iff_con comp_is_composite_of(2) composite_of_def con_sym cong_reflexive prfx_implies_con) lemma prfx_comp: assumes "arr u" and "t \ v = u" shows "t \ u" by (metis assms comp_is_composite_of(2) composable_def composable_iff_arr_comp composite_of_def) lemma comp_eqI: assumes "t \ v" and "u = v \\ t" shows "t \ u = v" by (metis assms prfx_decomp) lemma comp_assoc: assumes "composable (t \ u) v" shows "t \ (u \ v) = (t \ u) \ v" proof - have 1: "t \ (t \ u) \ v" by (meson assms composable_iff_arr_comp composableD prfx_comp prfx_transitive) moreover have "((t \ u) \ v) \\ t = u \ v" proof - have "((t \ u) \ v) \\ t = ((t \ u) \\ t) \ (v \\ (t \\ (t \ u)))" by (meson assms calculation con_sym prfx_implies_con resid_comp(2)) also have "... = u \ v" proof - have 2: "(t \ u) \\ t = u" by (metis assms comp_is_composite_of(2) composable_def composable_iff_arr_comp composable_imp_seq composite_of_def extensional seqE) moreover have "v \\ (t \\ (t \ u)) = v" using assms by (meson 1 con_comp_iff con_sym composable_imp_seq resid_arr_ide prfx_implies_con prfx_comp seqE) ultimately show ?thesis by simp qed finally show ?thesis by blast qed ultimately show "t \ (u \ v) = (t \ u) \ v" by (metis comp_eqI) qed text \ We note the following assymmetry: \composable (t \ u) v \ composable u v\ is true, but \composable t (u \ v) \ composable t u\ is not. \ lemma comp_cancel_left: assumes "arr (t \ u)" and "t \ u = t \ v" shows "u = v" using assms by (metis composable_def composable_iff_arr_comp composite_of_cancel_left extensional comp_is_composite_of(2)) lemma comp_resid_prfx [simp]: assumes "arr (t \ u)" shows "(t \ u) \\ t = u" using assms by (metis comp_cancel_left comp_eqI prfx_comp) lemma bounded_imp_con\<^sub>E: assumes "t \ u \ t' \ u'" shows "t \ t'" by (metis arr_resid_iff_con assms con_comp_iff con_implies_arr(2) prfx_implies_con con_sym) lemma join_of_unique: assumes "join_of t u v" and "join_of t u v'" shows "v = v'" using assms join_of_def composite_of_unique by blast definition join (infix "\" 52) where "t \ u \ if joinable t u then THE v. join_of t u v else null" lemma join_is_join_of: assumes "joinable t u" shows "join_of t u (t \ u)" using assms joinable_def join_def join_of_unique the1I2 [of "join_of t u" "join_of t u"] by force lemma joinable_iff_arr_join: shows "joinable t u \ arr (t \ u)" by (metis cong_char join_is_join_of join_of_un_upto_cong not_arr_null join_def) lemma joinable_iff_join_not_null: shows "joinable t u \ t \ u \ null" by (metis join_def joinable_iff_arr_join not_arr_null) lemma join_sym: assumes "t \ u \ null" shows "t \ u = u \ t" using assms by (meson join_def join_is_join_of join_of_symmetric join_of_unique joinable_def) lemma src_join: assumes "joinable t u" shows "src (t \ u) = src t" using assms by (metis con_imp_eq_src con_prfx_composite_of(1) join_is_join_of join_of_def) lemma trg_join: assumes "joinable t u" shows "trg (t \ u) = trg (t \\ u)" using assms by (metis arr_resid_iff_con join_is_join_of joinable_iff_arr_join joinable_implies_con in_targetsE src_eqI targets_join_of(1) trg_in_targets) lemma resid_join\<^sub>E [simp]: assumes "joinable t u" and "v \ t \ u" shows "v \\ (t \ u) = (v \\ u) \\ (t \\ u)" and "v \\ (t \ u) = (v \\ t) \\ (u \\ t)" and "(t \ u) \\ v = (t \\ v) \ (u \\ v)" proof - show 1: "v \\ (t \ u) = (v \\ u) \\ (t \\ u)" by (meson assms con_sym join_of_def resid_composite_of(3) extensional join_is_join_of) show "v \\ (t \ u) = (v \\ t) \\ (u \\ t)" by (metis "1" cube) show "(t \ u) \\ v = (t \\ v) \ (u \\ v)" using assms joinable_def join_of_resid join_is_join_of extensional by (meson join_of_unique) qed lemma join_eqI: assumes "t \ v" and "u \ v" and "v \\ u = t \\ u" and "v \\ t = u \\ t" shows "t \ u = v" using assms composite_of_def cube ideE join_of_def joinable_def join_of_unique join_is_join_of trg_def by metis lemma comp_join: assumes "joinable (t \ u) (t \ u')" shows "composable t (u \ u')" and "t \ (u \ u') = t \ u \ t \ u'" proof - have "t \ t \ u \ t \ u'" using assms by (metis composable_def composite_of_def join_of_def join_is_join_of joinable_implies_con prfx_transitive comp_is_composite_of(2) con_comp_iff) moreover have "(t \ u \ t \ u') \\ t = u \ u'" by (metis arr_resid_iff_con assms calculation comp_resid_prfx con_implies_arr(2) joinable_implies_con resid_join\<^sub>E(3) con_implies_arr(1) ide_implies_arr) ultimately show "t \ (u \ u') = t \ u \ t \ u'" by (metis comp_eqI) thus "composable t (u \ u')" by (metis assms joinable_iff_join_not_null comp_def) qed lemma join_src: assumes "arr t" shows "src t \ t = t" using assms joinable_def join_of_arr_src join_is_join_of join_of_unique src_in_sources by meson lemma join_self: assumes "arr t" shows "t \ t = t" using assms joinable_def join_of_arr_self join_is_join_of join_of_unique by blast lemma arr_prfx_join_self: assumes "joinable t u" shows "t \ t \ u" using assms by (meson composite_of_def join_is_join_of join_of_def) text \ We note that it is not the case that the existence of either of \t \ (u \ v)\ or \(t \ u) \ v\ implies that of the other. For example, if \(t \ u) \ v \ null\, then it is not necessarily the case that \u \ v \ null\. \ end subsubsection "Extensional RTS with Joins" locale extensional_rts_with_joins = rts_with_joins + extensional_rts begin lemma joinable_iff_con: shows "joinable t u \ t \ u" by (meson has_joins joinable_implies_con) lemma src_join\<^sub>E\<^sub>J [simp]: assumes "t \ u" shows "src (t \ u) = src t" using assms by (meson has_joins src_join) lemma trg_join\<^sub>E\<^sub>J: assumes "t \ u" shows "trg (t \ u) = trg (t \\ u)" using assms by (meson has_joins trg_join) lemma resid_join\<^sub>E\<^sub>J [simp]: assumes "t \ u" and "v \ t \ u" shows "v \\ (t \ u) = (v \\ t) \\ (u \\ t)" and "(t \ u) \\ v = (t \\ v) \ (u \\ v)" using assms has_joins resid_join\<^sub>E by blast+ lemma join_assoc: shows "t \ (u \ v) = (t \ u) \ v" proof - have *: "\t u v. con (t \ u) v \ t \ (u \ v) = (t \ u) \ v" proof - fix t u v assume 1: "con (t \ u) v" have vt_ut: "v \\ t \ u \\ t" using 1 by (metis con_implies_arr(1) con_with_join_of_iff(2) join_is_join_of not_arr_null join_def) have tv_uv: "t \\ v \ u \\ v" using vt_ut cube con_sym by (metis arr_resid_iff_con) have 2: "(t \ u) \ v = (t \ (u \\ t)) \ (v \\ (t \ (u \\ t)))" using 1 by (metis comp_is_composite_of(2) con_implies_arr(1) has_joins join_is_join_of join_of_def joinable_iff_arr_join) also have "... = t \ ((u \\ t) \ (v \\ (t \ (u \\ t))))" using 1 by (metis calculation has_joins joinable_iff_join_not_null comp_assoc comp_def) also have "... = t \ ((u \\ t) \ ((v \\ t) \\ (u \\ t)))" using 1 by (metis 2 comp_null(2) con_compI(2) con_comp_iff has_joins resid_comp(1) conI joinable_iff_join_not_null) also have "... = t \ ((v \\ t) \ (u \\ t))" by (metis vt_ut comp_is_composite_of(2) has_joins join_of_def join_is_join_of) also have "... = t \ ((u \\ t) \ (v \\ t))" using join_sym by metis also have "... = t \ ((u \ v) \\ t)" by (metis tv_uv vt_ut con_implies_arr(2) con_sym con_with_join_of_iff(1) has_joins join_is_join_of arr_resid_iff_con resid_join\<^sub>E(3)) also have "... = t \ (u \ v)" by (metis comp_is_composite_of(2) comp_null(2) conI has_joins join_is_join_of join_of_def joinable_iff_join_not_null) finally show "t \ (u \ v) = (t \ u) \ v" by simp qed thus ?thesis by (metis (full_types) has_joins joinable_iff_join_not_null joinable_implies_con con_sym) qed lemma join_is_lub: assumes "t \ v" and "u \ v" shows "t \ u \ v" proof - have "(t \ u) \\ v = (t \\ v) \ (u \\ v)" using assms resid_join\<^sub>E(3) [of t u v] by (metis arr_prfx_join_self con_target con_sym join_assoc joinable_iff_con joinable_iff_join_not_null prfx_implies_con resid_reflects_con) also have "... = trg v \ trg v" using assms by (metis ideE prfx_implies_con src_resid trg_ide) also have "... = trg v" by (metis assms(2) ide_iff_src_self ide_implies_arr join_self prfx_implies_con src_resid) finally have "(t \ u) \\ v = trg v" by blast moreover have "ide (trg v)" using assms by (metis con_implies_arr(2) prfx_implies_con cong_char trg_def) ultimately show ?thesis by simp qed end subsubsection "Extensional RTS with Composites" text \ If an extensional RTS is assumed to have composites for all composable pairs of transitions, then the ``semantic'' property of transitions being composable can be replaced by the ``syntactic'' property of transitions being sequential. This results in simpler statements of a number of properties. \ locale extensional_rts_with_composites = rts_with_composites + extensional_rts begin lemma seq_implies_arr_comp: assumes "seq t u" shows "arr (t \ u)" using assms by (meson composable_iff_arr_comp composable_iff_seq) lemma arr_comp\<^sub>E\<^sub>C [intro, simp]: assumes "arr t" and "arr u" and "trg t = src u" shows "arr (t \ u)" using assms by (simp add: seq_implies_arr_comp) lemma arr_compE\<^sub>E\<^sub>C [elim]: assumes "arr (t \ u)" and "\arr t; arr u; trg t = src u\ \ T" shows T using assms composable_iff_arr_comp composable_iff_seq by blast lemma trg_comp\<^sub>E\<^sub>C [simp]: assumes "seq t u" shows "trg (t \ u) = trg u" by (meson assms has_composites trg_comp) lemma src_comp\<^sub>E\<^sub>C [simp]: assumes "seq t u" shows "src (t \ u) = src t" using assms src_comp has_composites by simp lemma con_comp_iff\<^sub>E\<^sub>C [simp]: shows "w \ t \ u \ seq t u \ u \ w \\ t" and "t \ u \ w \ seq t u \ u \ w \\ t" using composable_iff_seq con_comp_iff con_sym by meson+ lemma comp_assoc\<^sub>E\<^sub>C: shows "t \ (u \ v) = (t \ u) \ v" apply (cases "seq t u") apply (metis arr_comp comp_assoc comp_def not_arr_null arr_compE\<^sub>E\<^sub>C arr_comp\<^sub>E\<^sub>C seq_implies_arr_comp trg_comp\<^sub>E\<^sub>C) by (metis comp_def composable_iff_arr_comp seqI\<^sub>W\<^sub>E src_comp arr_compE\<^sub>E\<^sub>C) lemma diamond_commutes: shows "t \ (u \\ t) = u \ (t \\ u)" proof (cases "t \ u") show "\ t \ u \ ?thesis" by (metis comp_null(2) conI con_sym) assume con: "t \ u" have "(t \ (u \\ t)) \\ u = (t \\ u) \ ((u \\ t) \\ (u \\ t))" using con by (metis (no_types, lifting) arr_resid_iff_con con_compI(2) con_implies_arr(1) resid_comp(2) con_imp_arr_resid con_sym comp_def arr_comp\<^sub>E\<^sub>C src_resid conI) moreover have "u \ t \ (u \\ t)" by (metis arr_resid_iff_con calculation con cong_reflexive comp_arr_trg resid_arr_self resid_comp(1) trg_resid_sym) ultimately show ?thesis by (metis comp_eqI con comp_arr_trg resid_arr_self arr_resid trg_resid_sym) qed lemma mediating_transition: assumes "t \ v = u \ w" shows "v \\ (u \\ t) = w \\ (t \\ u)" proof (cases "seq t v") assume 1: "seq t v" hence 2: "arr (u \ w)" using assms by (metis arr_comp\<^sub>E\<^sub>C seqE\<^sub>W\<^sub>E) have 3: "v \\ (u \\ t) = ((t \ v) \\ t) \\ (u \\ t)" by (metis "1" comp_is_composite_of(1) composite_of_def obtains_composite_of extensional) also have "... = (t \ v) \\ (t \ (u \\ t))" by (metis (no_types, lifting) "2" assms con_comp_iff\<^sub>E\<^sub>C(2) con_imp_eq_src con_implies_arr(2) con_sym comp_resid_prfx prfx_comp resid_comp(1) arr_compE\<^sub>E\<^sub>C arr_comp\<^sub>E\<^sub>C prfx_implies_con) also have "... = (u \ w) \\ (u \ (t \\ u))" using assms diamond_commutes by presburger also have "... = ((u \ w) \\ u) \\ (t \\ u)" by (metis 3 assms calculation cube) also have "... = w \\ (t \\ u)" using 2 by simp finally show ?thesis by blast next assume 1: "\ seq t v" have "v \\ (u \\ t) = null" using 1 by (metis (mono_tags, lifting) arr_resid_iff_con coinitial_iff\<^sub>W\<^sub>E con_imp_coinitial seqI\<^sub>W\<^sub>E src_resid conI) also have "... = w \\ (t \\ u)" by (metis (no_types, lifting) "1" arr_comp\<^sub>E\<^sub>C assms composable_imp_seq con_imp_eq_src con_implies_arr(1) con_implies_arr(2) comp_def not_arr_null conI src_resid) finally show ?thesis by blast qed lemma induced_arrow: assumes "seq t u" and "t \ u = t' \ u'" shows "(t' \\ t) \ (u \\ (t' \\ t)) = u" and "(t \\ t') \ (u \\ (t' \\ t)) = u'" and "(t' \\ t) \ v = u \ v = u \\ (t' \\ t)" apply (metis assms comp_eqI arr_compE\<^sub>E\<^sub>C prfx_comp resid_comp(1) arr_resid_iff_con seq_implies_arr_comp) apply (metis assms comp_resid_prfx arr_compE\<^sub>E\<^sub>C resid_comp(2) arr_resid_iff_con seq_implies_arr_comp) by (metis assms(1) comp_resid_prfx seq_def) text \ If an extensional RTS has composites, then it automatically has joins. \ sublocale extensional_rts_with_joins proof fix t u assume con: "t \ u" have 1: "con u (t \ (u \\ t))" using con_compI(1) [of t "u \\ t" u] by (metis con con_implies_arr(1) con_sym diamond_commutes prfx_implies_con arr_resid prfx_comp src_resid arr_comp\<^sub>E\<^sub>C) have "t \ u = t \ (u \\ t)" proof (intro join_eqI) show "t \ t \ (u \\ t)" by (metis 1 composable_def comp_is_composite_of(2) composite_of_def con_comp_iff) moreover show 2: "u \ t \ (u \\ t)" using 1 arr_resid con con_sym prfx_reflexive resid_comp(1) by metis moreover show "(t \ (u \\ t)) \\ u = t \\ u" using 1 diamond_commutes induced_arrow(2) resid_comp(2) by force ultimately show "(t \ (u \\ t)) \\ t = u \\ t" by (metis con_comp_iff\<^sub>E\<^sub>C(1) con_sym prfx_implies_con resid_comp(2) induced_arrow(1)) qed thus "joinable t u" by (metis "1" con_implies_arr(2) joinable_iff_join_not_null not_arr_null) qed lemma join_expansion: assumes "t \ u" shows "t \ u = t \ (u \\ t)" and "seq t (u \\ t)" proof - show "t \ u = t \ (u \\ t)" by (metis assms comp_is_composite_of(2) has_joins join_is_join_of join_of_def) thus "seq t (u \\ t)" by (meson assms composable_def composable_iff_seq has_joins join_is_join_of join_of_def) qed lemma join3_expansion: assumes "t \ u" and "t \ v" and "u \ v" shows "(t \ u) \ v = (t \ (u \\ t)) \ ((v \\ t) \\ (u \\ t))" proof (cases "v \\ t \ u \\ t") show "\ v \\ t \ u \\ t \ ?thesis" by (metis assms(1) comp_null(2) join_expansion(1) joinable_implies_con resid_comp(1) join_def conI) assume 1: "v \\ t \ u \\ t " have "(t \ u) \ v = (t \ u) \ (v \\ (t \ u))" by (metis comp_null(1) diamond_commutes ex_un_null join_expansion(1) joinable_implies_con null_is_zero(2) join_def conI) also have "... = (t \ (u \\ t)) \ (v \\ (t \ u))" using join_expansion [of t u] assms(1) by presburger also have "... = (t \ (u \\ t)) \ ((v \\ u) \\ (t \\ u))" using assms 1 join_of_resid(1) [of t u v] cube [of v t u] by (metis con_compI(2) con_implies_arr(2) join_expansion(1) not_arr_null resid_comp(1) con_sym comp_def src_resid arr_comp\<^sub>E\<^sub>C) also have "... = (t \ (u \\ t)) \ ((v \\ t) \\ (u \\ t))" by (metis cube) finally show ?thesis by blast qed lemma resid_common_prefix: assumes "t \ u \ t \ v" shows "(t \ u) \\ (t \ v) = u \\ v" using assms by (metis con_comp_iff con_sym con_comp_iff\<^sub>E\<^sub>C(2) con_implies_arr(2) induced_arrow(1) resid_comp(1) resid_comp(2) residuation.arr_resid_iff_con residuation_axioms) end subsection "Confluence" text \ An RTS is \emph{confluent} if every coinitial pair of transitions is consistent. \ locale confluent_rts = rts + assumes confluence: "coinitial t u \ con t u" section "Simulations" text \ - \emph{Simulations} are morphism of residuated transition systems. + \emph{Simulations} are morphisms of residuated transition systems. They are assumed to preserve consistency and residuation. \ locale simulation = A: rts A + B: rts B for A :: "'a resid" (infixr "\\\<^sub>A" 70) and B :: "'b resid" (infixr "\\\<^sub>B" 70) and F :: "'a \ 'b" + assumes extensional: "\ A.arr t \ F t = B.null" and preserves_con [simp]: "A.con t u \ B.con (F t) (F u)" and preserves_resid [simp]: "A.con t u \ F (t \\\<^sub>A u) = F t \\\<^sub>B F u" begin lemma preserves_reflects_arr [iff]: shows "B.arr (F t) \ A.arr t" by (metis A.arr_def B.con_implies_arr(2) B.not_arr_null extensional preserves_con) lemma preserves_ide [simp]: assumes "A.ide a" shows "B.ide (F a)" by (metis A.ideE assms preserves_con preserves_resid B.ideI) lemma preserves_sources: shows "F ` A.sources t \ B.sources (F t)" using A.sources_def B.sources_def preserves_con preserves_ide by auto lemma preserves_targets: shows "F ` A.targets t \ B.targets (F t)" by (metis A.arrE B.arrE A.sources_resid B.sources_resid equals0D image_subset_iff A.arr_iff_has_target preserves_reflects_arr preserves_resid preserves_sources) lemma preserves_trg: assumes "A.arr t" shows "F (A.trg t) = B.trg (F t)" using assms A.trg_def B.trg_def by auto lemma preserves_composites: assumes "A.composite_of t u v" shows "B.composite_of (F t) (F u) (F v)" using assms by (metis A.composite_ofE A.prfx_implies_con B.composite_of_def preserves_ide preserves_resid A.con_sym) lemma preserves_joins: assumes "A.join_of t u v" shows "B.join_of (F t) (F u) (F v)" using assms A.join_of_def B.join_of_def A.joinable_def by (metis A.joinable_implies_con preserves_composites preserves_resid) lemma preserves_prfx: assumes "A.prfx t u" shows "B.prfx (F t) (F u)" using assms by (metis A.prfx_implies_con preserves_ide preserves_resid) lemma preserves_cong: assumes "A.cong t u" shows "B.cong (F t) (F u)" using assms preserves_prfx by simp end subsection "Identity Simulation" locale identity_simulation = rts begin abbreviation map where "map \ \t. if arr t then t else null" sublocale simulation resid resid map using con_implies_arr con_sym arr_resid_iff_con by unfold_locales auto end subsection "Composite of Simulations" lemma simulation_comp: assumes "simulation A B F" and "simulation B C G" shows "simulation A C (G o F)" proof - interpret F: simulation A B F using assms(1) by auto interpret G: simulation B C G using assms(2) by auto show "simulation A C (G o F)" using F.extensional G.extensional by unfold_locales auto qed locale composite_simulation = F: simulation A B F + G: simulation B C G for A :: "'a resid" and B :: "'b resid" and C :: "'c resid" and F :: "'a \ 'b" and G :: "'b \ 'c" begin abbreviation map where "map \ G o F" sublocale simulation A C map using simulation_comp F.simulation_axioms G.simulation_axioms by blast lemma is_simulation: shows "simulation A C map" .. end subsection "Simulations into a Weakly Extensional RTS" locale simulation_to_weakly_extensional_rts = simulation + B: weakly_extensional_rts B begin lemma preserves_src: shows "a \ A.sources t \ B.src (F t) = F a" by (metis equals0D image_subset_iff B.arr_iff_has_source preserves_sources B.arr_has_un_source B.src_in_sources) lemma preserves_trg: shows "b \ A.targets t \ B.trg (F t) = F b" by (metis equals0D image_subset_iff B.arr_iff_has_target preserves_targets B.arr_has_un_target B.trg_in_targets) end subsection "Simulations into an Extensional RTS" locale simulation_to_extensional_rts = simulation + B: extensional_rts B begin lemma preserves_comp: assumes "A.composite_of t u v" shows "F v = B.comp (F t) (F u)" using assms by (metis preserves_composites B.comp_is_composite_of(2)) lemma preserves_join: assumes "A.join_of t u v" shows "F v = B.join (F t) (F u)" using assms preserves_joins by (meson B.join_is_join_of B.join_of_unique B.joinable_def) end subsection "Simulations between Extensional RTS's" locale simulation_between_extensional_rts = simulation_to_extensional_rts + A: extensional_rts A begin lemma preserves_src: shows "B.src (F t) = F (A.src t)" by (metis A.arr_src_iff_arr A.src_in_sources extensional image_subset_iff preserves_reflects_arr preserves_sources B.arr_has_un_source B.src_def B.src_in_sources) lemma preserves_trg: shows "B.trg (F t) = F (A.trg t)" by (metis A.arr_trg_iff_arr A.residuation_axioms A.trg_def B.null_is_zero(2) B.trg_def extensional preserves_resid residuation.arrE) lemma preserves_comp: assumes "A.composable t u" shows "F (A.comp t u) = B.comp (F t) (F u)" using assms by (metis A.arr_comp A.comp_resid_prfx A.composableD(2) A.not_arr_null A.prfx_comp A.residuation_axioms B.comp_eqI preserves_prfx preserves_resid residuation.conI) lemma preserves_join: assumes "A.joinable t u" shows "F (A.join t u) = B.join (F t) (F u)" using assms by (meson A.join_is_join_of B.joinable_def preserves_joins B.join_is_join_of B.join_of_unique) end subsection "Transformations" text \ A \emph{transformation} is a morphism of simulations, analogously to how a natural transformation is a morphism of functors, except the normal commutativity condition for that ``naturality squares'' is replaced by the requirement that the arrows at the apex of such a square are given by residuation of the arrows at the base. If the codomain RTS is extensional, then this condition implies the commutativity of the square with respect to composition, as would be the case for a natural transformation between functors. The proper way to define a transformation when the domain and codomain are general RTS's is not yet clear to me. However, if the domain and codomain are weakly extensional, then we have unique sources and targets, so there is no problem. The definition below is limited to that case. I do not make any attempt here to develop facts about transformations. My main reason for including this definition here is so that in the subsequent application to the \\\-calculus, I can exhibit \\\-reduction as an example of a transformation. \ locale transformation = A: weakly_extensional_rts A + B: weakly_extensional_rts B + F: simulation A B F + G: simulation A B G for A :: "'a resid" (infixr "\\\<^sub>A" 70) and B :: "'b resid" (infixr "\\\<^sub>B" 70) and F :: "'a \ 'b" and G :: "'a \ 'b" and \ :: "'a \ 'b" + assumes extensional: "\ A.arr f \ \ f = B.null" and preserves_src: "A.ide f \ B.src (\ f) = F (A.src f)" and preserves_trg: "A.ide f \ B.trg (\ f) = G (A.trg f)" and naturality1: "A.arr f \ \ (A.src f) \\\<^sub>B F f = \ (A.trg f)" and naturality2: "A.arr f \ F f \\\<^sub>B \ (A.src f) = G f" and naturality3: "A.arr f \ B.join_of (\ (A.src f)) (F f) (\ f)" section "Normal Sub-RTS's and Congruence" text \ We now develop a general quotient construction on an RTS. We define a \emph{normal sub-RTS} of an RTS to be a collection of transitions \\\ having certain ``local'' closure properties. A normal sub-RTS induces an equivalence relation \\\<^sub>0\, which we call \emph{semi-congruence}, by defining \t \\<^sub>0 u\ to hold exactly when \t \ u\ and \u \ t\ are both in \\\. This relation generalizes the relation \\\ defined for an arbitrary RTS, in the sense that \\\ is obtained when \\\ consists of all and only the identity transitions. However, in general the relation \\\<^sub>0\ is fully substitutive only in the left argument position of residuation; for the right argument position, a somewhat weaker property is satisfied. We then coarsen \\\<^sub>0\ to a relation \\\, by defining \t \ u\ to hold exactly when \t\ and \u\ can be transported by residuation along transitions in \\\ to a common source, in such a way that the residuals are related by \\\<^sub>0\. To obtain full substitutivity of \\\ with respect to residuation, we need to impose an additional condition on \\\. This condition, which we call \emph{coherence}, states that transporting a transition \t\ along parallel transitions \u\ and \v\ in \\\ always yields residuals \t \ u\ and \u \ t\ that are related by \\\<^sub>0\. We show that, under the assumption of coherence, the relation \\\ is fully substitutive, and the quotient of the original RTS by this relation is an extensional RTS which has the \\\-connected components of the original RTS as identities. Although the coherence property has a somewhat \emph{ad hoc} feel to it, we show that, in the context of the other conditions assumed for \\\, coherence is in fact equivalent to substitutivity for \\\. \ subsection "Normal Sub-RTS's" locale normal_sub_rts = R: rts + fixes \ :: "'a set" assumes elements_are_arr: "t \ \ \ R.arr t" and ide_closed: "R.ide a \ a \ \" and forward_stable: "\ u \ \; R.coinitial t u \ \ u \\ t \ \" and backward_stable: "\ u \ \; t \\ u \ \ \ \ t \ \" and composite_closed_left: "\ u \ \; R.seq u t \ \ \v. R.composite_of u t v" and composite_closed_right: "\ u \ \; R.seq t u \ \ \v. R.composite_of t u v" begin lemma prfx_closed: assumes "u \ \" and "R.prfx t u" shows "t \ \" using assms backward_stable ide_closed by blast lemma composite_closed: assumes "t \ \" and "u \ \" and "R.composite_of t u v" shows "v \ \" using assms backward_stable R.composite_of_def prfx_closed by blast lemma factor_closed: assumes "R.composite_of t u v" and "v \ \" shows "t \ \" and "u \ \" apply (metis assms R.composite_of_def prfx_closed) by (meson assms R.composite_of_def R.con_imp_coinitial forward_stable prfx_closed R.prfx_implies_con) lemma resid_along_elem_preserves_con: assumes "t \ t'" and "R.coinitial t u" and "u \ \" shows "t \\ u \ t' \\ u" proof - have "R.coinitial (t \\ t') (u \\ t')" by (metis assms R.arr_resid_iff_con R.coinitialI R.con_imp_common_source forward_stable elements_are_arr R.con_implies_arr(2) R.sources_resid R.sources_eqI) hence "t \\ t' \ u \\ t'" by (metis assms(3) R.coinitial_iff R.con_imp_coinitial R.con_sym elements_are_arr forward_stable R.arr_resid_iff_con) thus ?thesis using assms R.cube forward_stable by fastforce qed end subsubsection "Normal Sub-RTS's of an Extensional RTS with Composites" locale normal_in_extensional_rts_with_composites = R: extensional_rts + R: rts_with_composites + normal_sub_rts begin lemma factor_closed\<^sub>E\<^sub>C: assumes "t \ u \ \" shows "t \ \" and "u \ \" using assms factor_closed by (metis R.arrE R.composable_def R.comp_is_composite_of(2) R.con_comp_iff elements_are_arr)+ lemma comp_in_normal_iff: shows "t \ u \ \ \ t \ \ \ u \ \ \ R.seq t u" by (metis R.comp_is_composite_of(2) composite_closed elements_are_arr factor_closed(1-2) R.composable_def R.has_composites R.rts_with_composites_axioms R.extensional_rts_axioms extensional_rts_with_composites.arr_compE\<^sub>E\<^sub>C extensional_rts_with_composites_def R.seqI\<^sub>W\<^sub>E) end subsection "Semi-Congruence" context normal_sub_rts begin text \ We will refer to the elements of \\\ as \emph{normal transitions}. Generalizing identity transitions to normal transitions in the definition of congruence, we obtain the notion of \emph{semi-congruence} of transitions with respect to a normal sub-RTS. \ abbreviation Cong\<^sub>0 (infix "\\<^sub>0" 50) where "t \\<^sub>0 t' \ t \\ t' \ \ \ t' \\ t \ \" lemma Cong\<^sub>0_reflexive: assumes "R.arr t" shows "t \\<^sub>0 t" using assms R.cong_reflexive ide_closed by simp lemma Cong\<^sub>0_symmetric: assumes "t \\<^sub>0 t'" shows "t' \\<^sub>0 t" using assms by simp lemma Cong\<^sub>0_transitive [trans]: assumes "t \\<^sub>0 t'" and "t' \\<^sub>0 t''" shows "t \\<^sub>0 t''" by (metis (full_types) R.arr_resid_iff_con assms backward_stable forward_stable elements_are_arr R.coinitialI R.cube R.sources_resid) lemma Cong\<^sub>0_imp_con: assumes "t \\<^sub>0 t'" shows "R.con t t'" using assms R.arr_resid_iff_con elements_are_arr by blast lemma Cong\<^sub>0_imp_coinitial: assumes "t \\<^sub>0 t'" shows "R.sources t = R.sources t'" using assms by (meson Cong\<^sub>0_imp_con R.coinitial_iff R.con_imp_coinitial) text \ Semi-congruence is preserved and reflected by residuation along normal transitions. \ lemma Resid_along_normal_preserves_Cong\<^sub>0: assumes "t \\<^sub>0 t'" and "u \ \" and "R.sources t = R.sources u" shows "t \\ u \\<^sub>0 t' \\ u" by (metis Cong\<^sub>0_imp_coinitial R.arr_resid_iff_con R.coinitialI R.coinitial_def R.cube R.sources_resid assms elements_are_arr forward_stable) lemma Resid_along_normal_reflects_Cong\<^sub>0: assumes "t \\ u \\<^sub>0 t' \\ u" and "u \ \" shows "t \\<^sub>0 t'" using assms by (metis backward_stable R.con_imp_coinitial R.cube R.null_is_zero(2) forward_stable R.conI) text \ Semi-congruence is substitutive for the left-hand argument of residuation. \ lemma Cong\<^sub>0_subst_left: assumes "t \\<^sub>0 t'" and "t \ u" shows "t' \ u" and "t \\ u \\<^sub>0 t' \\ u" proof - have 1: "t \ u \ t \ t' \ u \\ t \ t' \\ t" using assms by (metis Resid_along_normal_preserves_Cong\<^sub>0 Cong\<^sub>0_imp_con Cong\<^sub>0_reflexive R.con_sym R.null_is_zero(2) R.arr_resid_iff_con R.sources_resid R.conI) hence 2: "t' \ u \ u \\ t \ t' \\ t \ (t \\ u) \\ (t' \\ u) = (t \\ t') \\ (u \\ t') \ (t' \\ u) \\ (t \\ u) = (t' \\ t) \\ (u \\ t)" by (meson R.con_sym R.cube R.resid_reflects_con) show "t' \ u" using 2 by simp show "t \\ u \\<^sub>0 t' \\ u" using assms 1 2 by (metis R.arr_resid_iff_con R.con_imp_coinitial R.cube forward_stable) qed text \ Semi-congruence is not exactly substitutive for residuation on the right. Instead, the following weaker property is satisfied. Obtaining exact substitutivity on the right is the motivation for defining a coarser notion of congruence below. \ lemma Cong\<^sub>0_subst_right: assumes "u \\<^sub>0 u'" and "t \ u" shows "t \ u'" and "(t \\ u) \\ (u' \\ u) \\<^sub>0 (t \\ u') \\ (u \\ u')" using assms apply (meson Cong\<^sub>0_subst_left(1) R.con_sym) using assms by (metis R.sources_resid Cong\<^sub>0_imp_con Cong\<^sub>0_reflexive Resid_along_normal_preserves_Cong\<^sub>0 R.arr_resid_iff_con residuation.cube R.residuation_axioms) lemma Cong\<^sub>0_subst_Con: assumes "t \\<^sub>0 t'" and "u \\<^sub>0 u'" shows "t \ u \ t' \ u'" using assms by (meson Cong\<^sub>0_subst_left(1) Cong\<^sub>0_subst_right(1)) lemma Cong\<^sub>0_cancel_left: assumes "R.composite_of t u v" and "R.composite_of t u' v'" and "v \\<^sub>0 v'" shows "u \\<^sub>0 u'" proof - have "u \\<^sub>0 v \\ t" using assms(1) ide_closed by blast also have "v \\ t \\<^sub>0 v' \\ t" by (meson assms(1,3) Cong\<^sub>0_subst_left(2) R.composite_of_def R.con_sym R.prfx_implies_con) also have "v' \\ t \\<^sub>0 u'" using assms(2) ide_closed by blast finally show ?thesis by auto qed lemma Cong\<^sub>0_iff: shows "t \\<^sub>0 t' \ (\u u' v v'. u \ \ \ u' \ \ \ v \\<^sub>0 v' \ R.composite_of t u v \ R.composite_of t' u' v')" proof (intro iffI) show "\u u' v v'. u \ \ \ u' \ \ \ v \\<^sub>0 v' \ R.composite_of t u v \ R.composite_of t' u' v' \ t \\<^sub>0 t'" by (meson Cong\<^sub>0_transitive R.composite_of_def ide_closed prfx_closed) show "t \\<^sub>0 t' \ \u u' v v'. u \ \ \ u' \ \ \ v \\<^sub>0 v' \ R.composite_of t u v \ R.composite_of t' u' v'" by (metis Cong\<^sub>0_imp_con Cong\<^sub>0_transitive R.composite_of_def R.prfx_reflexive R.arrI R.ideE) qed lemma diamond_commutes_upto_Cong\<^sub>0: assumes "t \ u" and "R.composite_of t (u \\ t) v" and "R.composite_of u (t \\ u) v'" shows "v \\<^sub>0 v'" proof - have "v \\ v \\<^sub>0 v' \\ v \ v' \\ v' \\<^sub>0 v \\ v'" proof- have 1: "(v \\ t) \\ (u \\ t) \\<^sub>0 (v' \\ u) \\ (t \\ u)" using assms(2-3) R.cube [of v t u] by (metis R.con_target R.composite_ofE R.ide_imp_con_iff_cong ide_closed R.conI) have 2: "v \\ v \\<^sub>0 v' \\ v" proof - have "v \\ v \\<^sub>0 (v \\ t) \\ (u \\ t)" using assms R.composite_of_def ide_closed by (meson R.composite_of_unq_upto_cong R.prfx_implies_con R.resid_composite_of(3)) also have "(v \\ t) \\ (u \\ t) \\<^sub>0 (v' \\ u) \\ (t \\ u)" using 1 by simp also have "(v' \\ u) \\ (t \\ u) \\<^sub>0 (v' \\ t) \\ (u \\ t)" by (metis "1" Cong\<^sub>0_transitive R.cube) also have "(v' \\ t) \\ (u \\ t) \\<^sub>0 v' \\ v" using assms R.composite_of_def ide_closed by (metis "1" R.conI R.con_sym_ax R.cube R.null_is_zero(2) R.resid_composite_of(3)) finally show ?thesis by auto qed moreover have "v' \\ v' \\<^sub>0 v \\ v'" proof - have "v' \\ v' \\<^sub>0 (v' \\ u) \\ (t \\ u)" using assms R.composite_of_def ide_closed by (meson R.composite_of_unq_upto_cong R.prfx_implies_con R.resid_composite_of(3)) also have "(v' \\ u) \\ (t \\ u) \\<^sub>0 (v \\ t) \\ (u \\ t)" using 1 by simp also have "(v \\ t) \\ (u \\ t) \\<^sub>0 (v \\ u) \\ (t \\ u)" using R.cube [of v t u] ide_closed by (metis Cong\<^sub>0_reflexive R.arr_resid_iff_con assms(2) R.composite_of_def R.prfx_implies_con) also have "(v \\ u) \\ (t \\ u) \\<^sub>0 v \\ v'" using assms R.composite_of_def ide_closed by (metis 2 R.conI elements_are_arr R.not_arr_null R.null_is_zero(2) R.resid_composite_of(3)) finally show ?thesis by auto qed ultimately show ?thesis by blast qed thus ?thesis by (metis assms(2-3) R.composite_of_unq_upto_cong R.resid_arr_ide Cong\<^sub>0_imp_con) qed subsection "Congruence" text \ We use semi-congruence to define a coarser relation as follows. \ definition Cong (infix "\" 50) where "Cong t t' \ \u u'. u \ \ \ u' \ \ \ t \\ u \\<^sub>0 t' \\ u'" lemma CongI [intro]: assumes "u \ \" and "u' \ \" and "t \\ u \\<^sub>0 t' \\ u'" shows "Cong t t'" using assms Cong_def by auto lemma CongE [elim]: assumes "t \ t'" obtains u u' where "u \ \" and "u' \ \" and "t \\ u \\<^sub>0 t' \\ u'" using assms Cong_def by auto lemma Cong_imp_arr: assumes "t \ t'" shows "R.arr t" and "R.arr t'" using assms Cong_def by (meson R.arr_resid_iff_con R.con_implies_arr(2) R.con_sym elements_are_arr)+ lemma Cong_reflexive: assumes "R.arr t" shows "t \ t" by (metis CongI Cong\<^sub>0_reflexive assms R.con_imp_coinitial_ax ide_closed R.resid_arr_ide R.arrE R.con_sym) lemma Cong_symmetric: assumes "t \ t'" shows "t' \ t" using assms Cong_def by auto text \ The existence of composites of normal transitions is used in the following. \ lemma Cong_transitive [trans]: assumes "t \ t''" and "t'' \ t'" shows "t \ t'" proof - obtain u u'' where uu'': "u \ \ \ u'' \ \ \ t \\ u \\<^sub>0 t'' \\ u''" using assms Cong_def by blast obtain v' v'' where v'v'': "v' \ \ \ v'' \ \ \ t'' \\ v'' \\<^sub>0 t' \\ v'" using assms Cong_def by blast let ?w = "(t \\ u) \\ (v'' \\ u'')" let ?w' = "(t' \\ v') \\ (u'' \\ v'')" let ?w'' = "(t'' \\ v'') \\ (u'' \\ v'')" have w'': "?w'' = (t'' \\ u'') \\ (v'' \\ u'')" by (metis R.cube) have u''v'': "R.coinitial u'' v''" by (metis (full_types) R.coinitial_iff elements_are_arr R.con_imp_coinitial R.arr_resid_iff_con uu'' v'v'') hence v''u'': "R.coinitial v'' u''" by (meson R.con_imp_coinitial elements_are_arr forward_stable R.arr_resid_iff_con v'v'') have 1: "?w \\ ?w'' \ \" proof - have "(v'' \\ u'') \\ (t'' \\ u'') \ \" by (metis Cong\<^sub>0_transitive R.con_imp_coinitial forward_stable Cong\<^sub>0_imp_con resid_along_elem_preserves_con R.arrI R.arr_resid_iff_con u''v'' uu'' v'v'') thus ?thesis by (metis Cong\<^sub>0_subst_left(2) R.con_sym R.null_is_zero(1) uu'' w'' R.conI) qed have 2: "?w'' \\ ?w \ \" by (metis 1 Cong\<^sub>0_subst_left(2) uu'' w'' R.conI) have 3: "R.seq u (v'' \\ u'')" by (metis (full_types) 2 Cong\<^sub>0_imp_coinitial R.sources_resid Cong\<^sub>0_imp_con R.arr_resid_iff_con R.con_implies_arr(2) R.seqI uu'' R.conI) have 4: "R.seq v' (u'' \\ v'')" by (metis 1 Cong\<^sub>0_imp_coinitial Cong\<^sub>0_imp_con R.arr_resid_iff_con R.con_implies_arr(2) R.seq_def R.sources_resid v'v'' R.conI) obtain x where x: "R.composite_of u (v'' \\ u'') x" using 3 composite_closed_left uu'' by blast obtain x' where x': "R.composite_of v' (u'' \\ v'') x'" using 4 composite_closed_left v'v'' by presburger have "?w \\<^sub>0 ?w'" proof - have "?w \\<^sub>0 ?w'' \ ?w' \\<^sub>0 ?w''" using 1 2 by (metis Cong\<^sub>0_subst_left(2) R.null_is_zero(2) v'v'' R.conI) thus ?thesis using Cong\<^sub>0_transitive by blast qed moreover have "x \ \ \ ?w \\<^sub>0 t \\ x" apply (intro conjI) apply (meson composite_closed forward_stable u''v'' uu'' v'v'' x) apply (metis (full_types) R.arr_resid_iff_con R.con_implies_arr(2) R.con_sym ide_closed forward_stable R.composite_of_def R.resid_composite_of(3) Cong\<^sub>0_subst_right(1) prfx_closed u''v'' uu'' v'v'' x R.conI) by (metis (no_types, lifting) 1 R.con_composite_of_iff ide_closed R.resid_composite_of(3) R.arr_resid_iff_con R.con_implies_arr(1) R.con_sym x R.conI) moreover have "x' \ \ \ ?w' \\<^sub>0 t' \\ x'" apply (intro conjI) apply (meson composite_closed forward_stable uu'' v''u'' v'v'' x') apply (metis (full_types) Cong\<^sub>0_subst_right(1) R.composite_ofE R.con_sym ide_closed forward_stable R.con_imp_coinitial prfx_closed R.resid_composite_of(3) R.arr_resid_iff_con R.con_implies_arr(1) uu'' v'v'' x' R.conI) by (metis (full_types) Cong\<^sub>0_subst_left(1) R.composite_ofE R.con_sym ide_closed forward_stable R.con_imp_coinitial prfx_closed R.resid_composite_of(3) R.arr_resid_iff_con R.con_implies_arr(1) uu'' v'v'' x' R.conI) ultimately show "t \ t'" using Cong_def Cong\<^sub>0_transitive by metis qed lemma Cong_closure_props: shows "t \ u \ u \ t" and "\t \ u; u \ v\ \ t \ v" and "t \\<^sub>0 u \ t \ u" and "\u \ \; R.sources t = R.sources u\ \ t \ t \\ u" proof - show "t \ u \ u \ t" using Cong_symmetric by blast show "\t \ u; u \ v\ \ t \ v" using Cong_transitive by blast show "t \\<^sub>0 u \ t \ u" by (metis Cong\<^sub>0_subst_left(2) Cong_def Cong_reflexive R.con_implies_arr(1) R.null_is_zero(2) R.conI) show "\u \ \; R.sources t = R.sources u\ \ t \ t \\ u" proof - assume u: "u \ \" and coinitial: "R.sources t = R.sources u" obtain a where a: "a \ R.targets u" by (meson elements_are_arr empty_subsetI R.arr_iff_has_target subsetI subset_antisym u) have "t \\ u \\<^sub>0 (t \\ u) \\ a" proof - have "R.arr t" using R.arr_iff_has_source coinitial elements_are_arr u by presburger thus ?thesis by (meson u a R.arr_resid_iff_con coinitial ide_closed forward_stable elements_are_arr R.coinitial_iff R.composite_of_arr_target R.resid_composite_of(3)) qed thus ?thesis using Cong_def by (metis a R.composite_of_arr_target elements_are_arr factor_closed(2) u) qed qed lemma Cong\<^sub>0_implies_Cong: assumes "t \\<^sub>0 t'" shows "t \ t'" using assms Cong_closure_props(3) by simp lemma in_sources_respects_Cong: assumes "t \ t'" and "a \ R.sources t" and "a' \ R.sources t'" shows "a \ a'" proof - obtain u u' where uu': "u \ \ \ u' \ \ \ t \\ u \\<^sub>0 t' \\ u'" using assms Cong_def by blast show "a \ a'" proof show "u \ \" using uu' by simp show "u' \ \" using uu' by simp show "a \\ u \\<^sub>0 a' \\ u'" proof - have "a \\ u \ R.targets u" by (metis Cong\<^sub>0_imp_con R.arr_resid_iff_con assms(2) R.con_imp_common_source R.con_implies_arr(1) R.resid_source_in_targets R.sources_eqI uu') moreover have "a' \\ u' \ R.targets u'" by (metis Cong\<^sub>0_imp_con R.arr_resid_iff_con assms(3) R.con_imp_common_source R.resid_source_in_targets R.con_implies_arr(1) R.sources_eqI uu') moreover have "R.targets u = R.targets u'" by (metis Cong\<^sub>0_imp_coinitial Cong\<^sub>0_imp_con R.arr_resid_iff_con R.con_implies_arr(1) R.sources_resid uu') ultimately show ?thesis using ide_closed R.targets_are_cong by presburger qed qed qed lemma in_targets_respects_Cong: assumes "t \ t'" and "b \ R.targets t" and "b' \ R.targets t'" shows "b \ b'" proof - obtain u u' where uu': "u \ \ \ u' \ \ \ t \\ u \\<^sub>0 t' \\ u'" using assms Cong_def by blast have seq: "R.seq (u \\ t) ((t' \\ u') \\ (t \\ u)) \ R.seq (u' \\ t') ((t \\ u) \\ (t' \\ u'))" by (metis R.arr_iff_has_source R.arr_iff_has_target R.conI elements_are_arr R.not_arr_null R.seqI R.sources_resid R.targets_resid_sym uu') obtain v where v: "R.composite_of (u \\ t) ((t' \\ u') \\ (t \\ u)) v" using seq composite_closed_right uu' by presburger obtain v' where v': "R.composite_of (u' \\ t') ((t \\ u) \\ (t' \\ u')) v'" using seq composite_closed_right uu' by presburger show "b \ b'" proof show v_in_\: "v \ \" by (metis composite_closed R.con_imp_coinitial R.con_implies_arr(1) forward_stable R.composite_of_def R.prfx_implies_con R.arr_resid_iff_con R.con_sym uu' v) show v'_in_\: "v' \ \" by (metis backward_stable R.composite_of_def R.con_imp_coinitial forward_stable R.null_is_zero(2) prfx_closed uu' v' R.conI) show "b \\ v \\<^sub>0 b' \\ v'" using assms uu' v v' by (metis R.arr_resid_iff_con ide_closed R.seq_def R.sources_resid R.targets_resid_sym R.resid_source_in_targets seq R.sources_composite_of R.targets_are_cong R.targets_composite_of) qed qed lemma sources_are_Cong: assumes "a \ R.sources t" and "a' \ R.sources t" shows "a \ a'" using assms by (simp add: ide_closed R.sources_are_cong Cong_closure_props(3)) lemma targets_are_Cong: assumes "b \ R.targets t" and "b' \ R.targets t" shows "b \ b'" using assms by (simp add: ide_closed R.targets_are_cong Cong_closure_props(3)) text \ It is \emph{not} the case that sources and targets are \\\-closed; \emph{i.e.} \t \ t' \ sources t = sources t'\ and \t \ t' \ targets t = targets t'\ do not hold, in general. \ lemma Resid_along_normal_preserves_reflects_con: assumes "u \ \" and "R.sources t = R.sources u" shows "t \\ u \ t' \\ u \ t \ t'" by (metis R.arr_resid_iff_con assms R.con_implies_arr(1-2) elements_are_arr R.coinitial_iff R.resid_reflects_con resid_along_elem_preserves_con) text \ We can alternatively characterize \\\ as the least symmetric and transitive relation on transitions that extends \\\<^sub>0\ and has the property of being preserved by residuation along transitions in \\\. \ inductive Cong' where "\t u. Cong' t u \ Cong' u t" | "\t u v. \Cong' t u; Cong' u v\ \ Cong' t v" | "\t u. t \\<^sub>0 u \ Cong' t u" | "\t u. \R.arr t; u \ \; R.sources t = R.sources u\ \ Cong' t (t \\ u)" lemma Cong'_if: shows "\u \ \; u' \ \; t \\ u \\<^sub>0 t' \\ u'\ \ Cong' t t'" proof - assume u: "u \ \" and u': "u' \ \" and 1: "t \\ u \\<^sub>0 t' \\ u'" show "Cong' t t'" using u u' 1 by (metis (no_types, lifting) Cong'.simps Cong\<^sub>0_imp_con R.arr_resid_iff_con R.coinitial_iff R.con_imp_coinitial) qed lemma Cong_char: shows "Cong t t' \ Cong' t t'" proof - have "Cong t t' \ Cong' t t'" using Cong_def Cong'_if by blast moreover have "Cong' t t' \ Cong t t'" apply (induction rule: Cong'.induct) using Cong_symmetric apply simp using Cong_transitive apply simp using Cong_closure_props(3) apply simp using Cong_closure_props(4) by simp ultimately show ?thesis using Cong_def by blast qed lemma normal_is_Cong_closed: assumes "t \ \" and "t \ t'" shows "t' \ \" using assms by (metis (full_types) CongE R.con_imp_coinitial forward_stable R.null_is_zero(2) backward_stable R.conI) subsection "Congruence Classes" text \ Here we develop some notions relating to the congruence classes of \\\. \ definition Cong_class ("\_\") where "Cong_class t \ {t'. t \ t'}" definition is_Cong_class where "is_Cong_class \ \ \t. t \ \ \ \ = \t\" definition Cong_class_rep where "Cong_class_rep \ \ SOME t. t \ \" lemma Cong_class_is_nonempty: assumes "is_Cong_class \" shows "\ \ {}" using assms is_Cong_class_def Cong_class_def by auto lemma rep_in_Cong_class: assumes "is_Cong_class \" shows "Cong_class_rep \ \ \" using assms is_Cong_class_def Cong_class_rep_def someI_ex [of "\t. t \ \"] by metis lemma arr_in_Cong_class: assumes "R.arr t" shows "t \ \t\" using assms Cong_class_def Cong_reflexive by simp lemma is_Cong_classI: assumes "R.arr t" shows "is_Cong_class \t\" using assms Cong_class_def is_Cong_class_def Cong_reflexive by blast lemma is_Cong_classI' [intro]: assumes "\ \ {}" and "\t t'. \t \ \; t' \ \\ \ t \ t'" and "\t t'. \t \ \; t' \ t\ \ t' \ \" shows "is_Cong_class \" proof - obtain t where t: "t \ \" using assms by auto have "\ = \t\" unfolding Cong_class_def using assms(2-3) t by blast thus ?thesis using is_Cong_class_def t by blast qed lemma Cong_class_memb_is_arr: assumes "is_Cong_class \" and "t \ \" shows "R.arr t" using assms Cong_class_def is_Cong_class_def Cong_imp_arr(2) by force lemma Cong_class_membs_are_Cong: assumes "is_Cong_class \" and "t \ \" and "t' \ \" shows "Cong t t'" using assms Cong_class_def is_Cong_class_def by (metis CollectD Cong_closure_props(2) Cong_symmetric) lemma Cong_class_eqI: assumes "t \ t'" shows "\t\ = \t'\" using assms Cong_class_def by (metis (full_types) Collect_cong Cong'.intros(1-2) Cong_char) lemma Cong_class_eqI': assumes "is_Cong_class \" and "is_Cong_class \" and "\ \ \ \ {}" shows "\ = \" using assms is_Cong_class_def Cong_class_eqI Cong_class_membs_are_Cong by (metis (no_types, lifting) Int_emptyI) lemma is_Cong_classE [elim]: assumes "is_Cong_class \" and "\\ \ {}; \t t'. \t \ \; t' \ \\ \ t \ t'; \t t'. \t \ \; t' \ t\ \ t' \ \\ \ T" shows T proof - have \: "\ \ {}" using assms Cong_class_is_nonempty by simp moreover have 1: "\t t'. \t \ \; t' \ \\ \ t \ t'" using assms Cong_class_membs_are_Cong by metis moreover have "\t t'. \t \ \; t' \ t\ \ t' \ \" using assms Cong_class_def by (metis 1 Cong_class_eqI Cong_imp_arr(1) is_Cong_class_def arr_in_Cong_class) ultimately show ?thesis using assms by blast qed lemma Cong_class_rep [simp]: assumes "is_Cong_class \" shows "\Cong_class_rep \\ = \" by (metis Cong_class_membs_are_Cong Cong_class_eqI assms is_Cong_class_def rep_in_Cong_class) lemma Cong_class_memb_Cong_rep: assumes "is_Cong_class \" and "t \ \" shows "Cong t (Cong_class_rep \)" using assms Cong_class_membs_are_Cong rep_in_Cong_class by simp lemma composite_of_normal_arr: shows "\ R.arr t; u \ \; R.composite_of u t t' \ \ t' \ t" by (meson Cong'.intros(3) Cong_char R.composite_of_def R.con_implies_arr(2) ide_closed R.prfx_implies_con Cong_closure_props(2,4) R.sources_composite_of) lemma composite_of_arr_normal: shows "\ arr t; u \ \; R.composite_of t u t' \ \ t' \\<^sub>0 t" by (meson Cong_closure_props(3) R.composite_of_def ide_closed prfx_closed) end subsection "Coherent Normal Sub-RTS's" text \ A \emph{coherent} normal sub-RTS is one that satisfies a parallel moves property with respect to arbitrary transitions. The congruence \\\ induced by a coherent normal sub-RTS is fully substitutive with respect to consistency and residuation, and in fact coherence is equivalent to substitutivity in this context. \ locale coherent_normal_sub_rts = normal_sub_rts + assumes coherent: "\ R.arr t; u \ \; u' \ \; R.sources u = R.sources u'; R.targets u = R.targets u'; R.sources t = R.sources u \ \ t \\ u \\<^sub>0 t \\ u'" (* * TODO: Should coherence be part of normality, or is it an additional property that guarantees * the existence of the quotient? * * e.g. see http://nlab-pages.s3.us-east-2.amazonaws.com/nlab/show/normal+subobject * Maybe also http://www.tac.mta.ca/tac/volumes/36/3/36-03.pdf for recent work. *) context normal_sub_rts begin text \ The above ``parallel moves'' formulation of coherence is equivalent to the following formulation, which involves ``opposing spans''. \ lemma coherent_iff: shows "(\t u u'. R.arr t \ u \ \ \ u' \ \ \ R.sources t = R.sources u \ R.sources u = R.sources u' \ R.targets u = R.targets u' \ t \\ u \\<^sub>0 t \\ u') \ (\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w')" proof assume 1: "\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w'" show "\t u u'. R.arr t \ u \ \ \ u' \ \ \ R.sources t = R.sources u \ R.sources u = R.sources u' \ R.targets u = R.targets u' \ t \\ u \\<^sub>0 t \\ u'" proof (intro allI impI, elim conjE) fix t u u' assume t: "R.arr t" and u: "u \ \" and u': "u' \ \" and tu: "R.sources t = R.sources u" and sources: "R.sources u = R.sources u'" and targets: "R.targets u = R.targets u'" show "t \\ u \\<^sub>0 t \\ u'" by (metis 1 Cong\<^sub>0_reflexive Resid_along_normal_preserves_Cong\<^sub>0 sources t targets tu u u') qed next assume 1: "\t u u'. R.arr t \ u \ \ \ u' \ \ \ R.sources t = R.sources u \ R.sources u = R.sources u' \ R.targets u = R.targets u' \ t \\ u \\<^sub>0 t \\ u'" show "\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w'" proof (intro allI impI, elim conjE) fix t t' v v' w w' assume v: "v \ \" and v': "v' \ \" and w: "w \ \" and w': "w' \ \" and vw: "R.sources v = R.sources w" and v'w': "R.sources v' = R.sources w'" and ww': "R.targets w = R.targets w'" and tvt'v': "(t \\ v) \\ (t' \\ v') \ \" and t'v'tv: "(t' \\ v') \\ (t \\ v) \ \" show "t \\ w \\<^sub>0 t' \\ w'" proof - have 3: "R.sources t = R.sources v \ R.sources t' = R.sources v'" using R.con_imp_coinitial by (meson Cong\<^sub>0_imp_con tvt'v' t'v'tv R.coinitial_iff R.arr_resid_iff_con) have 2: "t \\ w \ t' \\ w'" using Cong_closure_props by (metis tvt'v' t'v'tv 3 vw v'w' v v' w w') obtain z z' where zz': "z \ \ \ z' \ \ \ (t \\ w) \\ z \\<^sub>0 (t' \\ w') \\ z'" using 2 by auto have "(t \\ w) \\ z \\<^sub>0 (t \\ w) \\ z'" proof - have "R.coinitial ((t \\ w) \\ z) ((t \\ w) \\ z')" by (metis Cong\<^sub>0_imp_coinitial Cong_imp_arr(1) Resid_along_normal_preserves_reflects_con R.arr_def R.coinitialI R.con_imp_common_source Cong_closure_props(3) R.arr_resid_iff_con R.sources_eqI R.sources_resid ww' zz') thus ?thesis apply (intro conjI) by (metis 1 R.coinitial_iff R.con_imp_coinitial R.arr_resid_iff_con R.sources_resid zz')+ qed hence "(t \\ w) \\ z' \\<^sub>0 (t' \\ w') \\ z'" using zz' Cong\<^sub>0_transitive Cong\<^sub>0_symmetric by blast thus ?thesis using zz' Resid_along_normal_reflects_Cong\<^sub>0 by metis qed qed qed end context coherent_normal_sub_rts begin text \ The proof of the substitutivity of \\\ with respect to residuation only uses coherence in the ``opposing spans'' form. \ lemma coherent': assumes "v \ \" and "v' \ \" and "w \ \" and "w' \ \" and "R.sources v = R.sources w" and "R.sources v' = R.sources w'" and "R.targets w = R.targets w'" and "t \\ v \\<^sub>0 t' \\ v'" shows "t \\ w \\<^sub>0 t' \\ w'" using assms coherent coherent_iff by metis (* 6 sec *) text \ The relation \\\ is substitutive with respect to both arguments of residuation. \ lemma Cong_subst: assumes "t \ t'" and "u \ u'" and "t \ u" and "R.sources t' = R.sources u'" shows "t' \ u'" and "t \\ u \ t' \\ u'" proof - obtain v v' where vv': "v \ \ \ v' \ \ \ t \\ v \\<^sub>0 t' \\ v'" using assms by auto obtain w w' where ww': "w \ \ \ w' \ \ \ u \\ w \\<^sub>0 u' \\ w'" using assms by auto let ?x = "t \\ v" and ?x' = "t' \\ v'" let ?y = "u \\ w" and ?y' = "u' \\ w'" have xx': "?x \\<^sub>0 ?x'" using assms vv' by blast have yy': "?y \\<^sub>0 ?y'" using assms ww' by blast have 1: "t \\ w \\<^sub>0 t' \\ w'" proof - have "R.sources v = R.sources w" by (metis (no_types, lifting) Cong\<^sub>0_imp_con R.arr_resid_iff_con assms(3) R.con_imp_common_source R.con_implies_arr(2) R.sources_eqI ww' xx') moreover have "R.sources v' = R.sources w'" by (metis (no_types, lifting) assms(4) R.coinitial_iff R.con_imp_coinitial Cong\<^sub>0_imp_con R.arr_resid_iff_con ww' xx') moreover have "R.targets w = R.targets w'" by (metis Cong\<^sub>0_implies_Cong Cong\<^sub>0_imp_coinitial Cong_imp_arr(1) R.arr_resid_iff_con R.sources_resid ww') ultimately show ?thesis using assms vv' ww' by (intro coherent' [of v v' w w' t]) auto qed have 2: "t' \\ w' \ u' \\ w'" using assms 1 ww' by (metis Cong\<^sub>0_subst_left(1) Cong\<^sub>0_subst_right(1) Resid_along_normal_preserves_reflects_con R.arr_resid_iff_con R.coinitial_iff R.con_imp_coinitial elements_are_arr) thus 3: "t' \ u'" using ww' R.cube by force have "t \\ u \ ((t \\ u) \\ (w \\ u)) \\ (?y' \\ ?y)" proof - have "t \\ u \ (t \\ u) \\ (w \\ u)" by (metis Cong_closure_props(4) assms(3) R.con_imp_coinitial elements_are_arr forward_stable R.arr_resid_iff_con R.con_implies_arr(1) R.sources_resid ww') also have "... \ ((t \\ u) \\ (w \\ u)) \\ (?y' \\ ?y)" by (metis Cong\<^sub>0_imp_con Cong_closure_props(4) Cong_imp_arr(2) R.arr_resid_iff_con calculation R.con_implies_arr(2) R.targets_resid_sym R.sources_resid ww') finally show ?thesis by simp qed also have "... \ (((t \\ w) \\ ?y) \\ (?y' \\ ?y))" using ww' by (metis Cong_imp_arr(2) Cong_reflexive calculation R.cube) also have "... \ (((t' \\ w') \\ ?y) \\ (?y' \\ ?y))" using 1 Cong\<^sub>0_subst_left(2) [of "t \\ w" "(t' \\ w')" ?y] Cong\<^sub>0_subst_left(2) [of "(t \\ w) \\ ?y" "(t' \\ w') \\ ?y" "?y' \\ ?y"] by (meson 2 Cong\<^sub>0_implies_Cong Cong\<^sub>0_subst_Con Cong_imp_arr(2) R.arr_resid_iff_con calculation ww') also have "... \ ((t' \\ w') \\ ?y') \\ (?y \\ ?y')" using 2 Cong\<^sub>0_implies_Cong Cong\<^sub>0_subst_right(2) ww' by presburger also have 4: "... \ (t' \\ u') \\ (w' \\ u')" using 2 ww' by (metis Cong\<^sub>0_imp_con Cong_closure_props(4) Cong_symmetric R.cube R.sources_resid) also have "... \ t' \\ u'" using ww' 3 4 by (metis Cong_closure_props(4) Cong_imp_arr(2) Cong_symmetric R.con_imp_coinitial R.con_implies_arr(2) forward_stable R.sources_resid R.arr_resid_iff_con) finally show "t \\ u \ t' \\ u'" by simp qed lemma Cong_subst_con: assumes "R.sources t = R.sources u" and "R.sources t' = R.sources u'" and "t \ t'" and "u \ u'" shows "t \ u \ t' \ u'" using assms by (meson Cong_subst(1) Cong_symmetric) lemma Cong\<^sub>0_composite_of_arr_normal: assumes "R.composite_of t u t'" and "u \ \" shows "t' \\<^sub>0 t" using assms backward_stable R.composite_of_def ide_closed by blast lemma Cong_composite_of_normal_arr: assumes "R.composite_of u t t'" and "u \ \" shows "t' \ t" using assms by (meson Cong_closure_props(2-4) R.arr_composite_of ide_closed R.composite_of_def R.sources_composite_of) end context normal_sub_rts begin text \ Coherence is not an arbitrary property: here we show that substitutivity of congruence in residuation is equivalent to the ``opposing spans'' form of coherence. \ lemma Cong_subst_iff_coherent': shows "(\t t' u u'. t \ t' \ u \ u' \ t \ u \ R.sources t' = R.sources u' \ t' \ u' \ t \\ u \ t' \\ u') \ (\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w')" proof assume 1: "\t t' u u'. t \ t' \ u \ u' \ t \ u \ R.sources t' = R.sources u' \ t' \ u' \ t \\ u \ t' \\ u'" show "\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w'" proof (intro allI impI, elim conjE) fix t t' v v' w w' assume v: "v \ \" and v': "v' \ \" and w: "w \ \" and w': "w' \ \" and sources_vw: "R.sources v = R.sources w" and sources_v'w': "R.sources v' = R.sources w'" and targets_ww': "R.targets w = R.targets w'" and tt': "(t \\ v) \\ (t' \\ v') \ \" and t't: "(t' \\ v') \\ (t \\ v) \ \" show "t \\ w \\<^sub>0 t' \\ w'" proof - have 2: "\t t' u u'. \t \ t'; u \ u'; t \ u; R.sources t' = R.sources u'\ \ t' \ u' \ t \\ u \ t' \\ u'" using 1 by blast have 3: "t \\ w \ t \\ v \ t' \\ w' \ t' \\ v'" by (metis tt' t't sources_vw sources_v'w' Cong\<^sub>0_subst_right(2) Cong_closure_props(4) Cong_def R.arr_resid_iff_con Cong_closure_props(3) Cong_imp_arr(1) normal_is_Cong_closed v w v' w') have "(t \\ w) \\ (t' \\ w') \ (t \\ v) \\ (t' \\ v')" using 2 [of "t \\ w" "t \\ v" "t' \\ w'" "t' \\ v'"] 3 by (metis tt' t't targets_ww' 1 Cong\<^sub>0_imp_con Cong_imp_arr(1) Cong_symmetric R.arr_resid_iff_con R.sources_resid) moreover have "(t' \\ w') \\ (t \\ w) \ (t' \\ v') \\ (t \\ v)" using 2 3 by (metis tt' t't targets_ww' Cong\<^sub>0_imp_con Cong_symmetric Cong_imp_arr(1) R.arr_resid_iff_con R.sources_resid) ultimately show ?thesis by (meson tt' t't normal_is_Cong_closed Cong_symmetric) qed qed next assume 1: "\t t' v v' w w'. v \ \ \ v' \ \ \ w \ \ \ w' \ \ \ R.sources v = R.sources w \ R.sources v' = R.sources w' \ R.targets w = R.targets w' \ t \\ v \\<^sub>0 t' \\ v' \ t \\ w \\<^sub>0 t' \\ w'" show "\t t' u u'. t \ t' \ u \ u' \ t \ u \ R.sources t' = R.sources u' \ t' \ u' \ t \\ u \ t' \\ u'" proof (intro allI impI, elim conjE, intro conjI) have *: "\t t' v v' w w'. \v \ \; v' \ \; w \ \; w' \ \; R.sources v = R.sources w; R.sources v' = R.sources w'; R.targets v = R.targets v'; R.targets w = R.targets w'; t \\ v \\<^sub>0 t' \\ v'\ \ t \\ w \\<^sub>0 t' \\ w'" using 1 by metis fix t t' u u' assume tt': "t \ t'" and uu': "u \ u'" and con: "t \ u" and t'u': "R.sources t' = R.sources u'" obtain v v' where vv': "v \ \ \ v' \ \ \ t \\ v \\<^sub>0 t' \\ v'" using tt' by auto obtain w w' where ww': "w \ \ \ w' \ \ \ u \\ w \\<^sub>0 u' \\ w'" using uu' by auto let ?x = "t \\ v" and ?x' = "t' \\ v'" let ?y = "u \\ w" and ?y' = "u' \\ w'" have xx': "?x \\<^sub>0 ?x'" using tt' vv' by blast have yy': "?y \\<^sub>0 ?y'" using uu' ww' by blast have 1: "t \\ w \\<^sub>0 t' \\ w'" proof - have "R.sources v = R.sources w \ R.sources v' = R.sources w'" proof show "R.sources v' = R.sources w'" using Cong\<^sub>0_imp_con R.arr_resid_iff_con R.coinitial_iff R.con_imp_coinitial t'u' vv' ww' by metis show "R.sources v = R.sources w" by (metis con elements_are_arr R.not_arr_null R.null_is_zero(2) R.conI R.con_imp_common_source rts.sources_eqI R.rts_axioms vv' ww') qed moreover have "R.targets v = R.targets v' \ R.targets w = R.targets w'" by (metis Cong\<^sub>0_imp_coinitial Cong\<^sub>0_imp_con R.arr_resid_iff_con R.con_implies_arr(2) R.sources_resid vv' ww') ultimately show ?thesis using vv' ww' xx' by (intro * [of v v' w w' t t']) auto qed have 2: "t' \\ w' \ u' \\ w'" using 1 tt' ww' by (meson Cong\<^sub>0_imp_con Cong\<^sub>0_subst_Con R.arr_resid_iff_con con R.con_imp_coinitial R.con_implies_arr(2) resid_along_elem_preserves_con) thus 3: "t' \ u'" using ww' R.cube by force have "t \\ u \ (t \\ u) \\ (w \\ u)" by (metis Cong_closure_props(4) R.arr_resid_iff_con con R.con_imp_coinitial elements_are_arr forward_stable R.con_implies_arr(2) R.sources_resid ww') also have "(t \\ u) \\ (w \\ u) \ ((t \\ u) \\ (w \\ u)) \\ (?y' \\ ?y)" using yy' by (metis Cong\<^sub>0_imp_con Cong_closure_props(4) Cong_imp_arr(2) R.arr_resid_iff_con calculation R.con_implies_arr(2) R.sources_resid R.targets_resid_sym) also have "... \ (((t \\ w) \\ ?y) \\ (?y' \\ ?y))" using ww' by (metis Cong_imp_arr(2) Cong_reflexive calculation R.cube) also have "... \ (((t' \\ w') \\ ?y) \\ (?y' \\ ?y))" proof - have "((t \\ w) \\ ?y) \\ (?y' \\ ?y) \\<^sub>0 ((t' \\ w') \\ ?y) \\ (?y' \\ ?y)" using 1 2 Cong\<^sub>0_subst_left(2) by (meson Cong\<^sub>0_subst_Con calculation Cong_imp_arr(2) R.arr_resid_iff_con ww') thus ?thesis using Cong\<^sub>0_implies_Cong by presburger qed also have "... \ ((t' \\ w') \\ ?y') \\ (?y \\ ?y')" by (meson "2" Cong\<^sub>0_implies_Cong Cong\<^sub>0_subst_right(2) ww') also have 4: "... \ (t' \\ u') \\ (w' \\ u')" using 2 ww' by (metis Cong\<^sub>0_imp_con Cong_closure_props(4) Cong_symmetric R.cube R.sources_resid) also have "... \ t' \\ u'" using ww' 2 3 4 by (metis Cong'.intros(1) Cong'.intros(4) Cong_char Cong_imp_arr(2) R.arr_resid_iff_con forward_stable R.con_imp_coinitial R.sources_resid R.con_implies_arr(2)) finally show "t \\ u \ t' \\ u'" by simp qed qed end subsection "Quotient by Coherent Normal Sub-RTS" text \ We now define the quotient of an RTS by a coherent normal sub-RTS and show that it is an extensional RTS. \ locale quotient_by_coherent_normal = R: rts + N: coherent_normal_sub_rts begin definition Resid (infix "\\\\" 70) where "\ \\\\ \ \ if N.is_Cong_class \ \ N.is_Cong_class \ \ (\t u. t \ \ \ u \ \ \ t \ u) then N.Cong_class (fst (SOME tu. fst tu \ \ \ snd tu \ \ \ fst tu \ snd tu) \\ snd (SOME tu. fst tu \ \ \ snd tu \ \ \ fst tu \ snd tu)) else {}" sublocale partial_magma Resid using N.Cong_class_is_nonempty Resid_def by unfold_locales metis lemma is_partial_magma: shows "partial_magma Resid" .. lemma null_char: shows "null = {}" using N.Cong_class_is_nonempty Resid_def by (metis null_is_zero(2)) lemma Resid_by_members: assumes "N.is_Cong_class \" and "N.is_Cong_class \" and "t \ \" and "u \ \" and "t \ u" shows "\ \\\\ \ = \t \\ u\" using assms Resid_def someI_ex [of "\tu. fst tu \ \ \ snd tu \ \ \ fst tu \ snd tu"] apply simp by (meson N.Cong_class_membs_are_Cong N.Cong_class_eqI N.Cong_subst(2) R.coinitial_iff R.con_imp_coinitial) abbreviation Con (infix "\\\" 50) where "\ \\\ \ \ \ \\\\ \ \ {}" lemma Con_char: shows "\ \\\ \ \ N.is_Cong_class \ \ N.is_Cong_class \ \ (\t u. t \ \ \ u \ \ \ t \ u)" by (metis (no_types, opaque_lifting) N.Cong_class_is_nonempty N.is_Cong_classI Resid_def Resid_by_members R.arr_resid_iff_con) lemma Con_sym: assumes "Con \ \" shows "Con \ \" using assms Con_char R.con_sym by meson lemma is_Cong_class_Resid: assumes "\ \\\ \" shows "N.is_Cong_class (\ \\\\ \)" using assms Con_char Resid_by_members R.arr_resid_iff_con N.is_Cong_classI by auto lemma Con_witnesses: assumes "\ \\\ \" and "t \ \" and "u \ \" shows "\v w. v \ \ \ w \ \ \ t \\ v \ u \\ w" proof - have 1: "N.is_Cong_class \ \ N.is_Cong_class \ \ (\t u. t \ \ \ u \ \ \ t \ u)" using assms Con_char by simp obtain t' u' where t'u': "t' \ \ \ u' \ \ \ t' \ u'" using 1 by auto have 2: "t' \ t \ u' \ u" using assms 1 t'u' N.Cong_class_membs_are_Cong by auto obtain v v' where vv': "v \ \ \ v' \ \ \ t' \\ v \\<^sub>0 t \\ v'" using 2 by auto obtain w w' where ww': "w \ \ \ w' \ \ \ u' \\ w \\<^sub>0 u \\ w'" using 2 by auto have 3: "w \ v" by (metis R.arr_resid_iff_con R.con_def R.con_imp_coinitial R.ex_un_null N.elements_are_arr R.null_is_zero(2) N.resid_along_elem_preserves_con t'u' vv' ww') have "R.seq v (w \\ v)" by (simp add: N.elements_are_arr R.seq_def 3 vv') obtain x where x: "R.composite_of v (w \\ v) x" using N.composite_closed_left \R.seq v (w \ v)\ vv' by blast obtain x' where x': "R.composite_of v' (w \\ v) x'" using x vv' N.composite_closed_left by (metis N.Cong\<^sub>0_implies_Cong N.Cong\<^sub>0_imp_coinitial N.Cong_imp_arr(1) R.composable_def R.composable_imp_seq R.con_implies_arr(2) R.seq_def R.sources_resid R.arr_resid_iff_con) have *: "t' \\ x \\<^sub>0 t \\ x'" by (metis N.coherent' N.composite_closed N.forward_stable R.con_imp_coinitial R.targets_composite_of 3 R.con_sym R.sources_composite_of vv' ww' x x') obtain y where y: "R.composite_of w (v \\ w) y" using x vv' ww' by (metis R.arr_resid_iff_con R.composable_def R.composable_imp_seq R.con_imp_coinitial R.seq_def R.sources_resid N.elements_are_arr N.forward_stable N.composite_closed_left) obtain y' where y': "R.composite_of w' (v \\ w) y'" using y ww' by (metis N.Cong\<^sub>0_imp_coinitial N.Cong_closure_props(3) N.Cong_imp_arr(1) R.composable_def R.composable_imp_seq R.con_implies_arr(2) R.seq_def R.sources_resid N.composite_closed_left R.arr_resid_iff_con) have **: "u' \\ y \\<^sub>0 u \\ y'" by (metis N.composite_closed N.forward_stable R.con_imp_coinitial R.targets_composite_of \w \ v\ N.coherent' R.sources_composite_of vv' ww' y y') have 4: "x \ \ \ y \ \" using x y vv' ww' * ** by (metis 3 N.composite_closed N.forward_stable R.con_imp_coinitial R.con_sym) have "t \\ x' \ u \\ y'" proof - have "t \\ x' \\<^sub>0 t' \\ x" using * by simp moreover have "t' \\ x \ u' \\ y" proof - have "t' \\ x \ u' \\ x" using t'u' vv' ww' 4 * by (metis N.Resid_along_normal_preserves_reflects_con N.elements_are_arr R.coinitial_iff R.con_imp_coinitial R.arr_resid_iff_con) moreover have "u' \\ x \\<^sub>0 u' \\ y" using ww' x y by (metis 4 N.Cong\<^sub>0_imp_coinitial N.Cong\<^sub>0_imp_con N.Cong\<^sub>0_transitive N.coherent' N.factor_closed(2) R.sources_composite_of R.targets_composite_of R.targets_resid_sym) ultimately show ?thesis using N.Cong\<^sub>0_subst_right by blast qed moreover have "u' \\ y \\<^sub>0 u \\ y'" using ** R.con_sym by simp ultimately show ?thesis using N.Cong\<^sub>0_subst_Con by auto qed moreover have "x' \ \ \ y' \ \" using x' y' vv' ww' by (metis N.Cong_composite_of_normal_arr N.Cong_imp_arr(2) N.composite_closed R.con_imp_coinitial N.forward_stable R.arr_resid_iff_con) ultimately show ?thesis by auto qed abbreviation Arr where "Arr \ \ Con \ \" lemma Arr_Resid: assumes "Con \ \" shows "Arr (\ \\\\ \)" by (metis Con_char N.Cong_class_memb_is_arr R.arrE N.rep_in_Cong_class assms is_Cong_class_Resid) lemma Cube: assumes "Con (\ \\\\ \) (\ \\\\ \)" shows "(\ \\\\ \) \\\\ (\ \\\\ \) = (\ \\\\ \) \\\\ (\ \\\\ \)" proof - obtain t u where tu: "t \ \ \ u \ \ \ t \ u \ \ \\\\ \ = \t \\ u\" using assms by (metis Con_char N.Cong_class_is_nonempty R.con_sym Resid_by_members) obtain t' v where t'v: "t' \ \ \ v \ \ \ t' \ v \ \ \\\\ \ = \t' \\ v\" using assms by (metis Con_char N.Cong_class_is_nonempty Resid_by_members Con_sym) have tt': "t \ t'" using assms by (metis N.Cong_class_membs_are_Cong N.Cong_class_is_nonempty Resid_def t'v tu) obtain w w' where ww': "w \ \ \ w' \ \ \ t \\ w \\<^sub>0 t' \\ w'" using tu t'v tt' by auto have 1: "\ \\\\ \ = \u \\ t\ \ \ \\\\ \ = \v \\ t'\" by (metis Con_char N.Cong_class_is_nonempty R.con_sym Resid_by_members assms t'v tu) obtain x x' where xx': "x \ \ \ x' \ \ \ (u \\ t) \\ x \ (v \\ t') \\ x'" using 1 Con_witnesses [of "\ \\\\ \" "\ \\\\ \" "u \\ t" "v \\ t'"] by (metis N.arr_in_Cong_class R.con_sym t'v tu assms Con_sym R.arr_resid_iff_con) have "R.seq t x" by (metis R.arr_resid_iff_con R.coinitial_iff R.con_imp_coinitial R.seqI R.sources_resid xx') have "R.seq t' x'" by (metis R.arr_resid_iff_con R.sources_resid R.coinitialE R.con_imp_coinitial R.seqI xx') obtain tx where tx: "R.composite_of t x tx" using xx' \R.seq t x\ N.composite_closed_right [of x t] R.composable_def by auto obtain t'x' where t'x': "R.composite_of t' x' t'x'" using xx' \R.seq t' x'\ N.composite_closed_right [of x' t'] R.composable_def by auto let ?tx_w = "tx \\ w" and ?t'x'_w' = "t'x' \\ w'" let ?w_tx = "(w \\ t) \\ x" and ?w'_t'x' = "(w' \\ t') \\ x'" let ?u_tx = "(u \\ t) \\ x" and ?v_t'x' = "(v \\ t') \\ x'" let ?u_w = "u \\ w" and ?v_w' = "v \\ w'" let ?w_u = "w \\ u" and ?w'_v = "w' \\ v" have w_tx_in_\: "?w_tx \ \" using tx ww' xx' R.con_composite_of_iff [of t x tx w] by (metis (full_types) N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) N.forward_stable R.null_is_zero(2) R.con_imp_coinitial R.conI R.con_sym) have w'_t'x'_in_\: "?w'_t'x' \ \" using t'x' ww' xx' R.con_composite_of_iff [of t' x' t'x' w'] by (metis (full_types) N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) R.con_sym N.forward_stable R.null_is_zero(2) R.con_imp_coinitial R.conI) have 2: "?tx_w \\<^sub>0 ?t'x'_w'" proof - have "?tx_w \\<^sub>0 t \\ w" using t'x' tx ww' xx' N.Cong\<^sub>0_composite_of_arr_normal [of t x tx] N.Cong\<^sub>0_subst_left(2) by (metis N.Cong\<^sub>0_transitive R.conI) also have "t \\ w \\<^sub>0 t' \\ w'" using ww' by blast also have "t' \\ w' \\<^sub>0 ?t'x'_w'" using t'x' tx ww' xx' N.Cong\<^sub>0_composite_of_arr_normal [of t' x' t'x'] N.Cong\<^sub>0_subst_left(2) by (metis N.Cong\<^sub>0_transitive R.conI) finally show ?thesis by blast qed obtain z where z: "R.composite_of ?tx_w (?t'x'_w' \\ ?tx_w) z" by (metis "2" R.arr_resid_iff_con R.con_implies_arr(2) N.elements_are_arr N.composite_closed_right R.seqI R.sources_resid) obtain z' where z': "R.composite_of ?t'x'_w' (?tx_w \\ ?t'x'_w') z'" by (metis "2" R.arr_resid_iff_con R.con_implies_arr(2) N.elements_are_arr N.composite_closed_right R.seqI R.sources_resid) have 3: "z \\<^sub>0 z'" using 2 N.diamond_commutes_upto_Cong\<^sub>0 N.Cong\<^sub>0_imp_con z z' by blast have "R.targets z = R.targets z'" by (metis R.targets_resid_sym z z' R.targets_composite_of R.conI) have Con_z_uw: "z \ ?u_w" proof - have "?tx_w \ ?u_w" by (meson 3 N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) R.bounded_imp_con R.con_implies_arr(1) R.con_imp_coinitial N.resid_along_elem_preserves_con tu tx ww' xx' z z' R.arr_resid_iff_con) thus ?thesis using 2 N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) z by blast qed moreover have Con_z'_vw': "z' \ ?v_w'" proof - have "?t'x'_w' \ ?v_w'" by (meson 3 N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) R.bounded_imp_con t'v t'x' ww' xx' z z' R.con_imp_coinitial N.resid_along_elem_preserves_con R.arr_resid_iff_con R.con_implies_arr(1)) thus ?thesis by (meson 2 N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_left(1) z') qed moreover have Con_z_vw': "z \ ?v_w'" using 3 Con_z'_vw' N.Cong\<^sub>0_subst_left(1) by blast moreover have *: "?u_w \\ z \ ?v_w' \\ z" proof - obtain y where y: "R.composite_of (w \\ tx) (?t'x'_w' \\ ?tx_w) y" by (metis 2 R.arr_resid_iff_con R.composable_def R.composable_imp_seq R.con_imp_coinitial N.elements_are_arr N.composite_closed_right R.seq_def R.targets_resid_sym ww' z N.forward_stable) obtain y' where y': "R.composite_of (w' \\ t'x') (?tx_w \\ ?t'x'_w') y'" by (metis 2 R.arr_resid_iff_con R.composable_def R.composable_imp_seq R.con_imp_coinitial N.elements_are_arr N.composite_closed_right R.targets_resid_sym ww' z' R.seq_def N.forward_stable) have y_comp: "R.composite_of (w \\ tx) ((t'x' \\ w') \\ (tx \\ w)) y" using y by simp have y_in_normal: "y \ \" by (metis 2 Con_z_uw R.arr_iff_has_source R.arr_resid_iff_con N.composite_closed R.con_imp_coinitial R.con_implies_arr(1) N.forward_stable R.sources_composite_of ww' y_comp z) have y_coinitial: "R.coinitial y (u \\ tx)" using y R.coinitial_def by (metis Con_z_uw R.con_def R.con_prfx_composite_of(2) R.con_sym R.cube R.sources_composite_of R.con_imp_common_source z) have y_con: "y \ u \\ tx" using y_in_normal y_coinitial by (metis R.coinitial_iff N.elements_are_arr N.forward_stable R.arr_resid_iff_con) have A: "?u_w \\ z \ (u \\ tx) \\ y" proof - have "(u \\ tx) \\ y \ ((u \\ tx) \\ (w \\ tx)) \\ (?t'x'_w' \\ ?tx_w)" using y_comp y_con R.resid_composite_of(3) [of "w \\ tx" "?t'x'_w' \\ ?tx_w" y "u \\ tx"] by simp also have "((u \\ tx) \\ (w \\ tx)) \\ (?t'x'_w' \\ ?tx_w) \ ?u_w \\ z" by (metis Con_z_uw R.resid_composite_of(3) z R.cube) finally show ?thesis by blast qed have y'_comp: "R.composite_of (w' \\ t'x') (?tx_w \\ ?t'x'_w') y'" using y' by simp have y'_in_normal: "y' \ \" by (metis 2 Con_z'_vw' R.arr_iff_has_source R.arr_resid_iff_con N.composite_closed R.con_imp_coinitial R.con_implies_arr(1) N.forward_stable R.sources_composite_of ww' y'_comp z') have y'_coinitial: "R.coinitial y' (v \\ t'x')" using y' R.coinitial_def by (metis Con_z'_vw' R.arr_resid_iff_con R.composite_ofE R.con_imp_coinitial R.con_implies_arr(1) R.cube R.prfx_implies_con R.resid_composite_of(1) R.sources_resid z') have y'_con: "y' \ v \\ t'x'" using y'_in_normal y'_coinitial by (metis R.coinitial_iff N.elements_are_arr N.forward_stable R.arr_resid_iff_con) have B: "?v_w' \\ z' \ (v \\ t'x') \\ y'" proof - have "(v \\ t'x') \\ y' \ ((v \\ t'x') \\ (w' \\ t'x')) \\ (?tx_w \\ ?t'x'_w')" using y'_comp y'_con R.resid_composite_of(3) [of "w' \\ t'x'" "?tx_w \\ ?t'x'_w'" y' "v \\ t'x'"] by blast also have "((v \\ t'x') \\ (w' \\ t'x')) \\ (?tx_w \\ ?t'x'_w') \ ?v_w' \\ z'" by (metis Con_z'_vw' R.cube R.resid_composite_of(3) z') finally show ?thesis by blast qed have C: "u \\ tx \ v \\ t'x'" using tx t'x' xx' R.con_sym R.cong_subst_right(1) R.resid_composite_of(3) by (meson R.coinitial_iff R.arr_resid_iff_con y'_coinitial y_coinitial) have D: "y \\<^sub>0 y'" proof - have "y \\<^sub>0 w \\ tx" using 2 N.Cong\<^sub>0_composite_of_arr_normal y_comp by blast also have "w \\ tx \\<^sub>0 w' \\ t'x'" proof - have "w \\ tx \ \ \ w' \\ t'x' \ \" using N.factor_closed(1) y_comp y_in_normal y'_comp y'_in_normal by blast moreover have "R.coinitial (w \\ tx) (w' \\ t'x')" by (metis C R.coinitial_def R.con_implies_arr(2) N.elements_are_arr R.sources_resid calculation R.con_imp_coinitial R.arr_resid_iff_con y_con) ultimately show ?thesis by (meson R.arr_resid_iff_con R.con_imp_coinitial N.forward_stable N.elements_are_arr) qed also have "w' \\ t'x' \\<^sub>0 y'" using 2 N.Cong\<^sub>0_composite_of_arr_normal y'_comp by blast finally show ?thesis by blast qed have par_y_y': "R.sources y = R.sources y' \ R.targets y = R.targets y'" using D N.Cong\<^sub>0_imp_coinitial R.targets_composite_of y'_comp y_comp z z' \R.targets z = R.targets z'\ by presburger have E: "(u \\ tx) \\ y \ (v \\ t'x') \\ y'" proof - have "(u \\ tx) \\ y \ (v \\ t'x') \\ y" using C N.Resid_along_normal_preserves_reflects_con R.coinitial_iff y_coinitial y_in_normal by presburger moreover have "(v \\ t'x') \\ y \\<^sub>0 (v \\ t'x') \\ y'" using par_y_y' N.coherent R.coinitial_iff y'_coinitial y'_in_normal y_in_normal by presburger ultimately show ?thesis using N.Cong\<^sub>0_subst_right(1) by blast qed hence "?u_w \\ z \ ?v_w' \\ z'" proof - have "(u \\ tx) \\ y \ ?u_w \\ z" using A by simp moreover have "(u \\ tx) \\ y \ (v \\ t'x') \\ y'" using E by blast moreover have "(v \\ t'x') \\ y' \ ?v_w' \\ z'" using B R.cong_symmetric by blast moreover have "R.sources ((u \\ w) \\ z) = R.sources ((v \\ w') \\ z')" by (simp add: Con_z'_vw' Con_z_uw R.con_sym \R.targets z = R.targets z'\) ultimately show ?thesis by (meson N.Cong\<^sub>0_subst_Con N.ide_closed) qed moreover have "?v_w' \\ z' \ ?v_w' \\ z" by (meson 3 Con_z_vw' N.CongI N.Cong\<^sub>0_subst_right(2) R.con_sym) moreover have "R.sources ((v \\ w') \\ z) = R.sources ((u \\ w) \\ z)" by (metis R.con_implies_arr(1) R.sources_resid calculation(1) calculation(2) N.Cong_imp_arr(2) R.arr_resid_iff_con) ultimately show ?thesis by (metis N.Cong_reflexive N.Cong_subst(1) R.con_implies_arr(1)) qed ultimately have **: "?v_w' \\ z \ ?u_w \\ z \ (?v_w' \\ z) \\ (?u_w \\ z) = (?v_w' \\ ?u_w) \\ (z \\ ?u_w)" by (meson R.con_sym R.cube) have Cong_t_z: "t \ z" by (metis 2 N.Cong\<^sub>0_composite_of_arr_normal N.Cong_closure_props(2-3) N.Cong_closure_props(4) N.Cong_imp_arr(2) R.coinitial_iff R.con_imp_coinitial tx ww' xx' z R.arr_resid_iff_con) have Cong_u_uw: "u \ ?u_w" by (meson Con_z_uw N.Cong_closure_props(4) R.coinitial_iff R.con_imp_coinitial ww' R.arr_resid_iff_con) have Cong_v_vw': "v \ ?v_w'" by (meson Con_z_vw' N.Cong_closure_props(4) R.coinitial_iff ww' R.con_imp_coinitial R.arr_resid_iff_con) have \: "N.is_Cong_class \ \ z \ \" by (metis (no_types, lifting) Cong_t_z N.Cong_class_eqI N.Cong_class_is_nonempty N.Cong_class_memb_Cong_rep N.Cong_class_rep N.Cong_imp_arr(2) N.arr_in_Cong_class tu assms Con_char) have \: "N.is_Cong_class \ \ ?u_w \ \" by (metis Con_char Con_z_uw Cong_u_uw Int_iff N.Cong_class_eqI' N.Cong_class_eqI N.arr_in_Cong_class R.con_implies_arr(2) N.is_Cong_classI tu assms empty_iff) have \: "N.is_Cong_class \ \ ?v_w' \ \" by (metis Con_char Con_z_vw' Cong_v_vw' Int_iff N.Cong_class_eqI' N.Cong_class_eqI N.arr_in_Cong_class R.con_implies_arr(2) N.is_Cong_classI t'v assms empty_iff) show "(\ \\\\ \) \\\\ (\ \\\\ \) = (\ \\\\ \) \\\\ (\ \\\\ \)" proof - have "(\ \\\\ \) \\\\ (\ \\\\ \) = \(?v_w' \\ z) \\ (?u_w \\ z)\" using \ \ \ * Resid_by_members by (metis ** Con_char N.arr_in_Cong_class R.arr_resid_iff_con assms R.con_implies_arr(2)) moreover have "(\ \\\\ \) \\\\ (\ \\\\ \) = \(?v_w' \\ ?u_w) \\ (z \\ ?u_w)\" using Resid_by_members [of \ \ ?v_w' ?u_w] Resid_by_members [of \ \ z ?u_w] Resid_by_members [of "\ \\\\ \" "\ \\\\ \" "?v_w' \\ ?u_w" "z \\ ?u_w"] by (metis \ \ \ * ** N.arr_in_Cong_class R.con_implies_arr(2) N.is_Cong_classI R.resid_reflects_con R.arr_resid_iff_con) ultimately show ?thesis using ** by simp qed qed sublocale residuation Resid using null_char Con_sym Arr_Resid Cube by unfold_locales metis+ lemma is_residuation: shows "residuation Resid" .. lemma arr_char: shows "arr \ \ N.is_Cong_class \" by (metis N.is_Cong_class_def arrI not_arr_null null_char N.Cong_class_memb_is_arr Con_char R.arrE arrE arr_resid conI) lemma ide_char: shows "ide \ \ arr \ \ \ \ \ \ {}" proof show "ide \ \ arr \ \ \ \ \ \ {}" apply (elim ideE) by (metis Con_char N.Cong\<^sub>0_reflexive Resid_by_members disjoint_iff null_char N.arr_in_Cong_class R.arrE R.arr_resid arr_resid conE) show "arr \ \ \ \ \ \ {} \ ide \" proof - assume \: "arr \ \ \ \ \ \ {}" obtain u where u: "R.arr u \ u \ \ \ \" using \ arr_char by (metis IntI N.Cong_class_memb_is_arr disjoint_iff) show ?thesis by (metis IntD1 IntD2 N.Cong_class_eqI N.Cong_closure_props(4) N.arr_in_Cong_class N.is_Cong_classI Resid_by_members \ arrE arr_char disjoint_iff ideI N.Cong_class_eqI' R.arrE u) qed qed lemma ide_char': shows "ide \ \ arr \ \ \ \ \" by (metis Int_absorb2 Int_emptyI N.Cong_class_memb_Cong_rep N.Cong_closure_props(1) ide_char not_arr_null null_char N.normal_is_Cong_closed arr_char subsetI) lemma con_char\<^sub>Q\<^sub>C\<^sub>N: shows "con \ \ \ N.is_Cong_class \ \ N.is_Cong_class \ \ (\t u. t \ \ \ u \ \ \ t \ u)" by (metis Con_char conE conI null_char) (* * TODO: Does the stronger form of con_char hold in this context? * I am currently only able to prove it for the more special context of paths, * but it doesn't seem like that should be required. * * The issue is that congruent paths have the same sets of sources, * but this does not necessarily hold in general. If we know that all representatives * of a congruence class have the same sets of sources, then we known that if any * pair of representatives is consistent, then the arbitrarily chosen representatives * of the congruence class are consistent. This is by substitutivity of congruence, * which has coinitiality as a hypothesis. * * In the general case, we have to reason as follows: if t and u are consistent * representatives of \ and \, and if t' and u' are arbitrary coinitial representatives * of \ and \, then we can obtain "opposing spans" connecting t u and t' u'. * The opposing span form of coherence then implies that t' and u' are consistent. * So we should be able to show that if congruence classes \ and \ are consistent, * then all pairs of coinitial representatives are consistent. *) lemma con_imp_coinitial_members_are_con: assumes "con \ \" and "t \ \" and "u \ \" and "R.sources t = R.sources u" shows "t \ u" by (meson assms N.Cong_subst(1) N.is_Cong_classE con_char\<^sub>Q\<^sub>C\<^sub>N) sublocale rts Resid proof show 1: "\\ \. \ide \; con \ \\ \ \ \\\\ \ = \" proof - fix \ \ assume \: "ide \" and con: "con \ \" obtain t a where ta: "t \ \ \ a \ \ \ R.con t a \ \ \\\\ \ = \t \\ a\" using con con_char\<^sub>Q\<^sub>C\<^sub>N Resid_by_members by auto have "a \ \" using \ ta ide_char' by auto hence "t \\ a \ t" by (meson N.Cong_closure_props(4) N.Cong_symmetric R.coinitialE R.con_imp_coinitial ta) thus "\ \\\\ \ = \" using ta by (metis N.Cong_class_eqI N.Cong_class_memb_Cong_rep N.Cong_class_rep con con_char\<^sub>Q\<^sub>C\<^sub>N) qed show "\\. arr \ \ ide (trg \)" by (metis N.Cong\<^sub>0_reflexive Resid_by_members disjoint_iff ide_char N.Cong_class_memb_is_arr N.arr_in_Cong_class N.is_Cong_class_def arr_char R.arrE R.arr_resid resid_arr_self) show "\\ \. \ide \; con \ \\ \ ide (\ \\\\ \)" by (metis 1 arrE arr_resid con_sym ideE ideI cube) show "\\ \. con \ \ \ \\. ide \ \ con \ \ \ con \ \" proof - fix \ \ assume \\: "con \ \" obtain t u where tu: "\ = \t\ \ \ = \u\ \ t \ u" using \\ con_char\<^sub>Q\<^sub>C\<^sub>N arr_char by (metis N.Cong_class_memb_Cong_rep N.Cong_class_eqI N.Cong_class_rep) obtain a where a: "a \ R.sources t" using \\ tu R.con_implies_arr(1) R.arr_iff_has_source by blast have "ide \a\ \ con \a\ \ \ con \a\ \" proof (intro conjI) have 2: "a \ \" using \\ tu a arr_char N.ide_closed R.sources_def by force show 3: "ide \a\" using \\ tu 2 a ide_char arr_char con_char\<^sub>Q\<^sub>C\<^sub>N by (metis IntI N.arr_in_Cong_class N.is_Cong_classI empty_iff N.elements_are_arr) show "con \a\ \" using \\ tu 2 3 a ide_char arr_char con_char\<^sub>Q\<^sub>C\<^sub>N by (metis N.arr_in_Cong_class R.composite_of_source_arr R.composite_of_def R.prfx_implies_con R.con_implies_arr(1)) show "con \a\ \" using \\ tu a ide_char arr_char con_char\<^sub>Q\<^sub>C\<^sub>N by (metis N.arr_in_Cong_class R.composite_of_source_arr R.con_prfx_composite_of N.is_Cong_classI R.con_implies_arr(1) R.con_implies_arr(2)) qed thus "\\. ide \ \ con \ \ \ con \ \" by auto qed show "\\ \ \. \ide (\ \\\\ \); con \ \\ \ con (\ \\\\ \) (\ \\\\ \)" proof - fix \ \ \ assume \\: "ide (\ \\\\ \)" assume \\: "con \ \" obtain t u where tu: "t \ \ \ u \ \ \ t \ u \ \ \\\\ \ = \t \\ u\" using \\ by (meson Resid_by_members ide_implies_arr quotient_by_coherent_normal.con_char\<^sub>Q\<^sub>C\<^sub>N quotient_by_coherent_normal_axioms arr_resid_iff_con) obtain v u' where vu': "v \ \ \ u' \ \ \ v \ u' \ \ \\\\ \ = \v \\ u'\" by (meson R.con_sym Resid_by_members \\ con_char\<^sub>Q\<^sub>C\<^sub>N) have 1: "u \ u'" using \\ tu vu' by (meson N.Cong_class_membs_are_Cong con_char\<^sub>Q\<^sub>C\<^sub>N) obtain w w' where ww': "w \ \ \ w' \ \ \ u \\ w \\<^sub>0 u' \\ w'" using 1 by auto have 2: "((t \\ u) \\ (w \\ u)) \\ ((u' \\ w') \\ (u \\ w)) \ ((v \\ u') \\ (w' \\ u')) \\ ((u \\ w) \\ (u' \\ w'))" proof - have "((t \\ u) \\ (w \\ u)) \\ ((u' \\ w') \\ (u \\ w)) \ \" proof - have "t \\ u \ \" using tu N.arr_in_Cong_class R.arr_resid_iff_con \\ ide_char' by blast hence "(t \\ u) \\ (w \\ u) \ \" by (metis N.Cong_closure_props(4) N.forward_stable R.null_is_zero(2) R.con_imp_coinitial R.sources_resid N.Cong_imp_arr(2) R.arr_resid_iff_con tu ww' R.conI) thus ?thesis by (metis N.Cong_closure_props(4) N.normal_is_Cong_closed R.sources_resid R.targets_resid_sym N.elements_are_arr R.arr_resid_iff_con ww' R.conI) qed moreover have "R.sources (((t \\ u) \\ (w \\ u)) \\ ((u' \\ w') \\ (u \\ w))) = R.sources (((v \\ u') \\ (w' \\ u')) \\ ((u \\ w) \\ (u' \\ w')))" proof - have "R.sources (((t \\ u) \\ (w \\ u)) \\ ((u' \\ w') \\ (u \\ w))) = R.targets ((u' \\ w') \\ (u \\ w))" using R.arr_resid_iff_con N.elements_are_arr R.sources_resid calculation by blast also have "... = R.targets ((u \\ w) \\ (u' \\ w'))" by (metis R.targets_resid_sym R.conI) also have "... = R.sources (((v \\ u') \\ (w' \\ u')) \\ ((u \\ w) \\ (u' \\ w')))" using R.arr_resid_iff_con N.elements_are_arr R.sources_resid by (metis N.Cong_closure_props(4) N.Cong_imp_arr(2) R.con_implies_arr(1) R.con_imp_coinitial N.forward_stable R.targets_resid_sym vu' ww') finally show ?thesis by simp qed ultimately show ?thesis by (metis (no_types, lifting) N.Cong\<^sub>0_imp_con N.Cong_closure_props(4) N.Cong_imp_arr(2) R.arr_resid_iff_con R.con_imp_coinitial N.forward_stable R.null_is_zero(2) R.conI) qed moreover have "t \\ u \ ((t \\ u) \\ (w \\ u)) \\ ((u' \\ w') \\ (u \\ w))" by (metis (no_types, opaque_lifting) N.Cong_closure_props(4) N.Cong_transitive N.forward_stable R.arr_resid_iff_con R.con_imp_coinitial R.rts_axioms calculation rts.coinitial_iff ww') moreover have "v \\ u' \ ((v \\ u') \\ (w' \\ u')) \\ ((u \\ w) \\ (u' \\ w'))" proof - have "w' \\ u' \ \" by (meson R.con_implies_arr(2) R.con_imp_coinitial N.forward_stable ww' N.Cong\<^sub>0_imp_con R.arr_resid_iff_con) moreover have "(u \\ w) \\ (u' \\ w') \ \" using ww' by blast ultimately show ?thesis by (meson 2 N.Cong_closure_props(2) N.Cong_closure_props(4) R.arr_resid_iff_con R.coinitial_iff R.con_imp_coinitial) qed ultimately show "con (\ \\\\ \) (\ \\\\ \)" using con_char\<^sub>Q\<^sub>C\<^sub>N N.Cong_class_def N.is_Cong_classI tu vu' R.arr_resid_iff_con by auto qed qed lemma is_rts: shows "rts Resid" .. sublocale extensional_rts Resid proof fix \ \ assume \\: "cong \ \" show "\ = \" proof - obtain t u where tu: "\ = \t\ \ \ = \u\ \ t \ u" by (metis Con_char N.Cong_class_eqI N.Cong_class_memb_Cong_rep N.Cong_class_rep \\ ide_char not_arr_null null_char) have "t \\<^sub>0 u" proof show "t \\ u \ \" using tu \\ Resid_by_members [of \ \ t u] by (metis (full_types) N.arr_in_Cong_class R.con_implies_arr(1-2) N.is_Cong_classI ide_char' R.arr_resid_iff_con subset_iff) show "u \\ t \ \" using tu \\ Resid_by_members [of \ \ u t] R.con_sym by (metis (full_types) N.arr_in_Cong_class R.con_implies_arr(1-2) N.is_Cong_classI ide_char' R.arr_resid_iff_con subset_iff) qed hence "t \ u" using N.Cong\<^sub>0_implies_Cong by simp thus "\ = \" by (simp add: N.Cong_class_eqI tu) qed qed theorem is_extensional_rts: shows "extensional_rts Resid" .. lemma sources_char\<^sub>Q\<^sub>C\<^sub>N: shows "sources \ = {\. arr \ \ \ = {a. \t a'. t \ \ \ a' \ R.sources t \ a' \ a}}" proof - let ?\ = "{a. \t a'. t \ \ \ a' \ R.sources t \ a' \ a}" have 1: "arr \ \ ide ?\" proof (unfold ide_char', intro conjI) assume \: "arr \" show "?\ \ \" using N.ide_closed N.normal_is_Cong_closed by blast show "arr ?\" proof - have "N.is_Cong_class ?\" proof show "?\ \ {}" by (metis (mono_tags, lifting) Collect_empty_eq N.Cong_class_def N.Cong_imp_arr(1) N.is_Cong_class_def N.sources_are_Cong R.arr_iff_has_source R.sources_def \ arr_char mem_Collect_eq) show "\t t'. \t \ ?\; t' \ t\ \ t' \ ?\" using N.Cong_symmetric N.Cong_transitive by blast show "\a a'. \a \ ?\; a' \ ?\\ \ a \ a'" proof - fix a a' assume a: "a \ ?\" and a': "a' \ ?\" obtain t b where b: "t \ \ \ b \ R.sources t \ b \ a" using a by blast obtain t' b' where b': "t' \ \ \ b' \ R.sources t' \ b' \ a'" using a' by blast have "b \ b'" using \ arr_char b b' by (meson IntD1 N.Cong_class_membs_are_Cong N.in_sources_respects_Cong) thus "a \ a'" by (meson N.Cong_symmetric N.Cong_transitive b b') qed qed thus ?thesis using arr_char by auto qed qed moreover have "arr \ \ con \ ?\" proof - assume \: "arr \" obtain t a where a: "t \ \ \ a \ R.sources t" using \ arr_char by (metis N.Cong_class_is_nonempty R.arr_iff_has_source empty_subsetI N.Cong_class_memb_is_arr subsetI subset_antisym) have "t \ \ \ a \ {a. \t a'. t \ \ \ a' \ R.sources t \ a' \ a} \ t \ a" using a N.Cong_reflexive R.sources_def R.con_implies_arr(2) by fast thus ?thesis using \ 1 arr_char con_char\<^sub>Q\<^sub>C\<^sub>N [of \ ?\] by auto qed ultimately have "arr \ \ ?\ \ sources \" using sources_def by blast thus ?thesis using "1" ide_char sources_char by auto qed lemma targets_char\<^sub>Q\<^sub>C\<^sub>N: shows "targets \ = {\. arr \ \ \ = \ \\\\ \}" proof - have "targets \ = {\. ide \ \ con (\ \\\\ \) \}" by (simp add: targets_def trg_def) also have "... = {\. arr \ \ ide \ \ (\t u. t \ \ \\\\ \ \ u \ \ \ t \ u)}" using arr_resid_iff_con con_char\<^sub>Q\<^sub>C\<^sub>N arr_char arr_def by auto also have "... = {\. arr \ \ ide \ \ (\t t' b u. t \ \ \ t' \ \ \ t \ t' \ b \ \t \\ t'\ \ u \ \ \ b \ u)}" using arr_char ide_char Resid_by_members [of \ \] N.Cong_class_memb_is_arr N.is_Cong_class_def R.arr_def by auto metis+ also have "... = {\. arr \ \ ide \ \ (\t t' b. t \ \ \ t' \ \ \ t \ t' \ b \ \t \\ t'\ \ b \ \)}" proof - have "\\ t t' b. \arr \; ide \; t \ \; t' \ \; t \ t'; b \ \t \\ t'\\ \ (\u. u \ \ \ b \ u) \ b \ \" proof - fix \ t t' b assume \: "arr \" and \: "ide \" and t: "t \ \" and t': "t' \ \" and tt': "t \ t'" and b: "b \ \t \\ t'\" have 0: "b \ \" by (metis Resid_by_members \ b ide_char' ide_trg arr_char subsetD t t' trg_def tt') show "(\u. u \ \ \ b \ u) \ b \ \" using 0 by (meson N.Cong_closure_props(3) N.forward_stable N.elements_are_arr \ arr_char R.con_imp_coinitial N.is_Cong_classE ide_char' R.arrE R.con_sym subsetD) qed thus ?thesis using ide_char arr_char by (metis (no_types, lifting)) qed also have "... = {\. arr \ \ ide \ \ (\t t'. t \ \ \ t' \ \ \ t \ t' \ \t \\ t'\ \ \)}" proof - have "\\ t t' b. \arr \; ide \; t \ \; t' \ \; t \ t'\ \ (\b. b \ \t \\ t'\ \ b \ \) \ \t \\ t'\ \ \" using ide_char arr_char apply (intro iffI) apply (metis IntI N.Cong_class_eqI' R.arr_resid_iff_con N.is_Cong_classI empty_iff set_eq_subset) by (meson N.arr_in_Cong_class R.arr_resid_iff_con subsetD) thus ?thesis using ide_char arr_char by (metis (no_types, lifting)) qed also have "... = {\. arr \ \ ide \ \ \ \\\\ \ \ \}" using arr_char ide_char Resid_by_members [of \ \] by (metis (no_types, opaque_lifting) arrE con_char\<^sub>Q\<^sub>C\<^sub>N) also have "... = {\. arr \ \ \ = \ \\\\ \}" by (metis (no_types, lifting) arr_has_un_target calculation con_ide_are_eq cong_reflexive mem_Collect_eq targets_def trg_def) finally show ?thesis by blast qed lemma src_char\<^sub>Q\<^sub>C\<^sub>N: shows "src \ = {a. arr \ \ (\t a'. t \ \ \ a' \ R.sources t \ a' \ a)}" using sources_char\<^sub>Q\<^sub>C\<^sub>N [of \] by (simp add: null_char src_def) lemma trg_char\<^sub>Q\<^sub>C\<^sub>N: shows "trg \ = \ \\\\ \" unfolding trg_def by blast subsubsection "Quotient Map" abbreviation quot where "quot t \ \t\" sublocale quot: simulation resid Resid quot proof show "\t. \ R.arr t \ \t\ = null" using N.Cong_class_def N.Cong_imp_arr(1) null_char by force show "\t u. t \ u \ con \t\ \u\" by (meson N.arr_in_Cong_class N.is_Cong_classI R.con_implies_arr(1-2) con_char\<^sub>Q\<^sub>C\<^sub>N) show "\t u. t \ u \ \t \\ u\ = \t\ \\\\ \u\" by (metis N.arr_in_Cong_class N.is_Cong_classI R.con_implies_arr(1-2) Resid_by_members) qed lemma quotient_is_simulation: shows "simulation resid Resid quot" .. (* * TODO: Show couniversality. *) end subsection "Identities form a Coherent Normal Sub-RTS" text \ We now show that the collection of identities of an RTS form a coherent normal sub-RTS, and that the associated congruence \\\ coincides with \\\. Thus, every RTS can be factored by the relation \\\ to obtain an extensional RTS. Although we could have shown that fact much earlier, we have delayed proving it so that we could simply obtain it as a special case of our general quotient result without redundant work. \ context rts begin interpretation normal_sub_rts resid \Collect ide\ proof show "\t. t \ Collect ide \ arr t" by blast show 1: "\a. ide a \ a \ Collect ide" by blast show "\u t. \u \ Collect ide; coinitial t u\ \ u \\ t \ Collect ide" by (metis 1 CollectD arr_def coinitial_iff con_sym in_sourcesE in_sourcesI resid_ide_arr) show "\u t. \u \ Collect ide; t \\ u \ Collect ide\ \ t \ Collect ide" using ide_backward_stable by blast show "\u t. \u \ Collect ide; seq u t\ \ \v. composite_of u t v" by (metis composite_of_source_arr ide_def in_sourcesI mem_Collect_eq seq_def resid_source_in_targets) show "\u t. \u \ Collect ide; seq t u\ \ \v. composite_of t u v" by (metis arrE composite_of_arr_target in_sourcesI seqE mem_Collect_eq) qed lemma identities_form_normal_sub_rts: shows "normal_sub_rts resid (Collect ide)" .. interpretation coherent_normal_sub_rts resid \Collect ide\ apply unfold_locales by (metis CollectD Cong\<^sub>0_reflexive Cong_closure_props(4) Cong_imp_arr(2) arr_resid_iff_con resid_arr_ide) lemma identities_form_coherent_normal_sub_rts: shows "coherent_normal_sub_rts resid (Collect ide)" .. lemma Cong_iff_cong: shows "Cong t u \ t \ u" by (metis CollectD Cong_def ide_closed resid_arr_ide Cong_closure_props(3) Cong_imp_arr(2) arr_resid_iff_con) end section "Paths" text \ A \emph{path} in an RTS is a nonempty list of arrows such that the set of targets of each arrow suitably matches the set of sources of its successor. The residuation on the given RTS extends inductively to a residuation on paths, so that paths also form an RTS. The append operation on lists yields a composite for each pair of compatible paths. \ locale paths_in_rts = R: rts begin fun Srcs where "Srcs [] = {}" | "Srcs [t] = R.sources t" | "Srcs (t # T) = R.sources t" fun Trgs where "Trgs [] = {}" | "Trgs [t] = R.targets t" | "Trgs (t # T) = Trgs T" fun Arr where "Arr [] = False" | "Arr [t] = R.arr t" | "Arr (t # T) = (R.arr t \ Arr T \ R.targets t \ Srcs T)" fun Ide where "Ide [] = False" | "Ide [t] = R.ide t" | "Ide (t # T) = (R.ide t \ Ide T \ R.targets t \ Srcs T)" lemma set_Arr_subset_arr: shows "Arr T \ set T \ Collect R.arr" apply (induct T) apply auto using Arr.elims(2) apply blast by (metis Arr.simps(3) Ball_Collect list.set_cases) lemma Arr_imp_arr_hd [simp]: assumes "Arr T" shows "R.arr (hd T)" using assms by (metis Arr.simps(1) CollectD hd_in_set set_Arr_subset_arr subset_code(1)) lemma Arr_imp_arr_last [simp]: assumes "Arr T" shows "R.arr (last T)" using assms by (metis Arr.simps(1) CollectD in_mono last_in_set set_Arr_subset_arr) lemma Arr_imp_Arr_tl [simp]: assumes "Arr T" and "tl T \ []" shows "Arr (tl T)" using assms by (metis Arr.simps(3) list.exhaust_sel list.sel(2)) lemma set_Ide_subset_ide: shows "Ide T \ set T \ Collect R.ide" apply (induct T) apply auto using Ide.elims(2) apply blast by (metis Ide.simps(3) Ball_Collect list.set_cases) lemma Ide_imp_Ide_hd [simp]: assumes "Ide T" shows "R.ide (hd T)" using assms by (metis Ide.simps(1) CollectD hd_in_set set_Ide_subset_ide subset_code(1)) lemma Ide_imp_Ide_last [simp]: assumes "Ide T" shows "R.ide (last T)" using assms by (metis Ide.simps(1) CollectD in_mono last_in_set set_Ide_subset_ide) lemma Ide_imp_Ide_tl [simp]: assumes "Ide T" and "tl T \ []" shows "Ide (tl T)" using assms by (metis Ide.simps(3) list.exhaust_sel list.sel(2)) lemma Ide_implies_Arr: shows "Ide T \ Arr T" apply (induct T) apply simp using Ide.elims(2) by fastforce lemma const_ide_is_Ide: shows "\T \ []; R.ide (hd T); set T \ {hd T}\ \ Ide T" apply (induct T) apply auto by (metis Ide.simps(2-3) R.ideE R.sources_resid Srcs.simps(2-3) empty_iff insert_iff list.exhaust_sel list.set_sel(1) order_refl subset_singletonD) lemma Ide_char: shows "Ide T \ Arr T \ set T \ Collect R.ide" apply (induct T) apply auto[1] by (metis Arr.simps(3) Ide.simps(2-3) Ide_implies_Arr empty_subsetI insert_subset list.simps(15) mem_Collect_eq neq_Nil_conv set_empty) lemma IdeI [intro]: assumes "Arr T" and "set T \ Collect R.ide" shows "Ide T" using assms Ide_char by force lemma Arr_has_Src: shows "Arr T \ Srcs T \ {}" apply (cases T) apply simp by (metis R.arr_iff_has_source Srcs.elims Arr.elims(2) list.distinct(1) list.sel(1)) lemma Arr_has_Trg: shows "Arr T \ Trgs T \ {}" using R.arr_iff_has_target apply (induct T) apply simp by (metis Arr.simps(2) Arr.simps(3) Trgs.simps(2-3) list.exhaust_sel) lemma Srcs_are_ide: shows "Srcs T \ Collect R.ide" apply (cases T) apply simp by (metis (no_types, lifting) Srcs.elims list.distinct(1) mem_Collect_eq R.sources_def subsetI) lemma Trgs_are_ide: shows "Trgs T \ Collect R.ide" apply (induct T) apply simp by (metis R.arr_iff_has_target R.sources_resid Srcs.simps(2) Trgs.simps(2-3) Srcs_are_ide empty_subsetI list.exhaust R.arrE) lemma Srcs_are_con: assumes "a \ Srcs T" and "a' \ Srcs T" shows "a \ a'" using assms by (metis Srcs.elims empty_iff R.sources_are_con) lemma Srcs_con_closed: assumes "a \ Srcs T" and "R.ide a'" and "a \ a'" shows "a' \ Srcs T" using assms R.sources_con_closed apply (cases T, auto) by (metis Srcs.simps(2-3) neq_Nil_conv) lemma Srcs_eqI: assumes "Srcs T \ Srcs T' \ {}" shows "Srcs T = Srcs T'" using assms R.sources_eqI apply (cases T; cases T') apply auto apply (metis IntI Srcs.simps(2-3) empty_iff neq_Nil_conv) by (metis Srcs.simps(2-3) assms neq_Nil_conv) lemma Trgs_are_con: shows "\b \ Trgs T; b' \ Trgs T\ \ b \ b'" apply (induct T) apply auto by (metis R.targets_are_con Trgs.simps(2-3) list.exhaust_sel) lemma Trgs_con_closed: shows "\b \ Trgs T; R.ide b'; b \ b'\ \ b' \ Trgs T" apply (induct T) apply auto by (metis R.targets_con_closed Trgs.simps(2-3) neq_Nil_conv) lemma Trgs_eqI: assumes "Trgs T \ Trgs T' \ {}" shows "Trgs T = Trgs T'" using assms Trgs_are_ide Trgs_are_con Trgs_con_closed by blast lemma Srcs_simp\<^sub>P: assumes "Arr T" shows "Srcs T = R.sources (hd T)" using assms by (metis Arr_has_Src Srcs.simps(1) Srcs.simps(2) Srcs.simps(3) list.exhaust_sel) lemma Trgs_simp\<^sub>P: shows "Arr T \ Trgs T = R.targets (last T)" apply (induct T) apply simp by (metis Arr.simps(3) Trgs.simps(2) Trgs.simps(3) last_ConsL last_ConsR neq_Nil_conv) subsection "Residuation on Paths" text \ It was more difficult than I thought to get a correct formal definition for residuation on paths and to prove things from it. Straightforward attempts to write a single recursive definition ran into problems with being able to prove termination, as well as getting the cases correct so that the domain of definition was symmetric. Eventually I found the definition below, which simplifies the termination proof to some extent through the use of two auxiliary functions, and which has a symmetric form that makes symmetry easier to prove. However, there was still some difficulty in proving the recursive expansions with respect to cons and append that I needed. \ text \ The following defines residuation of a single transition along a path, yielding a transition. \ fun Resid1x (infix "\<^sup>1\\\<^sup>*" 70) where "t \<^sup>1\\\<^sup>* [] = R.null" | "t \<^sup>1\\\<^sup>* [u] = t \\ u" | "t \<^sup>1\\\<^sup>* (u # U) = (t \\ u) \<^sup>1\\\<^sup>* U" text \ Next, we have residuation of a path along a single transition, yielding a path. \ fun Residx1 (infix "\<^sup>*\\\<^sup>1" 70) where "[] \<^sup>*\\\<^sup>1 u = []" | "[t] \<^sup>*\\\<^sup>1 u = (if t \ u then [t \\ u] else [])" | "(t # T) \<^sup>*\\\<^sup>1 u = (if t \ u \ T \<^sup>*\\\<^sup>1 (u \\ t) \ [] then (t \\ u) # T \<^sup>*\\\<^sup>1 (u \\ t) else [])" text \ Finally, residuation of a path along a path, yielding a path. \ function (sequential) Resid (infix "\<^sup>*\\\<^sup>*" 70) where "[] \<^sup>*\\\<^sup>* _ = []" | "_ \<^sup>*\\\<^sup>* [] = []" | "[t] \<^sup>*\\\<^sup>* [u] = (if t \ u then [t \\ u] else [])" | "[t] \<^sup>*\\\<^sup>* (u # U) = (if t \ u \ (t \\ u) \<^sup>1\\\<^sup>* U \ R.null then [(t \\ u) \<^sup>1\\\<^sup>* U] else [])" | "(t # T) \<^sup>*\\\<^sup>* [u] = (if t \ u \ T \<^sup>*\\\<^sup>1 (u \\ t) \ [] then (t \\ u) # (T \<^sup>*\\\<^sup>1 (u \\ t)) else [])" | "(t # T) \<^sup>*\\\<^sup>* (u # U) = (if t \ u \ (t \\ u) \<^sup>1\\\<^sup>* U \ R.null \ (T \<^sup>*\\\<^sup>1 (u \\ t)) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>1 (t \\ u)) \ [] then (t \\ u) \<^sup>1\\\<^sup>* U # (T \<^sup>*\\\<^sup>1 (u \\ t)) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>1 (t \\ u)) else [])" by pat_completeness auto text \ Residuation of a path along a single transition is length non-increasing. Actually, it is length-preserving, except in case the path and the transition are not consistent. We will show that later, but for now this is what we need to establish termination for (\\\). \ lemma length_Residx1: shows "length (T \<^sup>*\\\<^sup>1 u) \ length T" proof (induct T arbitrary: u) show "\u. length ([] \<^sup>*\\\<^sup>1 u) \ length []" by simp fix t T u assume ind: "\u. length (T \<^sup>*\\\<^sup>1 u) \ length T" show "length ((t # T) \<^sup>*\\\<^sup>1 u) \ length (t # T)" using ind by (cases T, cases "t \ u", cases "T \<^sup>*\\\<^sup>1 (u \\ t)") auto qed termination Resid proof (relation "measure (\(T, U). length T + length U)") show "wf (measure (\(T, U). length T + length U))" by simp fix t t' T u U have "length ((t' # T) \<^sup>*\\\<^sup>1 (u \\ t)) + length (U \<^sup>*\\\<^sup>1 (t \\ u)) < length (t # t' # T) + length (u # U)" using length_Residx1 by (metis add_less_le_mono impossible_Cons le_neq_implies_less list.size(4) trans_le_add1) thus 1: "(((t' # T) \<^sup>*\\\<^sup>1 (u \\ t), U \<^sup>*\\\<^sup>1 (t \\ u)), t # t' # T, u # U) \ measure (\(T, U). length T + length U)" by simp show "(((t' # T) \<^sup>*\\\<^sup>1 (u \\ t), U \<^sup>*\\\<^sup>1 (t \\ u)), t # t' # T, u # U) \ measure (\(T, U). length T + length U)" using 1 length_Residx1 by blast have "length (T \<^sup>*\\\<^sup>1 (u \\ t)) + length (U \<^sup>*\\\<^sup>1 (t \\ u)) \ length T + length U" using length_Residx1 by (simp add: add_mono) thus 2: "((T \<^sup>*\\\<^sup>1 (u \\ t), U \<^sup>*\\\<^sup>1 (t \\ u)), t # T, u # U) \ measure (\(T, U). length T + length U)" by simp show "((T \<^sup>*\\\<^sup>1 (u \\ t), U \<^sup>*\\\<^sup>1 (t \\ u)), t # T, u # U) \ measure (\(T, U). length T + length U)" using 2 length_Residx1 by blast qed lemma Resid1x_null: shows "R.null \<^sup>1\\\<^sup>* T = R.null" apply (induct T) apply auto by (metis R.null_is_zero(1) Resid1x.simps(2-3) list.exhaust) lemma Resid1x_ide: shows "\R.ide a; a \<^sup>1\\\<^sup>* T \ R.null\ \ R.ide (a \<^sup>1\\\<^sup>* T)" proof (induct T arbitrary: a) show "\a. a \<^sup>1\\\<^sup>* [] \ R.null \ R.ide (a \<^sup>1\\\<^sup>* [])" by simp fix a t T assume a: "R.ide a" assume ind: "\a. \R.ide a; a \<^sup>1\\\<^sup>* T \ R.null\ \ R.ide (a \<^sup>1\\\<^sup>* T)" assume con: "a \<^sup>1\\\<^sup>* (t # T) \ R.null" have 1: "a \ t" using con by (metis R.con_def Resid1x.simps(2-3) Resid1x_null list.exhaust) show "R.ide (a \<^sup>1\\\<^sup>* (t # T))" using a 1 con ind R.resid_ide_arr by (metis Resid1x.simps(2-3) list.exhaust) qed (* * TODO: Try to make this a definition, rather than an abbreviation. * * I made an attempt at this, but there are many, many places where the * definition needs to be unwound. It is not clear how valuable it might * end up being to have this as a definition. *) abbreviation Con (infix "\<^sup>*\\<^sup>*" 50) where "T \<^sup>*\\<^sup>* U \ T \<^sup>*\\\<^sup>* U \ []" lemma Con_sym1: shows "T \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* T \ R.null" proof (induct T arbitrary: u) show "\u. [] \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* [] \ R.null" by simp show "\t T u. (\u. T \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* T \ R.null) \ (t # T) \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* (t # T) \ R.null" proof - fix t T u assume ind: "\u. T \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* T \ R.null" show "(t # T) \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* (t # T) \ R.null" proof show "(t # T) \<^sup>*\\\<^sup>1 u \ [] \ u \<^sup>1\\\<^sup>* (t # T) \ R.null" by (metis R.con_sym Resid1x.simps(2-3) Residx1.simps(2-3) ind neq_Nil_conv R.conE) show "u \<^sup>1\\\<^sup>* (t # T) \ R.null \ (t # T) \<^sup>*\\\<^sup>1 u \ []" using ind R.con_sym apply (cases T) apply auto by (metis R.conI Resid1x_null) qed qed qed lemma Con_sym_ind: shows "length T + length U \ n \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" proof (induct n arbitrary: T U) show "\T U. length T + length U \ 0 \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" by simp fix n and T U :: "'a list" assume ind: "\T U. length T + length U \ n \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" assume 1: "length T + length U \ Suc n" show "T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" using R.con_sym Con_sym1 apply (cases T; cases U) apply auto[3] proof - fix t u T' U' assume T: "T = t # T'" and U: "U = u # U'" show "T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" proof (cases "T' = []") show "T' = [] \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" using T U Con_sym1 R.con_sym by (cases U') auto show "T' \ [] \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" proof (cases "U' = []") show "\T' \ []; U' = []\ \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" using T U R.con_sym Con_sym1 by (cases T') auto show "\T' \ []; U' \ []\ \ T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" proof - assume T': "T' \ []" and U': "U' \ []" have 2: "length (U' \<^sup>*\\\<^sup>1 (t \\ u)) + length (T' \<^sup>*\\\<^sup>1 (u \\ t)) \ n" proof - have "length (U' \<^sup>*\\\<^sup>1 (t \\ u)) + length (T' \<^sup>*\\\<^sup>1 (u \\ t)) \ length U' + length T'" by (simp add: add_le_mono length_Residx1) also have "... \ length T' + length U'" using T' add.commute not_less_eq_eq by auto also have "... \ n" using 1 T U by simp finally show ?thesis by blast qed show "T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" proof assume Con: "T \<^sup>*\\<^sup>* U" have 3: "t \ u \ T' \<^sup>*\\\<^sup>1 (u \\ t) \ [] \ (t \\ u) \<^sup>1\\\<^sup>* U' \ R.null \ (T' \<^sup>*\\\<^sup>1 (u \\ t)) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>1 (t \\ u)) \ []" using Con T U T' U' Con_sym1 apply (cases T', cases U') apply simp_all by (metis Resid.simps(1) Resid.simps(6) neq_Nil_conv) hence "u \ t \ U' \<^sup>*\\\<^sup>1 (t \\ u) \ [] \ (u \\ t) \<^sup>1\\\<^sup>* T' \ R.null" using T' U' R.con_sym Con_sym1 by simp moreover have "(U' \<^sup>*\\\<^sup>1 (t \\ u)) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>1 (u \\ t)) \ []" using 2 3 ind by simp ultimately show "U \<^sup>*\\<^sup>* T" using T U T' U' by (cases T'; cases U') auto next assume Con: "U \<^sup>*\\<^sup>* T" have 3: "u \ t \ U' \<^sup>*\\\<^sup>1 (t \\ u) \ [] \ (u \\ t) \<^sup>1\\\<^sup>* T' \ R.null \ (U' \<^sup>*\\\<^sup>1 (t \\ u)) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>1 (u \\ t)) \ []" using Con T U T' U' Con_sym1 apply (cases T'; cases U') apply auto apply argo by force hence "t \ u \ T' \<^sup>*\\\<^sup>1 (u \\ t) \ [] \ (t \\ u) \<^sup>1\\\<^sup>* U' \ R.null" using T' U' R.con_sym Con_sym1 by simp moreover have "(T' \<^sup>*\\\<^sup>1 (u \\ t)) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>1 (t \\ u)) \ []" using 2 3 ind by simp ultimately show "T \<^sup>*\\<^sup>* U" using T U T' U' by (cases T'; cases U') auto qed qed qed qed qed qed lemma Con_sym: shows "T \<^sup>*\\<^sup>* U \ U \<^sup>*\\<^sup>* T" using Con_sym_ind by blast lemma Residx1_as_Resid: shows "T \<^sup>*\\\<^sup>1 u = T \<^sup>*\\\<^sup>* [u]" proof (induct T) show "[] \<^sup>*\\\<^sup>1 u = [] \<^sup>*\\\<^sup>* [u]" by simp fix t T assume ind: "T \<^sup>*\\\<^sup>1 u = T \<^sup>*\\\<^sup>* [u]" show "(t # T) \<^sup>*\\\<^sup>1 u = (t # T) \<^sup>*\\\<^sup>* [u]" by (cases T) auto qed lemma Resid1x_as_Resid': shows "t \<^sup>1\\\<^sup>* U = (if [t] \<^sup>*\\\<^sup>* U \ [] then hd ([t] \<^sup>*\\\<^sup>* U) else R.null)" proof (induct U) show "t \<^sup>1\\\<^sup>* [] = (if [t] \<^sup>*\\\<^sup>* [] \ [] then hd ([t] \<^sup>*\\\<^sup>* []) else R.null)" by simp fix u U assume ind: "t \<^sup>1\\\<^sup>* U = (if [t] \<^sup>*\\\<^sup>* U \ [] then hd ([t] \<^sup>*\\\<^sup>* U) else R.null)" show "t \<^sup>1\\\<^sup>* (u # U) = (if [t] \<^sup>*\\\<^sup>* (u # U) \ [] then hd ([t] \<^sup>*\\\<^sup>* (u # U)) else R.null)" using Resid1x_null by (cases U) auto qed text \ The following recursive expansion for consistency of paths is an intermediate result that is not yet quite in the form we really want. \ lemma Con_rec: shows "[t] \<^sup>*\\<^sup>* [u] \ t \ u" and "T \ [] \ t # T \<^sup>*\\<^sup>* [u] \ t \ u \ T \<^sup>*\\<^sup>* [u \\ t]" and "U \ [] \ [t] \<^sup>*\\<^sup>* (u # U) \ t \ u \ [t \\ u] \<^sup>*\\<^sup>* U" and "\T \ []; U \ []\ \ t # T \<^sup>*\\<^sup>* u # U \ t \ u \ T \<^sup>*\\<^sup>* [u \\ t] \ [t \\ u] \<^sup>*\\<^sup>* U \ T \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t \\ u]" proof - show "[t] \<^sup>*\\<^sup>* [u] \ t \ u" by simp show "T \ [] \ t # T \<^sup>*\\<^sup>* [u] \ t \ u \ T \<^sup>*\\<^sup>* [u \\ t]" using Residx1_as_Resid by (cases T) auto show "U \ [] \ [t] \<^sup>*\\<^sup>* (u # U) \ t \ u \ [t \\ u] \<^sup>*\\<^sup>* U" using Resid1x_as_Resid' Con_sym Con_sym1 Resid1x.simps(3) Residx1_as_Resid by (cases U) auto show "\T \ []; U \ []\ \ t # T \<^sup>*\\<^sup>* u # U \ t \ u \ T \<^sup>*\\<^sup>* [u \\ t] \ [t \\ u] \<^sup>*\\<^sup>* U \ T \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t \\ u]" using Residx1_as_Resid Resid1x_as_Resid' Con_sym1 Con_sym R.con_sym by (cases T; cases U) auto qed text \ This version is a more appealing form of the previously proved fact \Resid1x_as_Resid'\. \ lemma Resid1x_as_Resid: assumes "[t] \<^sup>*\\\<^sup>* U \ []" shows "[t] \<^sup>*\\\<^sup>* U = [t \<^sup>1\\\<^sup>* U]" using assms Con_rec(2,4) apply (cases U; cases "tl U") apply auto by argo+ (* TODO: Why can auto no longer complete this proof? *) text \ The following is an intermediate version of a recursive expansion for residuation, to be improved subsequently. \ lemma Resid_rec: shows [simp]: "[t] \<^sup>*\\<^sup>* [u] \ [t] \<^sup>*\\\<^sup>* [u] = [t \\ u]" and "\T \ []; t # T \<^sup>*\\<^sup>* [u]\ \ (t # T) \<^sup>*\\\<^sup>* [u] = (t \\ u) # (T \<^sup>*\\\<^sup>* [u \\ t])" and "\U \ []; Con [t] (u # U)\ \ [t] \<^sup>*\\\<^sup>* (u # U) = [t \\ u] \<^sup>*\\\<^sup>* U" and "\T \ []; U \ []; Con (t # T) (u # U)\ \ (t # T) \<^sup>*\\\<^sup>* (u # U) = ([t \\ u] \<^sup>*\\\<^sup>* U) @ ((T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u]))" proof - show "[t] \<^sup>*\\<^sup>* [u] \ Resid [t] [u] = [t \\ u]" by (meson Resid.simps(3)) show "\T \ []; t # T \<^sup>*\\<^sup>* [u]\ \ (t # T) \<^sup>*\\\<^sup>* [u] = (t \\ u) # (T \<^sup>*\\\<^sup>* [u \\ t])" using Residx1_as_Resid by (metis Residx1.simps(3) list.exhaust_sel) show 1: "\U \ []; [t] \<^sup>*\\<^sup>* u # U\ \ [t] \<^sup>*\\\<^sup>* (u # U) = [t \\ u] \<^sup>*\\\<^sup>* U" by (metis Con_rec(3) Resid1x.simps(3) Resid1x_as_Resid list.exhaust) show "\T \ []; U \ []; t # T \<^sup>*\\<^sup>* u # U\ \ (t # T) \<^sup>*\\\<^sup>* (u # U) = ([t \\ u] \<^sup>*\\\<^sup>* U) @ ((T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u]))" proof - assume T: "T \ []" and U: "U \ []" and Con: "Con (t # T) (u # U)" have tu: "t \ u" using Con Con_rec by metis have "(t # T) \<^sup>*\\\<^sup>* (u # U) = ((t \\ u) \<^sup>1\\\<^sup>* U) # ((T \<^sup>*\\\<^sup>1 (u \\ t)) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>1 (t \\ u)))" using T U Con tu by (cases T; cases U) auto also have "... = ([t \\ u] \<^sup>*\\\<^sup>* U) @ ((T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u]))" using T U Con tu Con_rec(4) Resid1x_as_Resid Residx1_as_Resid by force finally show ?thesis by simp qed qed text \ For consistent paths, residuation is length-preserving. \ lemma length_Resid_ind: shows "\length T + length U \ n; T \<^sup>*\\<^sup>* U\ \ length (T \<^sup>*\\\<^sup>* U) = length T" apply (induct n arbitrary: T U) apply simp proof - fix n T U assume ind: "\T U. \length T + length U \ n; T \<^sup>*\\<^sup>* U\ \ length (T \<^sup>*\\\<^sup>* U) = length T" assume Con: "T \<^sup>*\\<^sup>* U" assume len: "length T + length U \ Suc n" show "length (T \<^sup>*\\\<^sup>* U) = length T" using Con len ind Resid1x_as_Resid length_Cons Con_rec(2) Resid_rec(2) apply (cases T; cases U) apply auto apply (cases "tl T = []"; cases "tl U = []") apply auto apply metis apply fastforce proof - fix t T' u U' assume T: "T = t # T'" and U: "U = u # U'" assume T': "T' \ []" and U': "U' \ []" show "length ((t # T') \<^sup>*\\\<^sup>* (u # U')) = Suc (length T')" using Con Con_rec(4) Con_sym Resid_rec(4) T T' U U' ind len by auto qed qed lemma length_Resid: assumes "T \<^sup>*\\<^sup>* U" shows "length (T \<^sup>*\\\<^sup>* U) = length T" using assms length_Resid_ind by auto lemma Con_initial_left: shows "t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U" apply (induct U) apply simp by (metis Con_rec(1-4)) lemma Con_initial_right: shows "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u]" apply (induct T) apply simp by (metis Con_rec(1-4)) lemma Resid_cons_ind: shows "\T \ []; U \ []; length T + length U \ n\ \ (\t. t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U) \ (\t. t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U)" proof (induct n arbitrary: T U) show "\T U. \T \ []; U \ []; length T + length U \ 0\ \ (\t. t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U) \ (\t. t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U)" by simp fix n and T U :: "'a list" assume ind: "\T U. \T \ []; U \ []; length T + length U \ n\ \ (\t. t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U) \ (\t. t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U)" assume T: "T \ []" and U: "U \ []" assume len: "length T + length U \ Suc n" show "(\t. t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U) \ (\t. t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \ (\u. T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U)" proof (intro allI conjI iffI impI) fix t show 1: "t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" proof (cases U) show "U = [] \ ?thesis" using U by simp fix u U' assume U: "U = u # U'" assume Con: "t # T \<^sup>*\\<^sup>* U" show ?thesis proof (cases "U' = []") show "U' = [] \ ?thesis" using T U Con R.con_sym Con_rec(2) Resid_rec(2) by auto assume U': "U' \ []" have "(t # T) \<^sup>*\\\<^sup>* U = [t \\ u] \<^sup>*\\\<^sup>* U' @ (T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" using T U U' Con Resid_rec(4) by fastforce also have 1: "... = [t] \<^sup>*\\\<^sup>* U @ (T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" using T U U' Con Con_rec(3-4) Resid_rec(3) by auto also have "... = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* ((u \\ t) # (U' \<^sup>*\\\<^sup>* [t \\ u]))" proof - have "T \<^sup>*\\\<^sup>* ((u \\ t) # (U' \<^sup>*\\\<^sup>* [t \\ u])) = (T \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" using T U U' ind [of T "U' \<^sup>*\\\<^sup>* [t \\ u]"] Con Con_rec(4) Con_sym len length_Resid by fastforce thus ?thesis by auto qed also have "... = [t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" using T U U' 1 Con Con_rec(4) Con_sym1 Residx1_as_Resid Resid1x_as_Resid Resid_rec(2) Con_sym Con_initial_left by auto finally show ?thesis by simp qed qed show "t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U" by (simp add: Con_initial_left) show "t # T \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" by (metis "1" Suc_inject T append_Nil2 length_0_conv length_Cons length_Resid) show "[t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t] \ t # T \<^sup>*\\<^sup>* U" proof (cases U) show "\[t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]; U = []\ \ t # T \<^sup>*\\<^sup>* U" using U by simp fix u U' assume U: "U = u # U'" assume Con: "[t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]" show "t # T \<^sup>*\\<^sup>* U" proof (cases "U' = []") show "U' = [] \ ?thesis" using T U Con by (metis Con_rec(2) Resid.simps(3) R.con_sym) assume U': "U' \ []" show ?thesis proof - have "t \ u" using T U U' Con Con_rec(3) by blast moreover have "T \<^sup>*\\<^sup>* [u \\ t]" using T U U' Con Con_initial_right Con_sym1 Residx1_as_Resid Resid1x_as_Resid Resid_rec(2) by (metis Con_sym) moreover have "[t \\ u] \<^sup>*\\<^sup>* U'" using T U U' Con Resid_rec(3) by force moreover have "T \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" by (metis (no_types, opaque_lifting) Con Con_sym Resid_rec(2) Suc_le_mono T U U' add_Suc_right calculation(3) ind len length_Cons length_Resid) ultimately show ?thesis using T U U' Con_rec(4) by simp qed qed qed next fix u show 1: "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U" proof (cases T) show 2: "\T \<^sup>*\\<^sup>* u # U; T = []\ \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U" using T by simp fix t T' assume T: "T = t # T'" assume Con: "T \<^sup>*\\<^sup>* u # U" show ?thesis proof (cases "T' = []") show "T' = [] \ ?thesis" using T U Con Con_rec(3) Resid1x_as_Resid Resid_rec(3) by force assume T': "T' \ []" have "T \<^sup>*\\\<^sup>* (u # U) = [t \\ u] \<^sup>*\\\<^sup>* U @ (T' \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u])" using T U T' Con Resid_rec(4) [of T' U t u] by simp also have "... = ((t \\ u) # (T' \<^sup>*\\\<^sup>* [u \\ t])) \<^sup>*\\\<^sup>* U" proof - have "length (T' \<^sup>*\\\<^sup>* [u \\ t]) + length U \ n" by (metis (no_types, lifting) Con Con_rec(4) One_nat_def Suc_eq_plus1 Suc_leI T T' U add_Suc le_less_trans len length_Resid lessI list.size(4) not_le) thus ?thesis using ind [of "T' \<^sup>*\\\<^sup>* [u \\ t]" U] Con Con_rec(4) T T' U by auto qed also have "... = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U" using T U T' Con Con_rec(2,4) Resid_rec(2) by force finally show ?thesis by simp qed qed show "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u]" using 1 by force show "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U" using 1 by fastforce show "T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* u # U" proof (cases T) show "\T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U; T = []\ \ T \<^sup>*\\<^sup>* u # U" using T by simp fix t T' assume T: "T = t # T'" assume Con: "T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U" show "Con T (u # U)" proof (cases "T' = []") show "T' = [] \ ?thesis" using Con T U Con_rec(1,3) by auto assume T': "T' \ []" have "t \ u" using Con T U T' Con_rec(2) by blast moreover have 2: "T' \<^sup>*\\<^sup>* [u \\ t]" using Con T U T' Con_rec(2) by blast moreover have "[t \\ u] \<^sup>*\\<^sup>* U" using Con T U T' by (metis Con_initial_left Resid_rec(2)) moreover have "T' \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t \\ u]" proof - have 0: "length (U \<^sup>*\\\<^sup>* [t \\ u]) = length U" using Con T U T' length_Resid Con_sym calculation(3) by blast hence 1: "length T' + length (U \<^sup>*\\\<^sup>* [t \\ u]) \ n" using Con T U T' len length_Resid Con_sym by simp have "length ((T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U) = length ([t \\ u] \<^sup>*\\\<^sup>* U) + length ((T' \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u]))" proof - have "(T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U = [t \\ u] \<^sup>*\\\<^sup>* U @ (T' \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t \\ u])" by (metis 0 1 2 Con Resid_rec(2) T T' U ind length_Resid) thus ?thesis using Con T U T' length_Resid by simp qed moreover have "length ((T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U) = length T" using Con T U T' length_Resid by metis moreover have "length ([t \\ u] \<^sup>*\\\<^sup>* U) \ 1" using Con T U T' Resid1x_as_Resid by (metis One_nat_def length_Cons list.size(3) order_refl zero_le) ultimately show ?thesis using Con T U T' length_Resid by auto qed ultimately show "T \<^sup>*\\<^sup>* u # U" using T Con_rec(4) [of T' U t u] by fastforce qed qed qed qed text \ The following are the final versions of recursive expansion for consistency and residuation on paths. These are what I really wanted the original definitions to look like, but if this is tried, then \Con\ and \Resid\ end up having to be mutually recursive, expressing the definitions so that they are single-valued becomes an issue, and proving termination is more problematic. \ lemma Con_cons: assumes "T \ []" and "U \ []" shows "t # T \<^sup>*\\<^sup>* U \ [t] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]" and "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\<^sup>* [u] \ T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U" using assms Resid_cons_ind [of T U] by blast+ lemma Con_consI [intro, simp]: shows "\T \ []; U \ []; [t] \<^sup>*\\<^sup>* U; T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]\ \ t # T \<^sup>*\\<^sup>* U" and "\T \ []; U \ []; T \<^sup>*\\<^sup>* [u]; T \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* U\ \ T \<^sup>*\\<^sup>* u # U" using Con_cons by auto (* TODO: Making this a simp currently seems to produce undesirable breakage. *) lemma Resid_cons: assumes "U \ []" shows "t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = ([t] \<^sup>*\\\<^sup>* U) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" and "T \<^sup>*\\<^sup>* u # U \ T \<^sup>*\\\<^sup>* (u # U) = (T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U" using assms Resid_cons_ind [of T U] Resid.simps(1) by blast+ text \ The following expansion of residuation with respect to the first argument is stated in terms of the more primitive cons, rather than list append, but as a result \\<^sup>1\\<^sup>*\ has to be used. \ (* TODO: Making this a simp seems to produce similar breakage as above. *) lemma Resid_cons': assumes "T \ []" shows "t # T \<^sup>*\\<^sup>* U \ (t # T) \<^sup>*\\\<^sup>* U = (t \<^sup>1\\\<^sup>* U) # (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" using assms by (metis Con_sym Resid.simps(1) Resid1x_as_Resid Resid_cons(1) append_Cons append_Nil) lemma Srcs_Resid_Arr_single: assumes "T \<^sup>*\\<^sup>* [u]" shows "Srcs (T \<^sup>*\\\<^sup>* [u]) = R.targets u" proof (cases T) show "T = [] \ Srcs (T \<^sup>*\\\<^sup>* [u]) = R.targets u" using assms by simp fix t T' assume T: "T = t # T'" show "Srcs (T \<^sup>*\\\<^sup>* [u]) = R.targets u" proof (cases "T' = []") show "T' = [] \ ?thesis" using assms T R.sources_resid by auto assume T': "T' \ []" have "Srcs (T \<^sup>*\\\<^sup>* [u]) = Srcs ((t # T') \<^sup>*\\\<^sup>* [u])" using T by simp also have "... = Srcs ((t \\ u) # (T' \<^sup>*\\\<^sup>* ([u] \<^sup>*\\\<^sup>* T')))" using assms T by (metis Resid_rec(2) Srcs.elims T' list.distinct(1) list.sel(1)) also have "... = R.sources (t \\ u)" using Srcs.elims by blast also have "... = R.targets u" using assms Con_rec(2) T T' R.sources_resid by force finally show ?thesis by blast qed qed lemma Srcs_Resid_single_Arr: shows "[u] \<^sup>*\\<^sup>* T \ Srcs ([u] \<^sup>*\\\<^sup>* T) = Trgs T" proof (induct T arbitrary: u) show "\u. [u] \<^sup>*\\<^sup>* [] \ Srcs ([u] \<^sup>*\\\<^sup>* []) = Trgs []" by simp fix t u T assume ind: "\u. [u] \<^sup>*\\<^sup>* T \ Srcs ([u] \<^sup>*\\\<^sup>* T) = Trgs T" assume Con: "[u] \<^sup>*\\<^sup>* t # T" show "Srcs ([u] \<^sup>*\\\<^sup>* (t # T)) = Trgs (t # T)" proof (cases "T = []") show "T = [] \ ?thesis" using Con Srcs_Resid_Arr_single Trgs.simps(2) by presburger assume T: "T \ []" have "Srcs ([u] \<^sup>*\\\<^sup>* (t # T)) = Srcs ([u \\ t] \<^sup>*\\\<^sup>* T)" using Con Resid_rec(3) T by force also have "... = Trgs T" using Con ind Con_rec(3) T by auto also have "... = Trgs (t # T)" by (metis T Trgs.elims Trgs.simps(3)) finally show ?thesis by simp qed qed lemma Trgs_Resid_sym_Arr_single: shows "T \<^sup>*\\<^sup>* [u] \ Trgs (T \<^sup>*\\\<^sup>* [u]) = Trgs ([u] \<^sup>*\\\<^sup>* T)" proof (induct T arbitrary: u) show "\u. [] \<^sup>*\\<^sup>* [u] \ Trgs ([] \<^sup>*\\\<^sup>* [u]) = Trgs ([u] \<^sup>*\\\<^sup>* [])" by simp fix t u T assume ind: "\u. T \<^sup>*\\<^sup>* [u] \ Trgs (T \<^sup>*\\\<^sup>* [u]) = Trgs ([u] \<^sup>*\\\<^sup>* T)" assume Con: "t # T \<^sup>*\\<^sup>* [u]" show "Trgs ((t # T) \<^sup>*\\\<^sup>* [u]) = Trgs ([u] \<^sup>*\\\<^sup>* (t # T))" proof (cases "T = []") show "T = [] \ ?thesis" using R.targets_resid_sym by (simp add: R.con_sym) assume T: "T \ []" show ?thesis proof - have "Trgs ((t # T) \<^sup>*\\\<^sup>* [u]) = Trgs ((t \\ u) # (T \<^sup>*\\\<^sup>* [u \\ t]))" using Con Resid_rec(2) T by auto also have "... = Trgs (T \<^sup>*\\\<^sup>* [u \\ t])" using T Con Con_rec(2) [of T t u] by (metis Trgs.elims Trgs.simps(3)) also have "... = Trgs ([u \\ t] \<^sup>*\\\<^sup>* T)" using T Con ind Con_sym by metis also have "... = Trgs ([u] \<^sup>*\\\<^sup>* (t # T))" using T Con Con_sym Resid_rec(3) by presburger finally show ?thesis by blast qed qed qed lemma Srcs_Resid [simp]: shows "T \<^sup>*\\<^sup>* U \ Srcs (T \<^sup>*\\\<^sup>* U) = Trgs U" proof (induct U arbitrary: T) show "\T. T \<^sup>*\\<^sup>* [] \ Srcs (T \<^sup>*\\\<^sup>* []) = Trgs []" using Con_sym Resid.simps(1) by blast fix u U T assume ind: "\T. T \<^sup>*\\<^sup>* U \ Srcs (T \<^sup>*\\\<^sup>* U) = Trgs U" assume Con: "T \<^sup>*\\<^sup>* u # U" show "Srcs (T \<^sup>*\\\<^sup>* (u # U)) = Trgs (u # U)" by (metis Con Resid_cons(2) Srcs_Resid_Arr_single Trgs.simps(2-3) ind list.exhaust_sel) qed lemma Trgs_Resid_sym [simp]: shows "T \<^sup>*\\<^sup>* U \ Trgs (T \<^sup>*\\\<^sup>* U) = Trgs (U \<^sup>*\\\<^sup>* T)" proof (induct U arbitrary: T) show "\T. T \<^sup>*\\<^sup>* [] \ Trgs (T \<^sup>*\\\<^sup>* []) = Trgs ([] \<^sup>*\\\<^sup>* T)" by (meson Con_sym Resid.simps(1)) fix u U T assume ind: "\T. T \<^sup>*\\<^sup>* U \ Trgs (T \<^sup>*\\\<^sup>* U) = Trgs (U \<^sup>*\\\<^sup>* T)" assume Con: "T \<^sup>*\\<^sup>* u # U" show "Trgs (T \<^sup>*\\\<^sup>* (u # U)) = Trgs ((u # U) \<^sup>*\\\<^sup>* T)" proof (cases "U = []") show "U = [] \ ?thesis" using Con Trgs_Resid_sym_Arr_single by blast assume U: "U \ []" show ?thesis proof - have "Trgs (T \<^sup>*\\\<^sup>* (u # U)) = Trgs ((T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U)" using U by (metis Con Resid_cons(2)) also have "... = Trgs (U \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [u]))" using U Con by (metis Con_sym ind) also have "... = Trgs ((u # U) \<^sup>*\\\<^sup>* T)" by (metis (no_types, opaque_lifting) Con_cons(1) Con_sym Resid.simps(1) Resid_cons' Trgs.simps(3) U neq_Nil_conv) finally show ?thesis by simp qed qed qed lemma img_Resid_Srcs: shows "Arr T \ (\a. [a] \<^sup>*\\\<^sup>* T) ` Srcs T \ (\b. [b]) ` Trgs T" proof (induct T) show "Arr [] \ (\a. [a] \<^sup>*\\\<^sup>* []) ` Srcs [] \ (\b. [b]) ` Trgs []" by simp fix t :: 'a and T :: "'a list" assume tT: "Arr (t # T)" assume ind: "Arr T \ (\a. [a] \<^sup>*\\\<^sup>* T) ` Srcs T \ (\b. [b]) ` Trgs T" show "(\a. [a] \<^sup>*\\\<^sup>* (t # T)) ` Srcs (t # T) \ (\b. [b]) ` Trgs (t # T)" proof fix B assume B: "B \ (\a. [a] \<^sup>*\\\<^sup>* (t # T)) ` Srcs (t # T)" show "B \ (\b. [b]) ` Trgs (t # T)" proof (cases "T = []") assume T: "T = []" obtain a where a: "a \ R.sources t \ [a \\ t] = B" by (metis (no_types, lifting) B R.composite_of_source_arr R.con_prfx_composite_of(1) Resid_rec(1) Srcs.simps(2) T Arr.simps(2) Con_rec(1) imageE tT) have "a \\ t \ Trgs (t # T)" using tT T a by (simp add: R.resid_source_in_targets) thus ?thesis using B a image_iff by fastforce next assume T: "T \ []" obtain a where a: "a \ R.sources t \ [a] \<^sup>*\\\<^sup>* (t # T) = B" using tT T B Srcs.elims by blast have "[a \\ t] \<^sup>*\\\<^sup>* T = B" using tT T B a by (metis Con_rec(3) R.arrI R.resid_source_in_targets R.targets_are_cong Resid_rec(3) R.arr_resid_iff_con R.ide_implies_arr) moreover have "a \\ t \ Srcs T" using a tT by (metis Arr.simps(3) R.resid_source_in_targets T neq_Nil_conv subsetD) ultimately show ?thesis using T tT ind by (metis Trgs.simps(3) Arr.simps(3) image_iff list.exhaust_sel subsetD) qed qed qed lemma Resid_Arr_Src: shows "\Arr T; a \ Srcs T\ \ T \<^sup>*\\\<^sup>* [a] = T" proof (induct T arbitrary: a) show "\a. \Arr []; a \ Srcs []\ \ [] \<^sup>*\\\<^sup>* [a] = []" by simp fix a t T assume ind: "\a. \Arr T; a \ Srcs T\ \ T \<^sup>*\\\<^sup>* [a] = T" assume Arr: "Arr (t # T)" assume a: "a \ Srcs (t # T)" show "(t # T) \<^sup>*\\\<^sup>* [a] = t # T" proof (cases "T = []") show "T = [] \ ?thesis" using a R.resid_arr_ide R.sources_def by auto assume T: "T \ []" show "(t # T) \<^sup>*\\\<^sup>* [a] = t # T" proof - have 1: "R.arr t \ Arr T \ R.targets t \ Srcs T" using Arr T by (metis Arr.elims(2) list.sel(1) list.sel(3)) have 2: "t # T \<^sup>*\\<^sup>* [a]" using T a Arr Con_rec(2) by (metis (no_types, lifting) img_Resid_Srcs Con_sym imageE image_subset_iff list.distinct(1)) have "(t # T) \<^sup>*\\\<^sup>* [a] = (t \\ a) # (T \<^sup>*\\\<^sup>* [a \\ t])" using 2 T Resid_rec(2) by simp moreover have "t \\ a = t" using Arr a R.sources_def by (metis "2" CollectD Con_rec(2) T Srcs_are_ide in_mono R.resid_arr_ide) moreover have "T \<^sup>*\\\<^sup>* [a \\ t] = T" by (metis "1" "2" R.in_sourcesI R.resid_source_in_targets Srcs_are_ide T a Con_rec(2) in_mono ind mem_Collect_eq) ultimately show ?thesis by simp qed qed qed lemma Con_single_ide_ind: shows "R.ide a \ [a] \<^sup>*\\<^sup>* T \ Arr T \ a \ Srcs T" proof (induct T arbitrary: a) show "\a. [a] \<^sup>*\\<^sup>* [] \ Arr [] \ a \ Srcs []" by simp fix a t T assume ind: "\a. R.ide a \ [a] \<^sup>*\\<^sup>* T \ Arr T \ a \ Srcs T" assume a: "R.ide a" show "[a] \<^sup>*\\<^sup>* (t # T) \ Arr (t # T) \ a \ Srcs (t # T)" proof (cases "T = []") show "T = [] \ ?thesis" using a Con_sym by (metis Arr.simps(2) Resid_Arr_Src Srcs.simps(2) R.arr_iff_has_source Con_rec(1) empty_iff R.in_sourcesI list.distinct(1)) assume T: "T \ []" have 1: "[a] \<^sup>*\\<^sup>* (t # T) \ a \ t \ [a \\ t] \<^sup>*\\<^sup>* T" using a T Con_cons(2) [of "[a]" T t] by simp also have 2: "... \ a \ t \ Arr T \ a \\ t \ Srcs T" using a T ind R.resid_ide_arr by blast also have "... \ Arr (t # T) \ a \ Srcs (t # T)" using a T Con_sym R.con_sym Resid_Arr_Src R.con_implies_arr Srcs_are_ide apply (cases T) apply simp by (metis Arr.simps(3) R.resid_arr_ide R.targets_resid_sym Srcs.simps(3) Srcs_Resid_Arr_single calculation dual_order.eq_iff list.distinct(1) R.in_sourcesI) finally show ?thesis by simp qed qed lemma Con_single_ide_iff: assumes "R.ide a" shows "[a] \<^sup>*\\<^sup>* T \ Arr T \ a \ Srcs T" using assms Con_single_ide_ind by simp lemma Con_single_ideI [intro]: assumes "R.ide a" and "Arr T" and "a \ Srcs T" shows "[a] \<^sup>*\\<^sup>* T" and "T \<^sup>*\\<^sup>* [a]" using assms Con_single_ide_iff Con_sym by auto lemma Resid_single_ide: assumes "R.ide a" and "[a] \<^sup>*\\<^sup>* T" shows "[a] \<^sup>*\\\<^sup>* T \ (\b. [b]) ` Trgs T" and [simp]: "T \<^sup>*\\\<^sup>* [a] = T" using assms Con_single_ide_ind img_Resid_Srcs Resid_Arr_Src Con_sym by blast+ lemma Resid_Arr_Ide_ind: shows "\Ide A; T \<^sup>*\\<^sup>* A\ \ T \<^sup>*\\\<^sup>* A = T" proof (induct A) show "\Ide []; T \<^sup>*\\<^sup>* []\ \ T \<^sup>*\\\<^sup>* [] = T" by simp fix a A assume ind: "\Ide A; T \<^sup>*\\<^sup>* A\ \ T \<^sup>*\\\<^sup>* A = T" assume Ide: "Ide (a # A)" assume Con: "T \<^sup>*\\<^sup>* a # A" show "T \<^sup>*\\\<^sup>* (a # A) = T" by (metis (no_types, lifting) Con Con_initial_left Con_sym Ide Ide.elims(2) Resid_cons(2) Resid_single_ide(2) ind list.inject) qed lemma Resid_Ide_Arr_ind: shows "\Ide A; A \<^sup>*\\<^sup>* T\ \ Ide (A \<^sup>*\\\<^sup>* T)" proof (induct A) show "\Ide []; [] \<^sup>*\\<^sup>* T\ \ Ide ([] \<^sup>*\\\<^sup>* T)" by simp fix a A assume ind: "\Ide A; A \<^sup>*\\<^sup>* T\ \ Ide (A \<^sup>*\\\<^sup>* T)" assume Ide: "Ide (a # A)" assume Con: "a # A \<^sup>*\\<^sup>* T" have T: "Arr T" using Con Ide Con_single_ide_ind Con_initial_left Ide.elims(2) by blast show "Ide ((a # A) \<^sup>*\\\<^sup>* T)" proof (cases "A = []") show "A = [] \ ?thesis" by (metis Con Con_sym1 Ide Ide.simps(2) Resid1x_as_Resid Resid1x_ide Residx1_as_Resid Con_sym) assume A: "A \ []" show ?thesis proof - have "Ide ([a] \<^sup>*\\\<^sup>* T)" by (metis Con Con_initial_left Con_sym Con_sym1 Ide Ide.simps(3) Resid1x_as_Resid Residx1_as_Resid Ide.simps(2) Resid1x_ide list.exhaust_sel) moreover have "Trgs ([a] \<^sup>*\\\<^sup>* T) \ Srcs (A \<^sup>*\\\<^sup>* T)" using A T Ide Con by (metis (no_types, lifting) Con_sym Ide.elims(2) Ide.simps(2) Resid_Arr_Ide_ind Srcs_Resid Trgs_Resid_sym Con_cons(2) dual_order.eq_iff list.inject) moreover have "Ide (A \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [a]))" by (metis A Con Con_cons(1) Con_sym Ide Ide.simps(3) Resid_Arr_Ide_ind Resid_single_ide(2) ind list.exhaust_sel) moreover have "Ide ((a # A) \<^sup>*\\\<^sup>* T) \ Ide ([a] \<^sup>*\\\<^sup>* T) \ Ide (A \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [a])) \ Trgs ([a] \<^sup>*\\\<^sup>* T) \ Srcs (A \<^sup>*\\\<^sup>* T)" using calculation(1-3) by (metis Arr.simps(1) Con Ide Ide.simps(3) Resid1x_as_Resid Resid_cons' Trgs.simps(2) Con_single_ide_iff Ide.simps(2) Ide_implies_Arr Resid_Arr_Src list.exhaust_sel) ultimately show ?thesis by blast qed qed qed lemma Resid_Ide: assumes "Ide A" and "A \<^sup>*\\<^sup>* T" shows "T \<^sup>*\\\<^sup>* A = T" and "Ide (A \<^sup>*\\\<^sup>* T)" using assms Resid_Ide_Arr_ind Resid_Arr_Ide_ind Con_sym by auto lemma Con_Ide_iff: shows "Ide A \ A \<^sup>*\\<^sup>* T \ Arr T \ Srcs T = Srcs A" proof (induct A) show "Ide [] \ [] \<^sup>*\\<^sup>* T \ Arr T \ Srcs T = Srcs []" by simp fix a A assume ind: "Ide A \ A \<^sup>*\\<^sup>* T \ Arr T \ Srcs T = Srcs A" assume Ide: "Ide (a # A)" show "a # A \<^sup>*\\<^sup>* T \ Arr T \ Srcs T = Srcs (a # A)" proof (cases "A = []") show "A = [] \ ?thesis" using Con_single_ide_ind Ide by (metis Arr.simps(2) Con_sym Ide.simps(2) Ide_implies_Arr R.arrE Resid_Arr_Src Srcs.simps(2) Srcs_Resid R.in_sourcesI) assume A: "A \ []" have "a # A \<^sup>*\\<^sup>* T \ [a] \<^sup>*\\<^sup>* T \ A \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* [a]" using A Ide Con_cons(1) [of A T a] by fastforce also have 1: "... \ Arr T \ a \ Srcs T" by (metis A Arr_has_Src Con_single_ide_ind Ide Ide.elims(2) Resid_Arr_Src Srcs_Resid_Arr_single Con_sym Srcs_eqI ind inf.absorb_iff2 list.inject) also have "... \ Arr T \ Srcs T = Srcs (a # A)" by (metis A 1 Con_sym Ide Ide.simps(3) R.ideE R.sources_resid Resid_Arr_Src Srcs.simps(3) Srcs_Resid_Arr_single list.exhaust_sel R.in_sourcesI) finally show "a # A \<^sup>*\\<^sup>* T \ Arr T \ Srcs T = Srcs (a # A)" by blast qed qed lemma Con_IdeI: assumes "Ide A" and "Arr T" and "Srcs T = Srcs A" shows "A \<^sup>*\\<^sup>* T" and "T \<^sup>*\\<^sup>* A" using assms Con_Ide_iff Con_sym by auto lemma Con_Arr_self: shows "Arr T \ T \<^sup>*\\<^sup>* T" proof (induct T) show "Arr [] \ [] \<^sup>*\\<^sup>* []" by simp fix t T assume ind: "Arr T \ T \<^sup>*\\<^sup>* T" assume Arr: "Arr (t # T)" show "t # T \<^sup>*\\<^sup>* t # T" proof (cases "T = []") show "T = [] \ ?thesis" using Arr R.arrE by simp assume T: "T \ []" have "t \ t \ T \<^sup>*\\<^sup>* [t \\ t] \ [t \\ t] \<^sup>*\\<^sup>* T \ T \<^sup>*\\\<^sup>* [t \\ t] \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* [t \\ t]" proof - have "t \ t" using Arr Arr.elims(1) by auto moreover have "T \<^sup>*\\<^sup>* [t \\ t]" proof - have "Ide [t \\ t]" by (simp add: R.arr_def R.prfx_reflexive calculation) moreover have "Srcs [t \\ t] = Srcs T" by (metis Arr Arr.simps(2) Arr_has_Trg R.arrE R.sources_resid Srcs.simps(2) Srcs_eqI T Trgs.simps(2) Arr.simps(3) inf.absorb_iff2 list.exhaust) ultimately show ?thesis by (metis Arr Con_sym T Arr.simps(3) Con_Ide_iff neq_Nil_conv) qed ultimately show ?thesis by (metis Con_single_ide_ind Con_sym R.prfx_reflexive Resid_single_ide(2) ind R.con_implies_arr(1)) qed thus ?thesis using Con_rec(4) [of T T t t] by force qed qed lemma Resid_Arr_self: shows "Arr T \ Ide (T \<^sup>*\\\<^sup>* T)" proof (induct T) show "Arr [] \ Ide ([] \<^sup>*\\\<^sup>* [])" by simp fix t T assume ind: "Arr T \ Ide (T \<^sup>*\\\<^sup>* T)" assume Arr: "Arr (t # T)" show "Ide ((t # T) \<^sup>*\\\<^sup>* (t # T))" proof (cases "T = []") show "T = [] \ ?thesis" using Arr R.prfx_reflexive by auto assume T: "T \ []" have 1: "(t # T) \<^sup>*\\\<^sup>* (t # T) = t \<^sup>1\\\<^sup>* (t # T) # T \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* [t])" using Arr T Resid_cons' [of T t "t # T"] Con_Arr_self by presburger also have "... = (t \\ t) \<^sup>1\\\<^sup>* T # T \<^sup>*\\\<^sup>* (t \<^sup>1\\\<^sup>* [t] # T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t]))" using Arr T Resid_cons' [of T t "[t]"] by (metis Con_initial_right Resid1x.simps(3) calculation neq_Nil_conv) also have "... = (t \\ t) \<^sup>1\\\<^sup>* T # (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t]))" by (metis 1 Resid1x.simps(2) Residx1.simps(2) Residx1_as_Resid T calculation Con_cons(1) Con_rec(4) Resid_cons(2) list.distinct(1) list.inject) finally have 2: "(t # T) \<^sup>*\\\<^sup>* (t # T) = (t \\ t) \<^sup>1\\\<^sup>* T # (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t]))" by blast moreover have "Ide ..." proof - have "R.ide ((t \\ t) \<^sup>1\\\<^sup>* T)" using Arr T by (metis Con_initial_right Con_rec(2) Con_sym1 R.con_implies_arr(1) Resid1x_ide Con_Arr_self Residx1_as_Resid R.prfx_reflexive) moreover have "Ide ((T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])))" using Arr T by (metis Con_Arr_self Con_rec(4) Resid_single_ide(2) Con_single_ide_ind Resid.simps(3) ind R.prfx_reflexive R.con_implies_arr(2)) moreover have "R.targets ((t \\ t) \<^sup>1\\\<^sup>* T) \ Srcs ((T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [t])))" by (metis (no_types, lifting) 1 2 Con_cons(1) Resid1x_as_Resid T Trgs.simps(2) Trgs_Resid_sym Srcs_Resid dual_order.eq_iff list.discI list.inject) ultimately show ?thesis using Arr T by (metis Ide.simps(1,3) list.exhaust_sel) qed ultimately show ?thesis by auto qed qed lemma Con_imp_eq_Srcs: assumes "T \<^sup>*\\<^sup>* U" shows "Srcs T = Srcs U" proof (cases T) show "T = [] \ ?thesis" using assms by simp fix t T' assume T: "T = t # T'" show "Srcs T = Srcs U" proof (cases U) show "U = [] \ ?thesis" using assms T by simp fix u U' assume U: "U = u # U'" show "Srcs T = Srcs U" by (metis Con_initial_right Con_rec(1) Con_sym R.con_imp_common_source Srcs.simps(2-3) Srcs_eqI T Trgs.cases U assms) qed qed lemma Arr_iff_Con_self: shows "Arr T \ T \<^sup>*\\<^sup>* T" proof (induct T) show "Arr [] \ [] \<^sup>*\\<^sup>* []" by simp fix t T assume ind: "Arr T \ T \<^sup>*\\<^sup>* T" show "Arr (t # T) \ t # T \<^sup>*\\<^sup>* t # T" proof (cases "T = []") show "T = [] \ ?thesis" by auto assume T: "T \ []" show ?thesis proof show "Arr (t # T) \ t # T \<^sup>*\\<^sup>* t # T" using Con_Arr_self by simp show "t # T \<^sup>*\\<^sup>* t # T \ Arr (t # T)" proof - assume Con: "t # T \<^sup>*\\<^sup>* t # T" have "R.arr t" using T Con Con_rec(4) [of T T t t] by blast moreover have "Arr T" using T Con Con_rec(4) [of T T t t] ind R.arrI by (meson R.prfx_reflexive Con_single_ide_ind) moreover have "R.targets t \ Srcs T" using T Con by (metis Con_cons(2) Con_imp_eq_Srcs Trgs.simps(2) Srcs_Resid list.distinct(1) subsetI) ultimately show ?thesis by (cases T) auto qed qed qed qed lemma Arr_Resid_single: shows "T \<^sup>*\\<^sup>* [u] \ Arr (T \<^sup>*\\\<^sup>* [u])" proof (induct T arbitrary: u) show "\u. [] \<^sup>*\\<^sup>* [u] \ Arr ([] \<^sup>*\\\<^sup>* [u])" by simp fix t u T assume ind: "\u. T \<^sup>*\\<^sup>* [u] \ Arr (T \<^sup>*\\\<^sup>* [u])" assume Con: "t # T \<^sup>*\\<^sup>* [u]" show "Arr ((t # T) \<^sup>*\\\<^sup>* [u])" proof (cases "T = []") show "T = [] \ ?thesis" using Con Arr_iff_Con_self R.con_imp_arr_resid Con_rec(1) by fastforce assume T: "T \ []" have "Arr ((t # T) \<^sup>*\\\<^sup>* [u]) \ Arr ((t \\ u) # (T \<^sup>*\\\<^sup>* [u \\ t]))" using Con T Resid_rec(2) by auto also have "... \ R.arr (t \\ u) \ Arr (T \<^sup>*\\\<^sup>* [u \\ t]) \ R.targets (t \\ u) \ Srcs (T \<^sup>*\\\<^sup>* [u \\ t])" using Con T by (metis Arr.simps(3) Con_rec(2) neq_Nil_conv) also have "... \ R.con t u \ Arr (T \<^sup>*\\\<^sup>* [u \\ t])" using Con T by (metis Srcs_Resid_Arr_single Con_rec(2) R.arr_resid_iff_con subsetI R.targets_resid_sym) also have "... \ True" using Con ind T Con_rec(2) by blast finally show ?thesis by auto qed qed lemma Con_imp_Arr_Resid: shows "T \<^sup>*\\<^sup>* U \ Arr (T \<^sup>*\\\<^sup>* U)" proof (induct U arbitrary: T) show "\T. T \<^sup>*\\<^sup>* [] \ Arr (T \<^sup>*\\\<^sup>* [])" by (meson Con_sym Resid.simps(1)) fix u U T assume ind: "\T. T \<^sup>*\\<^sup>* U \ Arr (T \<^sup>*\\\<^sup>* U)" assume Con: "T \<^sup>*\\<^sup>* u # U" show "Arr (T \<^sup>*\\\<^sup>* (u # U))" by (metis Arr_Resid_single Con Resid_cons(2) ind) qed lemma Cube_ind: shows "\T \<^sup>*\\<^sup>* U; V \<^sup>*\\<^sup>* T; length T + length U + length V \ n\ \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U) \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U))" proof (induct n arbitrary: T U V) show "\T U V. \T \<^sup>*\\<^sup>* U; V \<^sup>*\\<^sup>* T; length T + length U + length V \ 0\ \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U) \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U))" by simp fix n and T U V :: "'a list" assume Con_TU: "T \<^sup>*\\<^sup>* U" and Con_VT: "V \<^sup>*\\<^sup>* T" have T: "T \ []" using Con_TU by auto have U: "U \ []" using Con_TU Con_sym Resid.simps(1) by blast have V: "V \ []" using Con_VT by auto assume len: "length T + length U + length V \ Suc n" assume ind: "\T U V. \T \<^sup>*\\<^sup>* U; V \<^sup>*\\<^sup>* T; length T + length U + length V \ n\ \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U) \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U))" show "(V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U) \ (V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U))" proof (cases V) show "V = [] \ ?thesis" using V by simp (* * TODO: I haven't found a better way to do this than just consider each combination * of T U V being a singleton. *) fix v V' assume V: "V = v # V'" show ?thesis proof (cases U) show "U = [] \ ?thesis" using U by simp fix u U' assume U: "U = u # U'" show ?thesis proof (cases T) show "T = [] \ ?thesis" using T by simp fix t T' assume T: "T = t # T'" show ?thesis proof (cases "V' = []", cases "U' = []", cases "T' = []") show "\V' = []; U' = []; T' = []\ \ ?thesis" using T U V R.cube Con_TU Resid.simps(2) Resid.simps(3) R.arr_resid_iff_con R.con_implies_arr Con_sym by metis assume T': "T' \ []" and V': "V' = []" and U': "U' = []" have 1: "U \<^sup>*\\<^sup>* [t]" using T Con_TU Con_cons(2) Con_sym Resid.simps(2) by metis have 2: "V \<^sup>*\\<^sup>* [t]" using V Con_VT Con_initial_right T by blast show ?thesis proof (intro conjI impI) have 3: "length [t] + length U + length V \ n" using T T' le_Suc_eq len by fastforce show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T'" using Con_TU Con_VT Con_sym Resid_cons(2) T T' by force also have "... \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t] \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" proof (intro iffI conjI) show "(V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]" using T U V T' U' V' 1 ind len Con_TU Con_rec(2) Resid_rec(1) Resid.simps(1) length_Cons Suc_le_mono add_Suc by (metis (no_types)) show "(V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" using T U V T' U' V' by (metis Con_sym Resid.simps(1) Resid_rec(1) Suc_le_mono ind len length_Cons list.size(3-4)) show "V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t] \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T'" using T U V T' U' V' 1 ind len Con_TU Con_VT Con_rec(1-3) by (metis (no_types, lifting) One_nat_def Resid_rec(1) Suc_le_mono add.commute list.size(3) list.size(4) plus_1_eq_Suc) qed also have "... \ (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" by (metis 2 3 Con_sym ind Resid.simps(1)) also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" using Con_rec(2) [of T' t] by (metis (no_types, lifting) "1" Con_TU Con_cons(2) Resid.simps(1) Resid.simps(3) Resid_rec(2) T T' U U') finally show ?thesis by simp qed assume Con: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T" show "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* ((U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T')" using Con_TU Con_VT Con_sym Resid_cons(2) T T' by force also have "... = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" using T U V T' U' V' 1 Con ind [of T' "Resid U [t]" "Resid V [t]"] by (metis One_nat_def add.commute calculation len length_0_conv length_Resid list.size(4) nat_add_left_cancel_le Con_sym plus_1_eq_Suc) also have "... = ((V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U)) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by (metis "1" "2" "3" Con_sym ind) also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" using T U T' U' Con * by (metis Con_sym Resid_rec(1-2) Resid.simps(1) Resid_cons(2)) finally show ?thesis by simp qed qed next assume U': "U' \ []" and V': "V' = []" show ?thesis proof (intro conjI impI) show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof (cases "T' = []") assume T': "T' = []" show ?thesis proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* (u \\ t) # (U' \<^sup>*\\\<^sup>* [t \\ u])" using Con_TU Con_sym Resid_rec(2) T T' U U' by auto also have "... \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" by (metis Con_TU Con_cons(2) Con_rec(3) Con_sym Resid.simps(1) T U U') also have "... \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" using T U V V' R.cube_ax apply simp by (metis R.con_implies_arr(1) R.not_arr_null R.con_def) also have "... \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U' \<^sup>*\\<^sup>* [t \\ u] \<^sup>*\\\<^sup>* U'" proof - have "length [t \\ u] + length U' + length (V \<^sup>*\\\<^sup>* [u]) \ n" using T U V V' len by force thus ?thesis by (metis Con_sym Resid.simps(1) add.commute ind) qed also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" by (metis Con_TU Resid_cons(2) Resid_rec(3) T T' U U' Con_cons(2) length_Resid length_0_conv) finally show ?thesis by simp qed next assume T': "T' \ []" show ?thesis proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* ((U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T')" using Con_TU Con_VT Con_sym Resid_cons(2) T T' by force also have "... \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" proof - have "length T' + length (U \<^sup>*\\\<^sup>* [t]) + length (V \<^sup>*\\\<^sup>* [t]) \ n" by (metis (no_types, lifting) Con_TU Con_VT Con_initial_right Con_sym One_nat_def Suc_eq_plus1 T ab_semigroup_add_class.add_ac(1) add_le_imp_le_left len length_Resid list.size(4) plus_1_eq_Suc) thus ?thesis by (metis Con_TU Con_VT Con_cons(1) Con_cons(2) T T' U V ind list.discI) qed also have "... \ (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" proof - have "length [t] + length U + length V \ n" using T T' le_Suc_eq len by fastforce thus ?thesis by (metis Con_TU Con_VT Con_initial_left Con_initial_right T ind) qed also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" by (metis Con_cons(2) Con_sym Resid.simps(1) Resid1x_as_Resid Residx1_as_Resid Resid_cons' T T') finally show ?thesis by blast qed qed show "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - assume Con: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T" show ?thesis proof (cases "T' = []") assume T': "T' = []" show ?thesis proof - have 1: "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* ((u \\ t) # (U'\<^sup>*\\\<^sup>* [t \\ u]))" using Con_TU Con_sym Resid_rec(2) T T' U U' by force also have "... = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" by (metis Con Con_TU Con_rec(2) Con_sym Resid_cons(2) T T' U U' calculation) also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" by (metis "*" Con Con_rec(3) R.cube Resid.simps(1,3) T T' U V V' calculation R.conI R.conE) also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t \\ u] \<^sup>*\\\<^sup>* U')" proof - have "length [t \\ u] + length (U' \<^sup>*\\\<^sup>* [t \\ u]) + length (V \<^sup>*\\\<^sup>* [u]) \ n" by (metis (no_types, lifting) Nat.le_diff_conv2 One_nat_def T U V V' add.commute add_diff_cancel_left' add_leD2 len length_Cons length_Resid list.size(3) plus_1_eq_Suc) thus ?thesis by (metis Con_sym add.commute Resid.simps(1) ind length_Resid) qed also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" by (metis Con_TU Con_cons(2) Resid_cons(2) T T' U U' Resid_rec(3) length_0_conv length_Resid) finally show ?thesis by blast qed next assume T': "T' \ []" show ?thesis proof - have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = ((V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* ([u] \<^sup>*\\\<^sup>* T)) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [u]))" by (metis Con Con_TU Resid.simps(2) Resid1x_as_Resid U U' Con_cons(2) Con_sym Resid_cons' Resid_cons(2)) also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [u])) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* [u]))" proof - have "length T + length [u] + length V \ n" using U U' antisym_conv len not_less_eq_eq by fastforce thus ?thesis by (metis Con_TU Con_VT Con_initial_right U ind) qed also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ((T \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U')" proof - have "length (T \<^sup>*\\\<^sup>* [u]) + length U' + length (V \<^sup>*\\\<^sup>* [u]) \ n" using Con_TU Con_initial_right U V V' len length_Resid by force thus ?thesis by (metis Con Con_TU Con_cons(2) U U' calculation ind length_0_conv length_Resid) qed also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" by (metis "*" Con Con_TU Resid_cons(2) U U' length_Resid length_0_conv) finally show ?thesis by blast qed qed qed qed next assume V': "V' \ []" show ?thesis proof (cases "U' = []") assume U': "U' = []" show ?thesis proof (cases "T' = []") assume T': "T' = []" show ?thesis proof (intro conjI impI) show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (v \\ t) # (V' \<^sup>*\\\<^sup>* [t \\ v]) \<^sup>*\\<^sup>* [u \\ t]" using Con_TU Con_VT Con_sym Resid_rec(1-2) T T' U U' V V' by metis also have "... \ [v \\ t] \<^sup>*\\<^sup>* [u \\ t] \ V' \<^sup>*\\\<^sup>* [t \\ v] \<^sup>*\\<^sup>* [u \\ v] \<^sup>*\\\<^sup>* [t \\ v]" by (metis T T' V V' Con_VT Con_rec(1-2) Con_sym R.con_def R.cube Resid.simps(3)) also have "... \ [v \\ t] \<^sup>*\\<^sup>* [u \\ t] \ V' \<^sup>*\\\<^sup>* [u \\ v] \<^sup>*\\<^sup>* [t \\ v] \<^sup>*\\\<^sup>* [u \\ v]" proof - have "length [t \\ v] + length [u \\ v] + length V' \ n" using T U V len by fastforce thus ?thesis by (metis Con_imp_Arr_Resid Arr_has_Src Con_VT T T' Trgs.simps(1) Trgs_Resid_sym V V' Con_rec(2) Srcs_Resid ind) qed also have "... \ [v \\ t] \<^sup>*\\<^sup>* [u \\ t] \ V' \<^sup>*\\\<^sup>* [u \\ v] \<^sup>*\\<^sup>* [t \\ u] \<^sup>*\\\<^sup>* [v \\ u]" by (simp add: R.con_def R.cube) also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof assume 1: "V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" have tu_vu: "t \\ u \ v \\ u" by (metis (no_types, lifting) 1 T T' U U' V V' Con_rec(3) Resid_rec(1-2) Con_sym length_Resid length_0_conv) have vt_ut: "v \\ t \ u \\ t" using 1 by (metis R.con_def R.con_sym R.cube tu_vu) show "[v \\ t] \<^sup>*\\<^sup>* [u \\ t] \ V' \<^sup>*\\\<^sup>* [u \\ v] \<^sup>*\\<^sup>* [t \\ u] \<^sup>*\\\<^sup>* [v \\ u]" by (metis (no_types, lifting) "1" Con_TU Con_cons(1) Con_rec(1-2) Resid_rec(1) T T' U U' V V' Resid_rec(2) length_Resid length_0_conv vt_ut) next assume 1: "[v \\ t] \<^sup>*\\<^sup>* [u \\ t] \ V' \<^sup>*\\\<^sup>* [u \\ v] \<^sup>*\\<^sup>* [t \\ u] \<^sup>*\\\<^sup>* [v \\ u]" have tu_vu: "t \\ u \ v \\ u \ v \\ t \ u \\ t" by (metis 1 Con_sym Resid.simps(1) Residx1.simps(2) Residx1_as_Resid) have tu: "t \ u" using Con_TU Con_rec(1) T T' U U' by blast show "V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" by (metis (no_types, opaque_lifting) 1 Con_rec(2) Con_sym R.con_implies_arr(2) Resid.simps(1,3) T T' U U' V V' Resid_rec(2) R.arr_resid_iff_con) qed finally show ?thesis by simp qed show "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - assume Con: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T" have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = ((v \\ t) # (V' \<^sup>*\\\<^sup>* [t \\ v])) \<^sup>*\\\<^sup>* [u \\ t]" using Con_TU Con_VT Con_sym Resid_rec(1-2) T T' U U' V V' by metis also have 1: "... = ((v \\ t) \\ (u \\ t)) # (V' \<^sup>*\\\<^sup>* [t \\ v]) \<^sup>*\\\<^sup>* ([u \\ v] \<^sup>*\\\<^sup>* [t \\ v])" apply simp by (metis Con Con_VT Con_rec(2) R.conE R.conI R.con_sym R.cube Resid_rec(2) T T' V V' calculation(1)) also have "... = ((v \\ t) \\ (u \\ t)) # (V' \<^sup>*\\\<^sup>* [u \\ v]) \<^sup>*\\\<^sup>* ([t \\ v] \<^sup>*\\\<^sup>* [u \\ v])" proof - have "length [t \\ v] + length [u \\ v] + length V' \ n" using T U V len by fastforce moreover have "u \\ v \ t \\ v" by (metis 1 Con_VT Con_rec(2) R.con_sym_ax T T' V V' list.discI R.conE R.conI R.cube) moreover have "t \\ v \ u \\ v" using R.con_sym calculation(2) by blast ultimately show ?thesis by (metis Con_VT Con_rec(2) T T' V V' Con_rec(1) ind) qed also have "... = ((v \\ t) \\ (u \\ t)) # ((V' \<^sup>*\\\<^sup>* [u \\ v]) \<^sup>*\\\<^sup>* ([t \\ u] \<^sup>*\\\<^sup>* [v \\ u]))" using R.cube by fastforce also have "... = ((v \\ u) \\ (t \\ u)) # ((V' \<^sup>*\\\<^sup>* [u \\ v]) \<^sup>*\\\<^sup>* ([t \\ u] \<^sup>*\\\<^sup>* [v \\ u]))" by (metis R.cube) also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - have "(V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U) = ((v \\ u) # ((V' \<^sup>*\\\<^sup>* [u \\ v]))) \<^sup>*\\\<^sup>* [t \\ u]" using T T' U U' V Resid_cons(1) [of "[u]" v V'] by (metis "*" Con Con_TU Resid.simps(1) Resid_rec(1) Resid_rec(2)) also have "... = ((v \\ u) \\ (t \\ u)) # ((V' \<^sup>*\\\<^sup>* [u \\ v]) \<^sup>*\\\<^sup>* ([t \\ u] \<^sup>*\\\<^sup>* [v \\ u]))" by (metis "*" Con Con_initial_left calculation Con_sym Resid.simps(1) Resid_rec(1-2)) finally show ?thesis by simp qed finally show ?thesis by simp qed qed next assume T': "T' \ []" show ?thesis proof (intro conjI impI) show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* [u \\ t] \<^sup>*\\\<^sup>* T'" using Con_TU Con_VT Con_sym Resid_cons(2) Resid_rec(3) T T' U U' by force also have "... \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* [u \\ t]" proof - have "length [u \\ t] + length T' + length (V \<^sup>*\\\<^sup>* [t]) \ n" using Con_VT Con_initial_right T U length_Resid len by fastforce thus ?thesis by (metis Con_TU Con_VT Con_rec(2) T T' U V add.commute Con_cons(2) ind list.discI) qed also have "... \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u] \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* [u \\ t]" proof - have "length [t] + length [u] + length V \ n" using T T' U le_Suc_eq len by fastforce hence "(V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* ([u] \<^sup>*\\\<^sup>* [t]) = (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* [u])" using ind [of "[t]" "[u]" V] by (metis Con_TU Con_VT Con_initial_left Con_initial_right T U) thus ?thesis by (metis (full_types) Con_TU Con_initial_left Con_sym Resid_rec(1) T U) qed also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" by (metis Con_TU Con_cons(2) Con_rec(2) Resid.simps(1) Resid_rec(2) T T' U U') finally show ?thesis by simp qed show "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - assume Con: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T" have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* ([u \\ t] \<^sup>*\\\<^sup>* T')" using Con_TU Con_VT Con_sym Resid_cons(2) Resid_rec(3) T T' U U' by force also have "... = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* [u \\ t])" proof - have "length [u \\ t] + length T' + length (Resid V [t]) \ n" using Con_VT Con_initial_right T U length_Resid len by fastforce thus ?thesis by (metis Con_TU Con_VT Con_cons(2) Con_rec(2) T T' U V add.commute ind list.discI) qed also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u]) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* [u \\ t])" proof - have "length [t] + length [u] + length V \ n" using T T' U le_Suc_eq len by fastforce thus ?thesis using ind [of "[t]" "[u]" V] by (metis Con_TU Con_VT Con_initial_left Con_sym Resid_rec(1) T U) qed also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" using * Con Con_TU Con_rec(2) Resid_cons(2) Resid_rec(2) T T' U U' by auto finally show ?thesis by simp qed qed qed next assume U': "U' \ []" show ?thesis proof (cases "T' = []") assume T': "T' = []" show ?thesis proof (intro conjI impI) show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* (u \\ t) # (U' \<^sup>*\\\<^sup>* [t \\ u])" using T U V T' U' V' Con_TU Con_VT Con_sym Resid_rec(2) by auto also have "... \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* [u \\ t] \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" by (metis Con_TU Con_VT Con_cons(2) Con_initial_right Con_rec(2) Con_sym T U U') also have "... \ V \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* [u \\ t] \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" proof - have "length [u] + length [t] + length V \ n" using T U V T' U' V' len not_less_eq_eq order_trans by fastforce thus ?thesis using ind [of "[t]" "[u]" V] by (metis Con_TU Con_VT Con_initial_right Resid_rec(1) T U Con_sym length_Cons) qed also have "... \ V \<^sup>*\\\<^sup>* [u] \<^sup>*\\<^sup>* [t \\ u] \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u] \<^sup>*\\<^sup>* U' \<^sup>*\\\<^sup>* [t \\ u]" proof - have "length [t] + length [u] + length V \ n" using T U V T' U' V' len antisym_conv not_less_eq_eq by fastforce thus ?thesis by (metis (full_types) Con_TU Con_VT Con_initial_right Con_sym Resid_rec(1) T U ind) qed also have "... \ (V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U' \<^sup>*\\<^sup>* [t \\ u] \<^sup>*\\\<^sup>* U'" proof - have "length [t \\ u] + length U' + length (V \<^sup>*\\\<^sup>* [u]) \ n" by (metis T T' U add.assoc add.right_neutral add_leD1 add_le_cancel_left length_Resid len length_Cons list.size(3) plus_1_eq_Suc) thus ?thesis by (metis (no_types, opaque_lifting) Con_sym Resid.simps(1) add.commute ind) qed also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" by (metis Con_TU Resid_cons(2) Resid_rec(3) T T' U U' Con_cons(2) length_Resid length_0_conv) finally show ?thesis by blast qed show "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - assume Con: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T" have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* ((u \\ t) # (U' \<^sup>*\\\<^sup>* [t \\ u]))" using Con_TU Con_sym Resid_rec(2) T T' U U' by auto also have "... = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u \\ t]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" by (metis Con Con_TU Con_rec(2) Con_sym T T' U U' calculation Resid_cons(2)) also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* [t \\ u]) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t \\ u])" proof - have "length [t] + length [u] + length V \ n" using T U U' le_Suc_eq len by fastforce thus ?thesis using T U Con_TU Con_VT Con_sym ind [of "[t]" "[u]" V] by (metis (no_types, opaque_lifting) Con_initial_right Resid.simps(3)) qed also have "... = ((V \<^sup>*\\\<^sup>* [u]) \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t \\ u] \<^sup>*\\\<^sup>* U')" proof - have "length [t \\ u] + length U' + length (V \<^sup>*\\\<^sup>* [u]) \ n" by (metis (no_types, opaque_lifting) T T' U add.left_commute add.right_neutral add_leD2 add_le_cancel_left len length_Cons length_Resid list.size(3) plus_1_eq_Suc) thus ?thesis by (metis Con Con_TU Con_rec(3) T T' U U' calculation ind length_0_conv length_Resid) qed also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" by (metis "*" Con Con_TU Resid_rec(3) T T' U U' Resid_cons(2) length_Resid length_0_conv) finally show ?thesis by blast qed qed next assume T': "T' \ []" show ?thesis proof (intro conjI impI) have 1: "U \<^sup>*\\<^sup>* [t]" using T Con_TU by (metis Con_cons(2) Con_sym Resid.simps(2)) have 2: "V \<^sup>*\\<^sup>* [t]" using V Con_VT Con_initial_right T by blast have 3: "length T' + length (U \<^sup>*\\\<^sup>* [t]) + length (V \<^sup>*\\\<^sup>* [t]) \ n" using "1" "2" T len length_Resid by force have 4: "length [t] + length U + length V \ n" using T T' len antisym_conv not_less_eq_eq by fastforce show *: "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" proof - have "V \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* T \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T' \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T'" using Con_TU Con_VT Con_sym Resid_cons(2) T T' by force also have "... \ (V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" by (metis 3 Con_TU Con_VT Con_cons(1) Con_cons(2) T T' U V ind list.discI) also have "... \ (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" by (metis 1 2 4 Con_sym ind) also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* hd ([t] \<^sup>*\\\<^sup>* U) # T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" by (metis 1 Con_TU Con_cons(1) Con_cons(2) Resid.simps(1) Resid1x_as_Resid T T' list.sel(1)) also have "... \ V \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T \<^sup>*\\\<^sup>* U" using 1 Resid_cons' [of T' t U] Con_TU T T' Resid1x_as_Resid Con_sym by force finally show ?thesis by simp qed show "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" proof - have "(V \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T) = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* ((U \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* T')" using Con_TU Con_VT Con_sym Resid_cons(2) T T' by force also have "... = ((V \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by (metis (no_types, lifting) "3" Con_TU Con_VT T T' U V Con_cons(1) Con_cons(2) ind list.simps(3)) also have "... = ((V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U)) \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by (metis 1 2 4 Con_sym ind) also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ((t # T') \<^sup>*\\\<^sup>* U)" by (metis "*" Con_TU Con_cons(1) Resid1x_as_Resid Resid_cons' T T' U calculation Resid_cons(2) list.distinct(1)) also have "... = (V \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* U)" using T by fastforce finally show ?thesis by simp qed qed qed qed qed qed qed qed qed lemma Cube: shows "T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* U \ T \<^sup>*\\\<^sup>* V \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* V" and "T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* U \ (T \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (V \<^sup>*\\\<^sup>* U) = (T \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* V)" proof - show "T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* U \ T \<^sup>*\\\<^sup>* V \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* V" using Cube_ind by (metis Con_sym Resid.simps(1) le_add2) show "T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* U \ (T \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* (V \<^sup>*\\\<^sup>* U) = (T \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* V)" using Cube_ind by (metis Con_sym Resid.simps(1) order_refl) qed lemma Con_implies_Arr: assumes "T \<^sup>*\\<^sup>* U" shows "Arr T" and "Arr U" using assms Con_sym by (metis Con_imp_Arr_Resid Arr_iff_Con_self Cube(1) Resid.simps(1))+ sublocale partial_magma Resid by (unfold_locales, metis Resid.simps(1) Con_sym) lemma is_partial_magma: shows "partial_magma Resid" .. lemma null_char: shows "null = []" by (metis null_is_zero(2) Resid.simps(1)) sublocale residuation Resid using null_char Con_sym Arr_iff_Con_self Con_imp_Arr_Resid Cube null_is_zero(2) by unfold_locales auto lemma is_residuation: shows "residuation Resid" .. lemma arr_char: shows "arr T \ Arr T" using null_char Arr_iff_Con_self by fastforce lemma arrI\<^sub>P [intro]: assumes "Arr T" shows "arr T" using assms arr_char by auto lemma ide_char: shows "ide T \ Ide T" by (metis Con_Arr_self Ide_implies_Arr Resid_Arr_Ide_ind Resid_Arr_self arr_char ide_def arr_def) lemma con_char: shows "con T U \ Con T U" using null_char by auto lemma conI\<^sub>P [intro]: assumes "Con T U" shows "con T U" using assms con_char by auto sublocale rts Resid proof show "\A T. \ide A; con T A\ \ T \<^sup>*\\\<^sup>* A = T" using Resid_Arr_Ide_ind ide_char null_char by auto show "\T. arr T \ ide (trg T)" by (metis arr_char Resid_Arr_self ide_char resid_arr_self) show "\A T. \ide A; con A T\ \ ide (A \<^sup>*\\\<^sup>* T)" by (simp add: Resid_Ide_Arr_ind con_char ide_char) show "\T U. con T U \ \A. ide A \ con A T \ con A U" proof - fix T U assume TU: "con T U" have 1: "Srcs T = Srcs U" using TU Con_imp_eq_Srcs con_char by force obtain a where a: "a \ Srcs T \ Srcs U" using 1 by (metis Int_absorb Int_emptyI TU arr_char Arr_has_Src con_implies_arr(1)) show "\A. ide A \ con A T \ con A U" using a 1 by (metis (full_types) Ball_Collect Con_single_ide_ind Ide.simps(2) Int_absorb TU Srcs_are_ide arr_char con_char con_implies_arr(1-2) ide_char) qed show "\T U V. \ide (Resid T U); con U V\ \ con (T \<^sup>*\\\<^sup>* U) (V \<^sup>*\\\<^sup>* U)" using null_char ide_char by (metis Con_imp_Arr_Resid Con_Ide_iff Srcs_Resid con_char con_sym arr_resid_iff_con ide_implies_arr) qed theorem is_rts: shows "rts Resid" .. notation cong (infix "\<^sup>*\\<^sup>*" 50) notation prfx (infix "\<^sup>*\\<^sup>*" 50) lemma sources_char\<^sub>P: shows "sources T = {A. Ide A \ Arr T \ Srcs A = Srcs T}" using Con_Ide_iff Con_sym con_char ide_char sources_def by fastforce lemma sources_cons: shows "Arr (t # T) \ sources (t # T) = sources [t]" apply (induct T) apply simp using sources_char\<^sub>P by auto lemma targets_char\<^sub>P: shows "targets T = {B. Ide B \ Arr T \ Srcs B = Trgs T}" unfolding targets_def by (metis (no_types, lifting) trg_def Arr.simps(1) Ide_implies_Arr Resid_Arr_self arr_char Con_Ide_iff Srcs_Resid con_char ide_char con_implies_arr(1)) lemma seq_char': shows "seq T U \ Arr T \ Arr U \ Trgs T \ Srcs U \ {}" proof show "seq T U \ Arr T \ Arr U \ Trgs T \ Srcs U \ {}" unfolding seq_def using Arr_has_Trg arr_char Con_Arr_self sources_char\<^sub>P trg_def trg_in_targets by fastforce assume 1: "Arr T \ Arr U \ Trgs T \ Srcs U \ {}" have "targets T = sources U" proof - obtain a where a: "R.ide a \ a \ Trgs T \ a \ Srcs U" using 1 Trgs_are_ide by blast have "Trgs [a] = Trgs T" using a 1 by (metis Con_single_ide_ind Con_sym Resid_Arr_Src Srcs_Resid Trgs_eqI) moreover have "Srcs [a] = Srcs U" using a 1 Con_single_ide_ind Con_imp_eq_Srcs by blast moreover have "Trgs [a] = Srcs [a]" using a by (metis R.residuation_axioms R.sources_resid Srcs.simps(2) Trgs.simps(2) residuation.ideE) ultimately show ?thesis using 1 sources_char\<^sub>P targets_char\<^sub>P by auto qed thus "seq T U" using 1 by blast qed lemma seq_char: shows "seq T U \ Arr T \ Arr U \ Trgs T = Srcs U" by (metis Int_absorb Srcs_Resid Arr_has_Src Arr_iff_Con_self Srcs_eqI seq_char') lemma seqI\<^sub>P [intro]: assumes "Arr T" and "Arr U" and "Trgs T \ Srcs U \ {}" shows "seq T U" using assms seq_char' by auto lemma Ide_imp_sources_eq_targets: assumes "Ide T" shows "sources T = targets T" using assms by (metis Resid_Arr_Ide_ind arr_iff_has_source arr_iff_has_target con_char arr_def sources_resid) subsection "Inclusion Map" text \ Inclusion of an RTS to the RTS of its paths. \ abbreviation incl where "incl \ \t. if R.arr t then [t] else null" lemma incl_is_simulation: shows "simulation resid Resid incl" using R.con_implies_arr(1-2) con_char R.arr_resid_iff_con null_char by unfold_locales auto lemma incl_is_injective: shows "inj_on incl (Collect R.arr)" by (intro inj_onI) simp lemma reflects_con: assumes "incl t \<^sup>*\\<^sup>* incl u" shows "t \ u" using assms by (metis (full_types) Arr.simps(1) Con_implies_Arr(1-2) Con_rec(1) null_char) end subsection "Composites of Paths" text \ The RTS of paths has composites, given by the append operation on lists. \ context paths_in_rts begin lemma Srcs_append [simp]: assumes "T \ []" shows "Srcs (T @ U) = Srcs T" by (metis Nil_is_append_conv Srcs.simps(2) Srcs.simps(3) assms hd_append list.exhaust_sel) lemma Trgs_append [simp]: shows "U \ [] \ Trgs (T @ U) = Trgs U" proof (induct T) show "U \ [] \ Trgs ([] @ U) = Trgs U" by auto show "\t T. \U \ [] \ Trgs (T @ U) = Trgs U; U \ []\ \ Trgs ((t # T) @ U) = Trgs U" by (metis Nil_is_append_conv Trgs.simps(3) append_Cons list.exhaust) qed lemma seq_implies_Trgs_eq_Srcs: shows "\Arr T; Arr U; Trgs T \ Srcs U\ \ Trgs T = Srcs U" by (metis inf.orderE Arr_has_Trg seqI\<^sub>P seq_char) lemma Arr_append_iff\<^sub>P: shows "\T \ []; U \ []\ \ Arr (T @ U) \ Arr T \ Arr U \ Trgs T \ Srcs U" proof (induct T arbitrary: U) show "\U. \[] \ []; U \ []\ \ Arr ([] @ U) = (Arr [] \ Arr U \ Trgs [] \ Srcs U)" by simp fix t T and U :: "'a list" assume ind: "\U. \T \ []; U \ []\ \ Arr (T @ U) = (Arr T \ Arr U \ Trgs T \ Srcs U)" assume U: "U \ []" show "Arr ((t # T) @ U) \ Arr (t # T) \ Arr U \ Trgs (t # T) \ Srcs U" proof (cases "T = []") show "T = [] \ ?thesis" using Arr.elims(1) U by auto assume T: "T \ []" have "Arr ((t # T) @ U) \ Arr (t # (T @ U))" by simp also have "... \ R.arr t \ Arr (T @ U) \ R.targets t \ Srcs (T @ U)" using T U by (metis Arr.simps(3) Nil_is_append_conv neq_Nil_conv) also have "... \ R.arr t \ Arr T \ Arr U \ Trgs T \ Srcs U \ R.targets t \ Srcs T" using T U ind by auto also have "... \ Arr (t # T) \ Arr U \ Trgs (t # T) \ Srcs U" using T U by (metis Arr.simps(3) Trgs.simps(3) neq_Nil_conv) finally show ?thesis by auto qed qed lemma Arr_consI\<^sub>P [intro, simp]: assumes "R.arr t" and "Arr U" and "R.targets t \ Srcs U" shows "Arr (t # U)" using assms Arr.elims(3) by blast lemma Arr_appendI\<^sub>P [intro, simp]: assumes "Arr T" and "Arr U" and "Trgs T \ Srcs U" shows "Arr (T @ U)" using assms by (metis Arr.simps(1) Arr_append_iff\<^sub>P) lemma Arr_appendE\<^sub>P [elim]: assumes "Arr (T @ U)" and "T \ []" and "U \ []" and "\Arr T; Arr U; Trgs T = Srcs U\ \ thesis" shows thesis using assms Arr_append_iff\<^sub>P seq_implies_Trgs_eq_Srcs by force lemma Ide_append_iff\<^sub>P: shows "\T \ []; U \ []\ \ Ide (T @ U) \ Ide T \ Ide U \ Trgs T \ Srcs U" using Ide_char by auto lemma Ide_appendI\<^sub>P [intro, simp]: assumes "Ide T" and "Ide U" and "Trgs T \ Srcs U" shows "Ide (T @ U)" using assms by (metis Ide.simps(1) Ide_append_iff\<^sub>P) lemma Resid_append_ind: shows "\T \ []; U \ []; V \ []\ \ (V @ T \<^sup>*\\<^sup>* U \ V \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* V) \ (T \<^sup>*\\<^sup>* V @ U \ T \<^sup>*\\<^sup>* V \ T \<^sup>*\\\<^sup>* V \<^sup>*\\<^sup>* U) \ (V @ T \<^sup>*\\<^sup>* U \ (V @ T) \<^sup>*\\\<^sup>* U = V \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* V)) \ (T \<^sup>*\\<^sup>* V @ U \ T \<^sup>*\\\<^sup>* (V @ U) = (T \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* U)" proof (induct V arbitrary: T U) show "\T U. \T \ []; U \ []; [] \ []\ \ ([] @ T \<^sup>*\\<^sup>* U \ [] \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* []) \ (T \<^sup>*\\<^sup>* [] @ U \ T \<^sup>*\\<^sup>* [] \ T \<^sup>*\\\<^sup>* [] \<^sup>*\\<^sup>* U) \ ([] @ T \<^sup>*\\<^sup>* U \ ([] @ T) \<^sup>*\\\<^sup>* U = [] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [])) \ (T \<^sup>*\\<^sup>* [] @ U \ T \<^sup>*\\\<^sup>* ([] @ U) = (T \<^sup>*\\\<^sup>* []) \<^sup>*\\\<^sup>* U)" by simp fix v :: 'a and T U V :: "'a list" assume ind: "\T U. \T \ []; U \ []; V \ []\ \ (V @ T \<^sup>*\\<^sup>* U \ V \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* V) \ (T \<^sup>*\\<^sup>* V @ U \ T \<^sup>*\\<^sup>* V \ T \<^sup>*\\\<^sup>* V \<^sup>*\\<^sup>* U) \ (V @ T \<^sup>*\\<^sup>* U \ (V @ T) \<^sup>*\\\<^sup>* U = V \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* V)) \ (T \<^sup>*\\<^sup>* V @ U \ T \<^sup>*\\\<^sup>* (V @ U) = (T \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* U)" assume T: "T \ []" and U: "U \ []" show "((v # V) @ T \<^sup>*\\<^sup>* U \ (v # V) \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* (v # V)) \ (T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\<^sup>* (v # V) \ T \<^sup>*\\\<^sup>* (v # V) \<^sup>*\\<^sup>* U) \ ((v # V) @ T \<^sup>*\\<^sup>* U \ ((v # V) @ T) \<^sup>*\\\<^sup>* U = (v # V) \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* (v # V))) \ (T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\\<^sup>* ((v # V) @ U) = (T \<^sup>*\\\<^sup>* (v # V)) \<^sup>*\\\<^sup>* U)" proof (intro conjI iffI impI) show 1: "(v # V) @ T \<^sup>*\\<^sup>* U \ ((v # V) @ T) \<^sup>*\\\<^sup>* U = (v # V) \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* (v # V))" proof (cases "V = []") show "V = [] \ (v # V) @ T \<^sup>*\\<^sup>* U \ ?thesis" using T U Resid_cons(1) U by auto assume V: "V \ []" assume Con: "(v # V) @ T \<^sup>*\\<^sup>* U" have "((v # V) @ T) \<^sup>*\\\<^sup>* U = (v # (V @ T)) \<^sup>*\\\<^sup>* U" by simp also have "... = [v] \<^sup>*\\\<^sup>* U @ (V @ T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [v])" using T U Con Resid_cons by simp also have "... = [v] \<^sup>*\\\<^sup>* U @ V \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [v]) @ T \<^sup>*\\\<^sup>* ((U \<^sup>*\\\<^sup>* [v]) \<^sup>*\\\<^sup>* V)" using T U V Con ind Resid_cons by (metis Con_sym Cons_eq_appendI append_is_Nil_conv Con_cons(1)) also have "... = (v # V) \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* (v # V))" by (metis Con Con_cons(2) Cons_eq_appendI Resid_cons(1) Resid_cons(2) T U V append.assoc append_is_Nil_conv Con_sym ind) finally show ?thesis by simp qed show 2: "T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\\<^sup>* ((v # V) @ U) = (T \<^sup>*\\\<^sup>* (v # V)) \<^sup>*\\\<^sup>* U" proof (cases "V = []") show "V = [] \ T \<^sup>*\\<^sup>* (v # V) @ U \ ?thesis" using Resid_cons(2) T U by auto assume V: "V \ []" assume Con: "T \<^sup>*\\<^sup>* (v # V) @ U" have "T \<^sup>*\\\<^sup>* ((v # V) @ U) = T \<^sup>*\\\<^sup>* (v # (V @ U))" by simp also have 1: "... = (T \<^sup>*\\\<^sup>* [v]) \<^sup>*\\\<^sup>* (V @ U)" using V Con Resid_cons(2) T by force also have "... = ((T \<^sup>*\\\<^sup>* [v]) \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* U" using T U V 1 Con ind by (metis Con_initial_right Cons_eq_appendI) also have "... = (T \<^sup>*\\\<^sup>* (v # V)) \<^sup>*\\\<^sup>* U" using T V Con by (metis Con_cons(2) Con_initial_right Cons_eq_appendI Resid_cons(2)) finally show ?thesis by blast qed show "(v # V) @ T \<^sup>*\\<^sup>* U \ v # V \<^sup>*\\<^sup>* U" by (metis 1 Con_sym Resid.simps(1) append_Nil) show "(v # V) @ T \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* (v # V)" using T U Con_sym by (metis 1 Con_initial_right Resid_cons(1-2) append.simps(2) ind self_append_conv) show "T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\<^sup>* v # V" using 2 by fastforce show "T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\\<^sup>* (v # V) \<^sup>*\\<^sup>* U" using 2 by fastforce show "T \<^sup>*\\<^sup>* v # V \ T \<^sup>*\\\<^sup>* (v # V) \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* (v # V) @ U" proof - assume Con: "T \<^sup>*\\<^sup>* v # V \ T \<^sup>*\\\<^sup>* (v # V) \<^sup>*\\<^sup>* U" have "T \<^sup>*\\<^sup>* (v # V) @ U \ T \<^sup>*\\<^sup>* v # (V @ U)" by simp also have "... \ T \<^sup>*\\<^sup>* [v] \ T \<^sup>*\\\<^sup>* [v] \<^sup>*\\<^sup>* V @ U" using T U Con_cons(2) by simp also have "... \ T \<^sup>*\\\<^sup>* [v] \<^sup>*\\<^sup>* V @ U" by fastforce also have "... \ True" using Con ind by (metis Con_cons(2) Resid_cons(2) T U self_append_conv2) finally show ?thesis by blast qed show "v # V \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* (v # V) \ (v # V) @ T \<^sup>*\\<^sup>* U" proof - assume Con: "v # V \<^sup>*\\<^sup>* U \ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* (v # V)" have "(v # V) @ T \<^sup>*\\<^sup>* U \v # (V @ T) \<^sup>*\\<^sup>* U" by simp also have "... \ [v] \<^sup>*\\<^sup>* U \ V @ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [v]" using T U Con_cons(1) by simp also have "... \ V @ T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [v]" by (metis Con Con_cons(1) U) also have "... \ True" using Con ind by (metis Con_cons(1) Con_sym Resid_cons(2) T U append_self_conv2) finally show ?thesis by blast qed qed qed lemma Con_append: assumes "T \ []" and "U \ []" and "V \ []" shows "T @ U \<^sup>*\\<^sup>* V \ T \<^sup>*\\<^sup>* V \ U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* T" and "T \<^sup>*\\<^sup>* U @ V \ T \<^sup>*\\<^sup>* U \ T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V" using assms Resid_append_ind by blast+ lemma Con_appendI [intro]: shows "\T \<^sup>*\\<^sup>* V; U \<^sup>*\\<^sup>* V \<^sup>*\\\<^sup>* T\ \ T @ U \<^sup>*\\<^sup>* V" and "\T \<^sup>*\\<^sup>* U; T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* V\ \ T \<^sup>*\\<^sup>* U @ V" by (metis Con_append(1) Con_sym Resid.simps(1))+ lemma Resid_append [intro, simp]: shows "\T \ []; T @ U \<^sup>*\\<^sup>* V\ \ (T @ U) \<^sup>*\\\<^sup>* V = (T \<^sup>*\\\<^sup>* V) @ (U \<^sup>*\\\<^sup>* (V \<^sup>*\\\<^sup>* T))" and "\U \ []; V \ []; T \<^sup>*\\<^sup>* U @ V\ \ T \<^sup>*\\\<^sup>* (U @ V) = (T \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* V" using Resid_append_ind apply (metis Con_sym Resid.simps(1) append_self_conv) using Resid_append_ind by (metis Resid.simps(1)) lemma Resid_append2 [simp]: assumes "T \ []" and "U \ []" and "V \ []" and "W \ []" and "T @ U \<^sup>*\\<^sup>* V @ W" shows "(T @ U) \<^sup>*\\\<^sup>* (V @ W) = (T \<^sup>*\\\<^sup>* V) \<^sup>*\\\<^sup>* W @ (U \<^sup>*\\\<^sup>* (V \<^sup>*\\\<^sup>* T)) \<^sup>*\\\<^sup>* (W \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* V))" using assms Resid_append by (metis Con_append(1-2) append_is_Nil_conv) lemma append_is_composite_of: assumes "seq T U" shows "composite_of T U (T @ U)" unfolding composite_of_def using assms apply (intro conjI) apply (metis Arr.simps(1) Resid_Arr_self Resid_Ide_Arr_ind Arr_appendI\<^sub>P Resid_append_ind ide_char order_refl seq_char) apply (metis Arr.simps(1) Arr_appendI\<^sub>P Con_Arr_self Resid_Arr_self Resid_append_ind ide_char seq_char order_refl) by (metis Arr.simps(1) Con_Arr_self Con_append(1) Resid_Arr_self Arr_appendI\<^sub>P Ide_append_iff\<^sub>P Resid_append(1) ide_char seq_char order_refl) sublocale rts_with_composites Resid using append_is_composite_of composable_def by unfold_locales blast theorem is_rts_with_composites: shows "rts_with_composites Resid" .. (* TODO: This stuff might be redundant. *) lemma arr_append [intro, simp]: assumes "seq T U" shows "arr (T @ U)" using assms arrI\<^sub>P seq_char by simp lemma arr_append_imp_seq: assumes "T \ []" and "U \ []" and "arr (T @ U)" shows "seq T U" using assms arr_char seq_char Arr_append_iff\<^sub>P seq_implies_Trgs_eq_Srcs by simp lemma sources_append [simp]: assumes "seq T U" shows "sources (T @ U) = sources T" using assms by (meson append_is_composite_of sources_composite_of) lemma targets_append [simp]: assumes "seq T U" shows "targets (T @ U) = targets U" using assms by (meson append_is_composite_of targets_composite_of) lemma cong_respects_seq\<^sub>P: assumes "seq T U" and "T \<^sup>*\\<^sup>* T'" and "U \<^sup>*\\<^sup>* U'" shows "seq T' U'" by (meson assms cong_respects_seq) lemma cong_append [intro]: assumes "seq T U" and "T \<^sup>*\\<^sup>* T'" and "U \<^sup>*\\<^sup>* U'" shows "T @ U \<^sup>*\\<^sup>* T' @ U'" proof have 1: "\T U T' U'. \seq T U; T \<^sup>*\\<^sup>* T'; U \<^sup>*\\<^sup>* U'\ \ seq T' U'" using assms cong_respects_seq\<^sub>P by simp have 2: "\T U T' U'. \seq T U; T \<^sup>*\\<^sup>* T'; U \<^sup>*\\<^sup>* U'\ \ T @ U \<^sup>*\\<^sup>* T' @ U'" proof - fix T U T' U' assume TU: "seq T U" and TT': "T \<^sup>*\\<^sup>* T'" and UU': "U \<^sup>*\\<^sup>* U'" have T'U': "seq T' U'" using TU TT' UU' cong_respects_seq\<^sub>P by simp have 3: "Ide (T \<^sup>*\\\<^sup>* T') \ Ide (T' \<^sup>*\\\<^sup>* T) \ Ide (U \<^sup>*\\\<^sup>* U') \ Ide (U' \<^sup>*\\\<^sup>* U)" using TU TT' UU' ide_char by blast have "(T @ U) \<^sup>*\\\<^sup>* (T' @ U') = ((T \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* U') @ U \<^sup>*\\\<^sup>* ((T' \<^sup>*\\\<^sup>* T) @ U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T'))" proof - have 4: "T \ [] \ U \ [] \ T' \ [] \ U' \ []" using TU TT' UU' Arr.simps(1) seq_char ide_char by auto moreover have "(T @ U) \<^sup>*\\\<^sup>* (T' @ U') \ []" proof (intro Con_appendI) show "T \<^sup>*\\\<^sup>* T' \ []" using "3" by force show "(T \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* U' \ []" using "3" T'U' \T \<^sup>*\\<^sup>* T' \ []\ Con_Ide_iff seq_char by fastforce show "U \<^sup>*\\\<^sup>* ((T' @ U') \<^sup>*\\\<^sup>* T) \ []" proof - have "U \<^sup>*\\\<^sup>* ((T' @ U') \<^sup>*\\\<^sup>* T) = U \<^sup>*\\\<^sup>* ((T' \<^sup>*\\\<^sup>* T) @ U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T'))" by (metis Con_appendI(1) Resid_append(1) \(T \<^sup>*\\<^sup>* T') \<^sup>*\\<^sup>* U' \ []\ \T \<^sup>*\\<^sup>* T' \ []\ calculation Con_sym) also have "... = (U \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* T)) \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T'))" by (metis Arr.simps(1) Con_append(2) Resid_append(2) \(T \<^sup>*\\<^sup>* T') \<^sup>*\\<^sup>* U' \ []\ Con_implies_Arr(1) Con_sym) also have "... = U \<^sup>*\\\<^sup>* U'" by (metis (mono_tags, lifting) "3" Ide.simps(1) Resid_Ide(1) Srcs_Resid TU \(T \<^sup>*\\<^sup>* T') \<^sup>*\\<^sup>* U' \ []\ Con_Ide_iff seq_char) finally show ?thesis using 3 UU' by force qed qed ultimately show ?thesis using Resid_append2 [of T U T' U'] seq_char by (metis Con_append(2) Con_sym Resid_append(2) Resid.simps(1)) qed moreover have "Ide ..." proof have 3: "Ide (T \<^sup>*\\\<^sup>* T') \ Ide (T' \<^sup>*\\\<^sup>* T) \ Ide (U \<^sup>*\\\<^sup>* U') \ Ide (U' \<^sup>*\\\<^sup>* U)" using TU TT' UU' ide_char by blast show 4: "Ide ((T \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* U')" using TU T'U' TT' UU' 1 3 by (metis (full_types) Srcs_Resid Con_Ide_iff Resid_Ide_Arr_ind seq_char) show 5: "Ide (U \<^sup>*\\\<^sup>* ((T' \<^sup>*\\\<^sup>* T) @ U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T')))" proof - have "U \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* T) = U" by (metis (full_types) "3" TT' TU Con_Ide_iff Resid_Ide(1) Srcs_Resid con_char seq_char prfx_implies_con) moreover have "U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T') = U'" by (metis "3" "4" Ide.simps(1) Resid_Ide(1)) ultimately show ?thesis by (metis "3" "4" Arr.simps(1) Con_append(2) Ide.simps(1) Resid_append(2) TU Con_sym seq_char) qed show "Trgs ((T \<^sup>*\\\<^sup>* T') \<^sup>*\\\<^sup>* U') \ Srcs (U \<^sup>*\\\<^sup>* (T' \<^sup>*\\\<^sup>* T @ U' \<^sup>*\\\<^sup>* (T \<^sup>*\\\<^sup>* T')))" by (metis 4 5 Arr_append_iff\<^sub>P Ide.simps(1) Nil_is_append_conv calculation Con_imp_Arr_Resid) qed ultimately show "T @ U \<^sup>*\\<^sup>* T' @ U'" using ide_char by presburger qed show "T @ U \<^sup>*\\<^sup>* T' @ U'" using assms 2 by simp show "T' @ U' \<^sup>*\\<^sup>* T @ U" using assms 1 2 cong_symmetric by blast qed lemma cong_cons [intro]: assumes "seq [t] U" and "t \ t'" and "U \<^sup>*\\<^sup>* U'" shows "t # U \<^sup>*\\<^sup>* t' # U'" using assms cong_append [of "[t]" U "[t']" U'] by (simp add: R.prfx_implies_con ide_char) lemma cong_append_ideI [intro]: assumes "seq T U" shows "ide T \ T @ U \<^sup>*\\<^sup>* U" and "ide U \ T @ U \<^sup>*\\<^sup>* T" and "ide T \ U \<^sup>*\\<^sup>* T @ U" and "ide U \ T \<^sup>*\\<^sup>* T @ U" proof - show 1: "ide T \ T @ U \<^sup>*\\<^sup>* U" using assms by (metis append_is_composite_of composite_ofE resid_arr_ide prfx_implies_con con_sym) show 2: "ide U \ T @ U \<^sup>*\\<^sup>* T" by (meson assms append_is_composite_of composite_ofE ide_backward_stable) show "ide T \ U \<^sup>*\\<^sup>* T @ U" using 1 cong_symmetric by auto show "ide U \ T \<^sup>*\\<^sup>* T @ U" using 2 cong_symmetric by auto qed lemma cong_cons_ideI [intro]: assumes "seq [t] U" and "R.ide t" shows "t # U \<^sup>*\\<^sup>* U" and "U \<^sup>*\\<^sup>* t # U" using assms cong_append_ideI [of "[t]" U] by (auto simp add: ide_char) lemma prfx_decomp: assumes "[t] \<^sup>*\\<^sup>* [u]" shows "[t] @ [u \\ t] \<^sup>*\\<^sup>* [u]" proof (* TODO: I really want these to be doable by auto. *) show 1: "[u] \<^sup>*\\<^sup>* [t] @ [u \\ t]" using assms by (metis Con_imp_Arr_Resid Con_rec(3) Resid.simps(3) Resid_rec(3) R.con_sym append.left_neutral append_Cons arr_char cong_reflexive list.distinct(1)) show "[t] @ [u \\ t] \<^sup>*\\<^sup>* [u]" proof - have "([t] @ [u \\ t]) \<^sup>*\\\<^sup>* [u] = ([t] \<^sup>*\\\<^sup>* [u]) @ ([u \\ t] \<^sup>*\\\<^sup>* [u \\ t])" using assms by (metis Arr_Resid_single Con_Arr_self Con_appendI(1) Con_sym Resid_append(1) Resid_rec(1) con_char list.discI prfx_implies_con) moreover have "Ide ..." using assms by (metis 1 Con_sym append_Nil2 arr_append_imp_seq calculation cong_append_ideI(4) ide_backward_stable Con_implies_Arr(2) Resid_Arr_self con_char ide_char prfx_implies_con arr_resid_iff_con) ultimately show ?thesis using ide_char by presburger qed qed lemma composite_of_single_single: assumes "R.composite_of t u v" shows "composite_of [t] [u] ([t] @ [u])" proof show "[t] \<^sup>*\\<^sup>* [t] @ [u]" proof - have "[t] \<^sup>*\\\<^sup>* ([t] @ [u]) = ([t] \<^sup>*\\\<^sup>* [t]) \<^sup>*\\\<^sup>* [u]" using assms by auto moreover have "Ide ..." by (metis (no_types, lifting) Con_implies_Arr(2) R.bounded_imp_con R.con_composite_of_iff R.con_prfx_composite_of(1) assms resid_ide_arr Con_rec(1) Resid.simps(3) Resid_Arr_self con_char ide_char) ultimately show ?thesis using ide_char by presburger qed show "([t] @ [u]) \<^sup>*\\\<^sup>* [t] \<^sup>*\\<^sup>* [u]" using assms by (metis \prfx [t] ([t] @ [u])\ append_is_composite_of arr_append_imp_seq composite_ofE con_def not_Cons_self2 Con_implies_Arr(2) arr_char null_char prfx_implies_con) qed end subsection "Paths in a Weakly Extensional RTS" locale paths_in_weakly_extensional_rts = R: weakly_extensional_rts + paths_in_rts begin lemma ex_un_Src: assumes "Arr T" shows "\!a. a \ Srcs T" using assms by (simp add: R.weakly_extensional_rts_axioms Srcs_simp\<^sub>P R.arr_has_un_source) fun Src where "Src T = R.src (hd T)" lemma Srcs_simp\<^sub>P\<^sub>W\<^sub>E: assumes "Arr T" shows "Srcs T = {Src T}" proof - have "[R.src (hd T)] \ sources T" by (metis Arr_imp_arr_hd Con_single_ide_ind Ide.simps(2) Srcs_simp\<^sub>P assms con_char ide_char in_sourcesI con_sym R.ide_src R.src_in_sources) hence "R.src (hd T) \ Srcs T" using assms by (metis Srcs.elims Arr_has_Src list.sel(1) R.arr_iff_has_source R.src_in_sources) thus ?thesis using assms ex_un_Src by auto qed lemma ex_un_Trg: assumes "Arr T" shows "\!b. b \ Trgs T" using assms apply (induct T) apply auto[1] by (metis Con_Arr_self Ide_implies_Arr Resid_Arr_self Srcs_Resid ex_un_Src) fun Trg where "Trg [] = R.null" | "Trg [t] = R.trg t" | "Trg (t # T) = Trg T" lemma Trg_simp [simp]: shows "T \ [] \ Trg T = R.trg (last T)" apply (induct T) apply auto by (metis Trg.simps(3) list.exhaust_sel) lemma Trgs_simp\<^sub>P\<^sub>W\<^sub>E [simp]: assumes "Arr T" shows "Trgs T = {Trg T}" using assms by (metis Arr_imp_arr_last Con_Arr_self Con_imp_Arr_Resid R.trg_in_targets Srcs.simps(1) Srcs_Resid Srcs_simp\<^sub>P\<^sub>W\<^sub>E Trg_simp insertE insert_absorb insert_not_empty Trgs_simp\<^sub>P) lemma Src_resid [simp]: assumes "T \<^sup>*\\<^sup>* U" shows "Src (T \<^sup>*\\\<^sup>* U) = Trg U" using assms Con_imp_Arr_Resid Con_implies_Arr(2) Srcs_Resid Srcs_simp\<^sub>P\<^sub>W\<^sub>E by force lemma Trg_resid_sym: assumes "T \<^sup>*\\<^sup>* U" shows "Trg (T \<^sup>*\\\<^sup>* U) = Trg (U \<^sup>*\\\<^sup>* T)" using assms Con_imp_Arr_Resid Con_sym Trgs_Resid_sym by auto lemma Src_append [simp]: assumes "seq T U" shows "Src (T @ U) = Src T" using assms by (metis Arr.simps(1) Src.simps hd_append seq_char) lemma Trg_append [simp]: assumes "seq T U" shows "Trg (T @ U) = Trg U" using assms by (metis Ide.simps(1) Resid.simps(1) Trg_simp append_is_Nil_conv ide_char ide_trg last_appendR seqE trg_def) lemma Arr_append_iff\<^sub>P\<^sub>W\<^sub>E: assumes "T \ []" and "U \ []" shows "Arr (T @ U) \ Arr T \ Arr U \ Trg T = Src U" using assms Arr_appendE\<^sub>P Srcs_simp\<^sub>P\<^sub>W\<^sub>E by auto lemma Arr_consI\<^sub>P\<^sub>W\<^sub>E [intro, simp]: assumes "R.arr t" and "Arr U" and "R.trg t = Src U" shows "Arr (t # U)" using assms by (metis Arr.simps(2) Srcs_simp\<^sub>P\<^sub>W\<^sub>E Trg.simps(2) Trgs.simps(2) Trgs_simp\<^sub>P\<^sub>W\<^sub>E dual_order.eq_iff Arr_consI\<^sub>P) lemma Arr_consE [elim]: assumes "Arr (t # U)" and "\R.arr t; U \ [] \ Arr U; U \ [] \ R.trg t = Src U\ \ thesis" shows thesis using assms by (metis Arr_append_iff\<^sub>P\<^sub>W\<^sub>E Trg.simps(2) append_Cons append_Nil list.distinct(1) Arr.simps(2)) lemma Arr_appendI\<^sub>P\<^sub>W\<^sub>E [intro, simp]: assumes "Arr T" and "Arr U" and "Trg T = Src U" shows "Arr (T @ U)" using assms by (metis Arr.simps(1) Arr_append_iff\<^sub>P\<^sub>W\<^sub>E) lemma Arr_appendE\<^sub>P\<^sub>W\<^sub>E [elim]: assumes "Arr (T @ U)" and "T \ []" and "U \ []" and "\Arr T; Arr U; Trg T = Src U\ \ thesis" shows thesis using assms Arr_append_iff\<^sub>P\<^sub>W\<^sub>E seq_implies_Trgs_eq_Srcs by force lemma Ide_append_iff\<^sub>P\<^sub>W\<^sub>E: assumes "T \ []" and "U \ []" shows "Ide (T @ U) \ Ide T \ Ide U \ Trg T = Src U" using assms Ide_char by auto lemma Ide_appendI\<^sub>P\<^sub>W\<^sub>E [intro, simp]: assumes "Ide T" and "Ide U" and "Trg T = Src U" shows "Ide (T @ U)" using assms by (metis Ide.simps(1) Ide_append_iff\<^sub>P\<^sub>W\<^sub>E) lemma Ide_appendE [elim]: assumes "Ide (T @ U)" and "T \ []" and "U \ []" and "\Ide T; Ide U; Trg T = Src U\ \ thesis" shows thesis using assms Ide_append_iff\<^sub>P\<^sub>W\<^sub>E by metis lemma Ide_consI [intro, simp]: assumes "R.ide t" and "Ide U" and "R.trg t = Src U" shows "Ide (t # U)" using assms by (simp add: Ide_char) lemma Ide_consE [elim]: assumes "Ide (t # U)" and "\R.ide t; U \ [] \ Ide U; U \ [] \ R.trg t = Src U\ \ thesis" shows thesis using assms by (metis Con_rec(4) Ide.simps(2) Ide_imp_Ide_hd Ide_imp_Ide_tl R.trg_def R.trg_ide Resid_Arr_Ide_ind Trg.simps(2) ide_char list.sel(1) list.sel(3) list.simps(3) Src_resid ide_def) lemma Ide_imp_Src_eq_Trg: assumes "Ide T" shows "Src T = Trg T" using assms by (metis Ide.simps(1) Src_resid ide_char ide_def) end subsection "Paths in a Confluent RTS" text \ Here we show that confluence of an RTS extends to confluence of the RTS of its paths. \ locale paths_in_confluent_rts = paths_in_rts + R: confluent_rts begin lemma confluence_single: assumes "\t u. R.coinitial t u \ t \ u" shows "\R.arr t; Arr U; R.sources t = Srcs U\ \ [t] \<^sup>*\\<^sup>* U" proof (induct U arbitrary: t) show "\t. \R.arr t; Arr []; R.sources t = Srcs []\ \ [t] \<^sup>*\\<^sup>* []" by simp fix t u U assume ind: "\t. \R.arr t; Arr U; R.sources t = Srcs U\ \ [t] \<^sup>*\\<^sup>* U" assume t: "R.arr t" assume uU: "Arr (u # U)" assume coinitial: "R.sources t = Srcs (u # U)" hence 1: "R.coinitial t u" using t uU by (metis Arr.simps(2) Con_implies_Arr(1) Con_imp_eq_Srcs Con_initial_left Srcs.simps(2) Con_Arr_self R.coinitial_iff) show "[t] \<^sup>*\\<^sup>* u # U" proof (cases "U = []") show "U = [] \ ?thesis" using assms t uU coinitial R.coinitial_iff by fastforce assume U: "U \ []" show ?thesis proof - have 2: "Arr [t \\ u] \ Arr U \ Srcs [t \\ u] = Srcs U" using assms 1 t uU U R.arr_resid_iff_con apply (intro conjI) apply simp apply (metis Con_Arr_self Con_implies_Arr(2) Resid_cons(2)) by (metis (full_types) Con_cons(2) Srcs.simps(2) Srcs_Resid Trgs.simps(2) Con_Arr_self Con_imp_eq_Srcs list.simps(3) R.sources_resid) have "[t] \<^sup>*\\<^sup>* u # U \ t \ u \ [t \\ u] \<^sup>*\\<^sup>* U" using U Con_rec(3) [of U t u] by simp also have "... \ True" using assms t uU U 1 2 ind by force finally show ?thesis by blast qed qed qed lemma confluence_ind: shows "\Arr T; Arr U; Srcs T = Srcs U\ \ T \<^sup>*\\<^sup>* U" proof (induct T arbitrary: U) show "\U. \Arr []; Arr U; Srcs [] = Srcs U\ \ [] \<^sup>*\\<^sup>* U" by simp fix t T U assume ind: "\U. \Arr T; Arr U; Srcs T = Srcs U\ \ T \<^sup>*\\<^sup>* U" assume tT: "Arr (t # T)" assume U: "Arr U" assume coinitial: "Srcs (t # T) = Srcs U" show "t # T \<^sup>*\\<^sup>* U" proof (cases "T = []") show "T = [] \ ?thesis" using U tT coinitial confluence_single [of t U] R.confluence by simp assume T: "T \ []" show ?thesis proof - have 1: "[t] \<^sup>*\\<^sup>* U" using tT U coinitial R.confluence by (metis R.arr_def Srcs.simps(2) T Con_Arr_self Con_imp_eq_Srcs Con_initial_right Con_rec(4) confluence_single) moreover have "T \<^sup>*\\<^sup>* U \<^sup>*\\\<^sup>* [t]" using 1 tT U T coinitial ind [of "U \<^sup>*\\\<^sup>* [t]"] by (metis (full_types) Con_imp_Arr_Resid Arr_iff_Con_self Con_implies_Arr(2) Con_imp_eq_Srcs Con_sym R.sources_resid Srcs.simps(2) Srcs_Resid Trgs.simps(2) Con_rec(4)) ultimately show ?thesis using Con_cons(1) [of T U t] by fastforce qed qed qed lemma confluence\<^sub>P: assumes "coinitial T U" shows "con T U" using assms confluence_ind sources_char\<^sub>P coinitial_def con_char by auto sublocale confluent_rts Resid apply (unfold_locales) using confluence\<^sub>P by simp lemma is_confluent_rts: shows "confluent_rts Resid" .. end subsection "Simulations Lift to Paths" text \ In this section we show that a simulation from RTS \A\ to RTS \B\ determines a simulation from the RTS of paths in \A\ to the RTS of paths in \B\. In other words, the path-RTS construction is functorial with respect to simulation. \ context simulation begin interpretation P\<^sub>A: paths_in_rts A .. interpretation P\<^sub>B: paths_in_rts B .. lemma map_Resid_single: shows "P\<^sub>A.con T [u] \ map F (P\<^sub>A.Resid T [u]) = P\<^sub>B.Resid (map F T) [F u]" apply (induct T arbitrary: u) apply simp proof - fix t u T assume ind: "\u. P\<^sub>A.con T [u] \ map F (P\<^sub>A.Resid T [u]) = P\<^sub>B.Resid (map F T) [F u]" assume 1: "P\<^sub>A.con (t # T) [u]" show "map F (P\<^sub>A.Resid (t # T) [u]) = P\<^sub>B.Resid (map F (t # T)) [F u]" proof (cases "T = []") show "T = [] \ ?thesis" using "1" P\<^sub>A.null_char by fastforce assume T: "T \ []" show ?thesis using T 1 ind P\<^sub>A.con_def P\<^sub>A.null_char P\<^sub>A.Con_rec(2) P\<^sub>A.Resid_rec(2) P\<^sub>B.Con_rec(2) P\<^sub>B.Resid_rec(2) apply simp by (metis A.con_sym Nil_is_map_conv preserves_con preserves_resid) qed qed lemma map_Resid: shows "P\<^sub>A.con T U \ map F (P\<^sub>A.Resid T U) = P\<^sub>B.Resid (map F T) (map F U)" apply (induct U arbitrary: T) using P\<^sub>A.Resid.simps(1) P\<^sub>A.con_char P\<^sub>A.con_sym apply blast proof - fix u U T assume ind: "\T. P\<^sub>A.con T U \ map F (P\<^sub>A.Resid T U) = P\<^sub>B.Resid (map F T) (map F U)" assume 1: "P\<^sub>A.con T (u # U)" show "map F (P\<^sub>A.Resid T (u # U)) = P\<^sub>B.Resid (map F T) (map F (u # U))" proof (cases "U = []") show "U = [] \ ?thesis" using "1" map_Resid_single by force assume U: "U \ []" have "P\<^sub>B.Resid (map F T) (map F (u # U)) = P\<^sub>B.Resid (P\<^sub>B.Resid (map F T) [F u]) (map F U)" using U 1 P\<^sub>B.Resid_cons(2) apply simp by (metis P\<^sub>B.Arr.simps(1) P\<^sub>B.Con_consI(2) P\<^sub>B.Con_implies_Arr(1) list.map_disc_iff) also have "... = map F (P\<^sub>A.Resid (P\<^sub>A.Resid T [u]) U)" using U 1 ind by (metis P\<^sub>A.Con_initial_right P\<^sub>A.Resid_cons(2) P\<^sub>A.con_char map_Resid_single) also have "... = map F (P\<^sub>A.Resid T (u # U))" using "1" P\<^sub>A.Resid_cons(2) P\<^sub>A.con_char U by auto finally show ?thesis by simp qed qed lemma preserves_paths: shows "P\<^sub>A.Arr T \ P\<^sub>B.Arr (map F T)" by (metis P\<^sub>A.Con_Arr_self P\<^sub>A.conI\<^sub>P P\<^sub>B.Arr_iff_Con_self map_Resid map_is_Nil_conv) interpretation Fx: simulation P\<^sub>A.Resid P\<^sub>B.Resid \\T. if P\<^sub>A.Arr T then map F T else []\ proof let ?Fx = "\T. if P\<^sub>A.Arr T then map F T else []" show "\T. \ P\<^sub>A.arr T \ ?Fx T = P\<^sub>B.null" by (simp add: P\<^sub>A.arr_char P\<^sub>B.null_char) show "\T U. P\<^sub>A.con T U \ P\<^sub>B.con (?Fx T) (?Fx U)" using P\<^sub>A.Con_implies_Arr(1) P\<^sub>A.Con_implies_Arr(2) P\<^sub>A.con_char map_Resid by fastforce show "\T U. P\<^sub>A.con T U \ ?Fx (P\<^sub>A.Resid T U) = P\<^sub>B.Resid (?Fx T) (?Fx U)" by (simp add: P\<^sub>A.Con_imp_Arr_Resid P\<^sub>A.Con_implies_Arr(1) P\<^sub>A.Con_implies_Arr(2) P\<^sub>A.con_char map_Resid) qed lemma lifts_to_paths: shows "simulation P\<^sub>A.Resid P\<^sub>B.Resid (\T. if P\<^sub>A.Arr T then map F T else [])" .. end subsection "Normal Sub-RTS's Lift to Paths" text \ Here we show that a normal sub-RTS \N\ of an RTS \R\ lifts to a normal sub-RTS of the RTS of paths in \N\, and that it is coherent if \N\ is. \ locale paths_in_rts_with_normal = R: rts + N: normal_sub_rts + paths_in_rts begin text \ We define a ``normal path'' to be a path that consists entirely of normal transitions. We show that the collection of all normal paths is a normal sub-RTS of the RTS of paths. \ definition NPath where "NPath T \ (Arr T \ set T \ \)" lemma Ide_implies_NPath: assumes "Ide T" shows "NPath T" using assms by (metis Ball_Collect NPath_def Ide_implies_Arr N.ide_closed set_Ide_subset_ide subsetI) lemma NPath_implies_Arr: assumes "NPath T" shows "Arr T" using assms NPath_def by simp lemma NPath_append: assumes "T \ []" and "U \ []" shows "NPath (T @ U) \ NPath T \ NPath U \ Trgs T \ Srcs U" using assms NPath_def by auto lemma NPath_appendI [intro, simp]: assumes "NPath T" and "NPath U" and "Trgs T \ Srcs U" shows "NPath (T @ U)" using assms NPath_def by simp lemma NPath_Resid_single_Arr: shows "\t \ \; Arr U; R.sources t = Srcs U\ \ NPath (Resid [t] U)" proof (induct U arbitrary: t) show "\t. \t \ \; Arr []; R.sources t = Srcs []\ \ NPath (Resid [t] [])" by simp fix t u U assume ind: "\t. \t \ \; Arr U; R.sources t = Srcs U\ \ NPath (Resid [t] U)" assume t: "t \ \" assume uU: "Arr (u # U)" assume src: "R.sources t = Srcs (u # U)" show "NPath (Resid [t] (u # U))" proof (cases "U = []") show "U = [] \ ?thesis" using NPath_def t src apply simp by (metis Arr.simps(2) R.arr_resid_iff_con R.coinitialI N.forward_stable N.elements_are_arr uU) assume U: "U \ []" show ?thesis proof - have "NPath (Resid [t] (u # U)) \ NPath (Resid [t \\ u] U)" using t U uU src by (metis Arr.simps(2) Con_implies_Arr(1) Resid_rec(3) Con_rec(3) R.arr_resid_iff_con) also have "... \ True" proof - have "t \\ u \ \" using t U uU src N.forward_stable [of t u] by (metis Con_Arr_self Con_imp_eq_Srcs Con_initial_left Srcs.simps(2) inf.idem Arr_has_Src R.coinitial_def) moreover have "Arr U" using U uU by (metis Arr.simps(3) neq_Nil_conv) moreover have "R.sources (t \\ u) = Srcs U" using t uU src by (metis Con_Arr_self Srcs.simps(2) U calculation(1) Con_imp_eq_Srcs Con_rec(4) N.elements_are_arr R.sources_resid R.arr_resid_iff_con) ultimately show ?thesis using ind [of "t \\ u"] by simp qed finally show ?thesis by blast qed qed qed lemma NPath_Resid_Arr_single: shows "\ NPath T; R.arr u; Srcs T = R.sources u \ \ NPath (Resid T [u])" proof (induct T arbitrary: u) show "\u. \NPath []; R.arr u; Srcs [] = R.sources u\ \ NPath (Resid [] [u])" by simp fix t u T assume ind: "\u. \NPath T; R.arr u; Srcs T = R.sources u\ \ NPath (Resid T [u])" assume tT: "NPath (t # T)" assume u: "R.arr u" assume src: "Srcs (t # T) = R.sources u" show "NPath (Resid (t # T) [u])" proof (cases "T = []") show "T = [] \ ?thesis" using tT u src NPath_def by (metis Arr.simps(2) NPath_Resid_single_Arr Srcs.simps(2) list.set_intros(1) subsetD) assume T: "T \ []" have "R.coinitial u t" by (metis R.coinitialI Srcs.simps(3) T list.exhaust_sel src u) hence con: "t \ u" using tT T u src R.con_sym NPath_def by (metis N.forward_stable N.elements_are_arr R.not_arr_null list.set_intros(1) R.conI subsetD) have 1: "NPath (Resid (t # T) [u]) \ NPath ((t \\ u) # Resid T [u \\ t])" proof - have "t # T \<^sup>*\\<^sup>* [u]" proof - have 2: "[t] \<^sup>*\\<^sup>* [u]" by (simp add: Con_rec(1) con) moreover have "T \<^sup>*\\<^sup>* Resid [u] [t]" proof - have "NPath T" using tT T NPath_def by (metis NPath_append append_Cons append_Nil) moreover have 3: "R.arr (u \\ t)" using con by (meson R.arr_resid_iff_con R.con_sym) moreover have "Srcs T = R.sources (u \\ t)" using tT T u src con by (metis "3" Arr_iff_Con_self Con_cons(2) Con_imp_eq_Srcs R.sources_resid Srcs_Resid Trgs.simps(2) NPath_implies_Arr list.discI R.arr_resid_iff_con) ultimately show ?thesis using 2 ind [of "u \\ t"] NPath_def by auto qed ultimately show ?thesis using tT T u src Con_cons(1) [of T "[u]" t] by simp qed thus ?thesis using tT T u src Resid_cons(1) [of T t "[u]"] Resid_rec(2) by presburger qed also have "... \ True" proof - have 2: "t \\ u \ \ \ R.arr (u \\ t)" using tT u src con NPath_def by (meson R.arr_resid_iff_con R.con_sym N.forward_stable \R.coinitial u t\ list.set_intros(1) subsetD) moreover have 3: "NPath (T \<^sup>*\\\<^sup>* [u \\ t])" using tT ind [of "u \\ t"] NPath_def by (metis Con_Arr_self Con_imp_eq_Srcs Con_rec(4) R.arr_resid_iff_con R.sources_resid Srcs.simps(2) T calculation insert_subset list.exhaust list.simps(15) Arr.simps(3)) moreover have "R.targets (t \\ u) \ Srcs (Resid T [u \\ t])" using tT T u src NPath_def by (metis "3" Arr.simps(1) R.targets_resid_sym Srcs_Resid_Arr_single con subset_refl) ultimately show ?thesis using NPath_def by (metis Arr_consI\<^sub>P N.elements_are_arr insert_subset list.simps(15)) qed finally show ?thesis by blast qed qed lemma NPath_Resid [simp]: shows "\NPath T; Arr U; Srcs T = Srcs U\ \ NPath (T \<^sup>*\\\<^sup>* U)" proof (induct T arbitrary: U) show "\U. \NPath []; Arr U; Srcs [] = Srcs U\ \ NPath ([] \<^sup>*\\\<^sup>* U)" by simp fix t T U assume ind: "\U. \NPath T; Arr U; Srcs T = Srcs U\ \ NPath (T \<^sup>*\\\<^sup>* U)" assume tT: "NPath (t # T)" assume U: "Arr U" assume Coinitial: "Srcs (t # T) = Srcs U" show "NPath ((t # T) \<^sup>*\\\<^sup>* U)" proof (cases "T = []") show "T = [] \ ?thesis" using tT U Coinitial NPath_Resid_single_Arr [of t U] NPath_def by force assume T: "T \ []" have 0: "NPath ((t # T) \<^sup>*\\\<^sup>* U) \ NPath ([t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" proof - have "U \ []" using U by auto moreover have "(t # T) \<^sup>*\\<^sup>* U" proof - have "t \ \" using tT NPath_def by auto moreover have "R.sources t = Srcs U" using Coinitial by (metis Srcs.elims U list.sel(1) Arr_has_Src) ultimately have 1: "[t] \<^sup>*\\<^sup>* U" using U NPath_Resid_single_Arr [of t U] NPath_def by auto moreover have "T \<^sup>*\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" proof - have "Srcs T = Srcs (U \<^sup>*\\\<^sup>* [t])" using tT U Coinitial 1 by (metis Con_Arr_self Con_cons(2) Con_imp_eq_Srcs Con_sym Srcs_Resid_Arr_single T list.discI NPath_implies_Arr) hence "NPath (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" using tT U Coinitial 1 Con_sym ind [of "Resid U [t]"] NPath_def by (metis Con_imp_Arr_Resid Srcs.elims T insert_subset list.simps(15) Arr.simps(3)) thus ?thesis using NPath_def by auto qed ultimately show ?thesis using Con_cons(1) [of T U t] by fastforce qed ultimately show ?thesis using tT U T Coinitial Resid_cons(1) by auto qed also have "... \ True" proof (intro iffI, simp_all) have 1: "NPath ([t] \<^sup>*\\\<^sup>* U)" by (metis Coinitial NPath_Resid_single_Arr Srcs_simp\<^sub>P U insert_subset list.sel(1) list.simps(15) NPath_def tT) moreover have 2: "NPath (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by (metis "0" Arr.simps(1) Con_cons(1) Con_imp_eq_Srcs Con_implies_Arr(1-2) NPath_def T append_Nil2 calculation ind insert_subset list.simps(15) tT) moreover have "Trgs ([t] \<^sup>*\\\<^sup>* U) \ Srcs (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by (metis Arr.simps(1) NPath_def Srcs_Resid Trgs_Resid_sym calculation(2) dual_order.refl) ultimately show "NPath ([t] \<^sup>*\\\<^sup>* U @ T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" using NPath_append [of "T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" "[t] \<^sup>*\\\<^sup>* U"] by fastforce qed finally show ?thesis by blast qed qed lemma Backward_stable_single: shows "\NPath U; NPath ([t] \<^sup>*\\\<^sup>* U)\ \ NPath [t]" proof (induct U arbitrary: t) show "\t. \NPath []; NPath ([t] \<^sup>*\\\<^sup>* [])\ \ NPath [t]" using NPath_def by simp fix t u U assume ind: "\t. \NPath U; NPath ([t] \<^sup>*\\\<^sup>* U)\ \ NPath [t]" assume uU: "NPath (u # U)" assume resid: "NPath ([t] \<^sup>*\\\<^sup>* (u # U))" show "NPath [t]" using uU ind NPath_def by (metis Arr.simps(1) Arr.simps(2) Con_implies_Arr(2) N.backward_stable N.elements_are_arr Resid_rec(1) Resid_rec(3) insert_subset list.simps(15) resid) qed lemma Backward_stable: shows "\NPath U; NPath (T \<^sup>*\\\<^sup>* U)\ \ NPath T" proof (induct T arbitrary: U) show "\U. \NPath U; NPath ([] \<^sup>*\\\<^sup>* U)\ \ NPath []" by simp fix t T U assume ind: "\U. \NPath U; NPath (T \<^sup>*\\\<^sup>* U)\ \ NPath T" assume U: "NPath U" assume resid: "NPath ((t # T) \<^sup>*\\\<^sup>* U)" show "NPath (t # T)" proof (cases "T = []") show "T = [] \ ?thesis" using U resid Backward_stable_single by blast assume T: "T \ []" have 1: "NPath ([t] \<^sup>*\\\<^sup>* U) \ NPath (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" using T U NPath_append resid NPath_def by (metis Arr.simps(1) Con_cons(1) Resid_cons(1)) have 2: "t \ \" using 1 U Backward_stable_single NPath_def by simp moreover have "NPath T" using 1 U resid ind by (metis 2 Arr.simps(2) Con_imp_eq_Srcs NPath_Resid N.elements_are_arr) moreover have "R.targets t \ Srcs T" using resid 1 Con_imp_eq_Srcs Con_sym Srcs_Resid_Arr_single NPath_def by (metis Arr.simps(1) dual_order.eq_iff) ultimately show ?thesis using NPath_def by (simp add: N.elements_are_arr) qed qed sublocale normal_sub_rts Resid \Collect NPath\ using Ide_implies_NPath NPath_implies_Arr arr_char ide_char coinitial_def sources_char\<^sub>P append_is_composite_of apply unfold_locales apply auto using Backward_stable by metis+ theorem normal_extends_to_paths: shows "normal_sub_rts Resid (Collect NPath)" .. lemma Resid_NPath_preserves_reflects_Con: assumes "NPath U" and "Srcs T = Srcs U" shows "T \<^sup>*\\\<^sup>* U \<^sup>*\\<^sup>* T' \<^sup>*\\\<^sup>* U \ T \<^sup>*\\<^sup>* T'" using assms NPath_def NPath_Resid con_char con_imp_coinitial resid_along_elem_preserves_con Con_implies_Arr(2) Con_sym Cube(1) by (metis Arr.simps(1) mem_Collect_eq) notation Cong\<^sub>0 (infix "\\<^sup>*\<^sub>0" 50) notation Cong (infix "\\<^sup>*" 50) (* * TODO: Leave these for now -- they still seem a little difficult to prove * in this context, but are probably useful. *) lemma Cong\<^sub>0_cancel_left\<^sub>C\<^sub>S: assumes "T @ U \\<^sup>*\<^sub>0 T @ U'" and "T \ []" and "U \ []" and "U' \ []" shows "U \\<^sup>*\<^sub>0 U'" using assms Cong\<^sub>0_cancel_left [of T U "T @ U" U' "T @ U'"] Cong\<^sub>0_reflexive append_is_composite_of by (metis Cong\<^sub>0_implies_Cong Cong_imp_arr(1) arr_append_imp_seq) lemma Srcs_respects_Cong: assumes "T \\<^sup>* T'" and "a \ Srcs T" and "a' \ Srcs T'" shows "[a] \\<^sup>* [a']" proof - obtain U U' where UU': "NPath U \ NPath U' \ T \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 T' \<^sup>*\\\<^sup>* U'" using assms(1) by blast show ?thesis proof show "U \ Collect NPath" using UU' by simp show "U' \ Collect NPath" using UU' by simp show "[a] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [a'] \<^sup>*\\\<^sup>* U'" proof - have "NPath ([a] \<^sup>*\\\<^sup>* U) \ NPath ([a'] \<^sup>*\\\<^sup>* U')" by (metis Arr.simps(1) Con_imp_eq_Srcs Con_implies_Arr(1) Con_single_ide_ind NPath_implies_Arr N.ide_closed R.in_sourcesE Srcs.simps(2) Srcs_simp\<^sub>P UU' assms(2-3) elements_are_arr not_arr_null null_char NPath_Resid_single_Arr) thus ?thesis using UU' by (metis Con_imp_eq_Srcs Cong\<^sub>0_imp_con NPath_Resid Srcs_Resid con_char NPath_implies_Arr mem_Collect_eq arr_resid_iff_con con_implies_arr(2)) qed qed qed lemma Trgs_respects_Cong: assumes "T \\<^sup>* T'" and "b \ Trgs T" and "b' \ Trgs T'" shows "[b] \\<^sup>* [b']" proof - have "[b] \ targets T \ [b'] \ targets T'" proof - have 1: "Ide [b] \ Ide [b']" using assms by (metis Ball_Collect Trgs_are_ide Ide.simps(2)) moreover have "Srcs [b] = Trgs T" using assms by (metis 1 Con_imp_Arr_Resid Con_imp_eq_Srcs Cong_imp_arr(1) Ide.simps(2) Srcs_Resid Con_single_ide_ind con_char arrE) moreover have "Srcs [b'] = Trgs T'" using assms by (metis Con_imp_Arr_Resid Con_imp_eq_Srcs Cong_imp_arr(2) Ide.simps(2) Srcs_Resid 1 Con_single_ide_ind con_char arrE) ultimately show ?thesis unfolding targets_char\<^sub>P using assms Cong_imp_arr(2) arr_char by blast qed thus ?thesis using assms targets_char in_targets_respects_Cong [of T T' "[b]" "[b']"] by simp qed lemma Cong\<^sub>0_append_resid_NPath: assumes "NPath (T \<^sup>*\\\<^sup>* U)" shows "Cong\<^sub>0 (T @ (U \<^sup>*\\\<^sup>* T)) U" proof (intro conjI) show 0: "(T @ U \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* U \ Collect NPath" proof - have 1: "(T @ U \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* U = T \<^sup>*\\\<^sup>* U @ (U \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* T)" by (metis Arr.simps(1) NPath_implies_Arr assms Con_append(1) Con_implies_Arr(2) Con_sym Resid_append(1) con_imp_arr_resid null_char) moreover have "NPath ..." using assms by (metis 1 Arr_append_iff\<^sub>P NPath_append NPath_implies_Arr Ide_implies_NPath Nil_is_append_conv Resid_Arr_self arr_char con_char arr_resid_iff_con self_append_conv) ultimately show ?thesis by simp qed show "U \<^sup>*\\\<^sup>* (T @ U \<^sup>*\\\<^sup>* T) \ Collect NPath" using assms 0 by (metis Arr.simps(1) Con_implies_Arr(2) Cong\<^sub>0_reflexive Resid_append(2) append.right_neutral arr_char Con_sym) qed end locale paths_in_rts_with_coherent_normal = R: rts + N: coherent_normal_sub_rts + paths_in_rts begin sublocale paths_in_rts_with_normal resid \ .. notation Cong\<^sub>0 (infix "\\<^sup>*\<^sub>0" 50) notation Cong (infix "\\<^sup>*" 50) text \ Since composites of normal transitions are assumed to exist, normal paths can be ``folded'' by composition down to single transitions. \ lemma NPath_folding: shows "NPath U \ \u. u \ \ \ R.sources u = Srcs U \ R.targets u = Trgs U \ (\t. con [t] U \ [t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t \\ u])" proof (induct U) show "NPath [] \ \u. u \ \ \ R.sources u = Srcs [] \ R.targets u = Trgs [] \ (\t. con [t] [] \ [t] \<^sup>*\\\<^sup>* [] \\<^sup>*\<^sub>0 [t \\ u])" using NPath_def by auto fix v U assume ind: "NPath U \ \u. u \ \ \ R.sources u = Srcs U \ R.targets u = Trgs U \ (\t. con [t] U \ [t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t \\ u])" assume vU: "NPath (v # U)" show "\vU. vU \ \ \ R.sources vU = Srcs (v # U) \ R.targets vU = Trgs (v # U) \ (\t. con [t] (v # U) \ [t] \<^sup>*\\\<^sup>* (v # U) \\<^sup>*\<^sub>0 [t \\ vU])" proof (cases "U = []") show "U = [] \ \vU. vU \ \ \ R.sources vU = Srcs (v # U) \ R.targets vU = Trgs (v # U) \ (\t. con [t] (v # U) \ [t] \<^sup>*\\\<^sup>* (v # U) \\<^sup>*\<^sub>0 [t \\ vU])" using vU Resid_rec(1) con_char by (metis Cong\<^sub>0_reflexive NPath_def Srcs.simps(2) Trgs.simps(2) arr_resid_iff_con insert_subset list.simps(15)) assume "U \ []" hence U: "NPath U" using vU by (metis NPath_append append_Cons append_Nil) obtain u where u: "u \ \ \ R.sources u = Srcs U \ R.targets u = Trgs U \ (\t. con [t] U \ [t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t \\ u])" using U ind by blast have seq: "R.seq v u" proof show "R.arr v" using vU by (metis Con_Arr_self Con_rec(4) NPath_implies_Arr \U \ []\ R.arrI) show "R.arr u" by (simp add: N.elements_are_arr u) show "R.targets v = R.sources u" by (metis (full_types) NPath_implies_Arr R.sources_resid Srcs.simps(2) \U \ []\ Con_Arr_self Con_imp_eq_Srcs Con_initial_right Con_rec(2) u vU) qed obtain vu where vu: "R.composite_of v u vu" using N.composite_closed_right seq u by presburger have "vu \ \ \ R.sources vu = Srcs (v # U) \ R.targets vu = Trgs (v # U) \ (\t. con [t] (v # U) \ [t] \<^sup>*\\\<^sup>* (v # U) \\<^sup>*\<^sub>0 [t \\ vu])" proof (intro conjI allI) show "vu \ \" by (meson NPath_def N.composite_closed list.set_intros(1) subsetD u vU vu) show "R.sources vu = Srcs (v # U)" by (metis Con_imp_eq_Srcs Con_initial_right NPath_implies_Arr R.sources_composite_of Srcs.simps(2) Arr_iff_Con_self vU vu) show "R.targets vu = Trgs (v # U)" by (metis R.targets_composite_of Trgs.simps(3) \U \ []\ list.exhaust_sel u vu) fix t show "con [t] (v # U) \ [t] \<^sup>*\\\<^sup>* (v # U) \\<^sup>*\<^sub>0 [t \\ vu]" proof (intro impI) assume t: "con [t] (v # U)" have 1: "[t] \<^sup>*\\\<^sup>* (v # U) = [t \\ v] \<^sup>*\\\<^sup>* U" using t Resid_rec(3) \U \ []\ con_char by force also have "... \\<^sup>*\<^sub>0 [(t \\ v) \\ u]" using 1 t u by force also have "[(t \\ v) \\ u] \\<^sup>*\<^sub>0 [t \\ vu]" proof - have "(t \\ v) \\ u \ t \\ vu" using vu R.resid_composite_of by (metis (no_types, lifting) N.Cong\<^sub>0_composite_of_arr_normal N.Cong\<^sub>0_subst_right(1) \U \ []\ Con_rec(3) con_char R.con_sym t u) thus ?thesis using Ide.simps(2) R.prfx_implies_con Resid.simps(3) ide_char ide_closed by presburger qed finally show "[t] \<^sup>*\\\<^sup>* (v # U) \\<^sup>*\<^sub>0 [t \\ vu]" by blast qed qed thus ?thesis by blast qed qed text \ Coherence for single transitions extends inductively to paths. \ lemma Coherent_single: assumes "R.arr t" and "NPath U" and "NPath U'" and "R.sources t = Srcs U" and "Srcs U = Srcs U'" and "Trgs U = Trgs U'" shows "[t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t] \<^sup>*\\\<^sup>* U'" proof - have 1: "con [t] U \ con [t] U'" using assms by (metis Arr.simps(1-2) Arr_iff_Con_self Resid_NPath_preserves_reflects_Con Srcs.simps(2) con_char) obtain u where u: "u \ \ \ R.sources u = Srcs U \ R.targets u = Trgs U \ (\t. con [t] U \ [t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t \\ u])" using assms NPath_folding by metis obtain u' where u': "u' \ \ \ R.sources u' = Srcs U' \ R.targets u' = Trgs U' \ (\t. con [t] U' \ [t] \<^sup>*\\\<^sup>* U' \\<^sup>*\<^sub>0 [t \\ u'])" using assms NPath_folding by metis have "[t] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [t \\ u]" using u 1 by blast also have "[t \\ u] \\<^sup>*\<^sub>0 [t \\ u']" using assms(1,4-6) N.Cong\<^sub>0_imp_con N.coherent u u' NPath_def by simp also have "[t \\ u'] \\<^sup>*\<^sub>0 [t] \<^sup>*\\\<^sup>* U'" using u' 1 by simp finally show ?thesis by simp qed lemma Coherent: shows "\ Arr T; NPath U; NPath U'; Srcs T = Srcs U; Srcs U = Srcs U'; Trgs U = Trgs U' \ \ T \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 T \<^sup>*\\\<^sup>* U'" proof (induct T arbitrary: U U') show "\U U'. \ Arr []; NPath U; NPath U'; Srcs [] = Srcs U; Srcs U = Srcs U'; Trgs U = Trgs U' \ \ [] \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 [] \<^sup>*\\\<^sup>* U'" by (simp add: arr_char) fix t T U U' assume tT: "Arr (t # T)" and U: "NPath U" and U': "NPath U'" and Srcs1: "Srcs (t # T) = Srcs U" and Srcs2: "Srcs U = Srcs U'" and Trgs: "Trgs U = Trgs U'" and ind: "\U U'. \ Arr T; NPath U; NPath U'; Srcs T = Srcs U; Srcs U = Srcs U'; Trgs U = Trgs U' \ \ T \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 T \<^sup>*\\\<^sup>* U'" have t: "R.arr t" using tT by (metis Arr.simps(2) Con_Arr_self Con_rec(4) R.arrI) show "(t # T) \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 (t # T) \<^sup>*\\\<^sup>* U'" proof (cases "T = []") show "T = [] \ ?thesis" by (metis Srcs.simps(2) Srcs1 Srcs2 Trgs U U' Coherent_single Arr.simps(2) tT) assume T: "T \ []" let ?t = "[t] \<^sup>*\\\<^sup>* U" and ?t' = "[t] \<^sup>*\\\<^sup>* U'" let ?T = "T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])" let ?T' = "T \<^sup>*\\\<^sup>* (U' \<^sup>*\\\<^sup>* [t])" have 0: "(t # T) \<^sup>*\\\<^sup>* U = ?t @ ?T \ (t # T) \<^sup>*\\\<^sup>* U' = ?t' @ ?T'" using tT U U' Srcs1 Srcs2 by (metis Arr_has_Src Arr_iff_Con_self Resid_cons(1) Srcs.simps(1) Resid_NPath_preserves_reflects_Con) have 1: "?t \\<^sup>*\<^sub>0 ?t'" by (metis Srcs1 Srcs2 Srcs_simp\<^sub>P Trgs U U' list.sel(1) Coherent_single t tT) have A: "?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t) = T \<^sup>*\\\<^sup>* ((U \<^sup>*\\\<^sup>* [t]) @ (?t' \<^sup>*\\\<^sup>* ?t))" using 1 Arr.simps(1) Con_append(2) Con_sym Resid_append(2) Con_implies_Arr(1) NPath_def by (metis arr_char elements_are_arr) have B: "?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t') = T \<^sup>*\\\<^sup>* ((U' \<^sup>*\\\<^sup>* [t]) @ (?t \<^sup>*\\\<^sup>* ?t'))" by (metis "1" Con_appendI(2) Con_sym Resid.simps(1) Resid_append(2) elements_are_arr not_arr_null null_char) have E: "?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t) \\<^sup>*\<^sub>0 ?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')" proof - have "Arr T" using Arr.elims(1) T tT by blast moreover have "NPath (U \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U))" using 1 U t tT Srcs1 Srcs_simp\<^sub>P apply (intro NPath_appendI) apply auto by (metis Arr.simps(1) NPath_def Srcs_Resid Trgs_Resid_sym) moreover have "NPath (U' \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U'))" using t U' 1 Con_imp_eq_Srcs Trgs_Resid_sym apply (intro NPath_appendI) apply auto apply (metis Arr.simps(2) NPath_Resid Resid.simps(1)) by (metis Arr.simps(1) NPath_def Srcs_Resid) moreover have "Srcs T = Srcs (U \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U))" using A B by (metis (full_types) "0" "1" Arr_has_Src Con_cons(1) Con_implies_Arr(1) Srcs.simps(1) Srcs_append T elements_are_arr not_arr_null null_char Con_imp_eq_Srcs) moreover have "Srcs (U \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U)) = Srcs (U' \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U'))" by (metis "1" Con_implies_Arr(2) Con_sym Cong\<^sub>0_imp_con Srcs_Resid Srcs_append arr_char con_char arr_resid_iff_con) moreover have "Trgs (U \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U)) = Trgs (U' \<^sup>*\\\<^sup>* [t] @ ([t] \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ([t] \<^sup>*\\\<^sup>* U'))" using "1" Cong\<^sub>0_imp_con con_char by force ultimately show ?thesis using A B ind [of "(U \<^sup>*\\\<^sup>* [t]) @ (?t' \<^sup>*\\\<^sup>* ?t)" "(U' \<^sup>*\\\<^sup>* [t]) @ (?t \<^sup>*\\\<^sup>* ?t')"] by simp qed have C: "NPath ((?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)) \<^sup>*\\\<^sup>* (?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')))" using E by blast have D: "NPath ((?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')) \<^sup>*\\\<^sup>* (?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)))" using E by blast show ?thesis proof have 2: "((t # T) \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* U') = ((?t \<^sup>*\\\<^sup>* ?t') \<^sup>*\\\<^sup>* ?T') @ ((?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)) \<^sup>*\\\<^sup>* (?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')))" proof - have "((t # T) \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* U') = (?t @ ?T) \<^sup>*\\\<^sup>* (?t' @ ?T')" using 0 by fastforce also have "... = ((?t @ ?T) \<^sup>*\\\<^sup>* ?t') \<^sup>*\\\<^sup>* ?T'" using tT T U U' Srcs1 Srcs2 0 by (metis Con_appendI(2) Con_cons(1) Con_sym Resid.simps(1) Resid_append(2)) also have "... = ((?t \<^sup>*\\\<^sup>* ?t') @ (?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t))) \<^sup>*\\\<^sup>* ?T'" by (metis (no_types, lifting) Arr.simps(1) Con_appendI(1) Con_implies_Arr(1) D NPath_def Resid_append(1) null_is_zero(2)) also have "... = ((?t \<^sup>*\\\<^sup>* ?t') \<^sup>*\\\<^sup>* ?T') @ ((?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)) \<^sup>*\\\<^sup>* (?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')))" proof - have "?t \<^sup>*\\\<^sup>* ?t' @ ?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t) \<^sup>*\\<^sup>* ?T'" using C D E Con_sym by (metis Con_append(2) Cong\<^sub>0_imp_con con_char arr_resid_iff_con con_implies_arr(2)) thus ?thesis using Resid_append(1) by (metis Con_sym append.right_neutral Resid.simps(1)) qed finally show ?thesis by simp qed moreover have 3: "NPath ..." proof - have "NPath ((?t \<^sup>*\\\<^sup>* ?t') \<^sup>*\\\<^sup>* ?T')" using 0 1 E by (metis Con_imp_Arr_Resid Con_imp_eq_Srcs NPath_Resid Resid.simps(1) ex_un_null mem_Collect_eq) moreover have "Trgs ((?t \<^sup>*\\\<^sup>* ?t') \<^sup>*\\\<^sup>* ?T') = Srcs ((?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)) \<^sup>*\\\<^sup>* (?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')))" using C by (metis NPath_implies_Arr Srcs.simps(1) Srcs_Resid Trgs_Resid_sym Arr_has_Src) ultimately show ?thesis using C by blast qed ultimately show "((t # T) \<^sup>*\\\<^sup>* U) \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* U') \ Collect NPath" by simp have 4: "((t # T) \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* U) = ((?t' \<^sup>*\\\<^sup>* ?t) \<^sup>*\\\<^sup>* ?T) @ ((?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')) \<^sup>*\\\<^sup>* (?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)))" by (metis "0" "2" "3" Arr.simps(1) Con_implies_Arr(1) Con_sym D NPath_def Resid_append2) moreover have "NPath ..." proof - have "NPath ((?t' \<^sup>*\\\<^sup>* ?t) \<^sup>*\\\<^sup>* ?T)" by (metis "1" CollectD Cong\<^sub>0_imp_con E con_imp_coinitial forward_stable arr_resid_iff_con con_implies_arr(2)) moreover have "NPath ((?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')) \<^sup>*\\\<^sup>* (?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)))" using U U' 1 D ind Coherent_single [of t U' U] by blast moreover have "Trgs ((?t' \<^sup>*\\\<^sup>* ?t) \<^sup>*\\\<^sup>* ?T) = Srcs ((?T' \<^sup>*\\\<^sup>* (?t \<^sup>*\\\<^sup>* ?t')) \<^sup>*\\\<^sup>* (?T \<^sup>*\\\<^sup>* (?t' \<^sup>*\\\<^sup>* ?t)))" by (metis Arr.simps(1) NPath_def Srcs_Resid Trgs_Resid_sym calculation(2)) ultimately show ?thesis by blast qed ultimately show "((t # T) \<^sup>*\\\<^sup>* U') \<^sup>*\\\<^sup>* ((t # T) \<^sup>*\\\<^sup>* U) \ Collect NPath" by simp qed qed qed sublocale rts_with_composites Resid using is_rts_with_composites by simp sublocale coherent_normal_sub_rts Resid \Collect NPath\ proof fix T U U' assume T: "arr T" and U: "U \ Collect NPath" and U': "U' \ Collect NPath" assume sources_UU': "sources U = sources U'" and targets_UU': "targets U = targets U'" and TU: "sources T = sources U" have "Srcs T = Srcs U" using TU sources_char\<^sub>P T arr_iff_has_source by auto moreover have "Srcs U = Srcs U'" by (metis Con_imp_eq_Srcs T TU con_char con_imp_coinitial_ax con_sym in_sourcesE in_sourcesI arr_def sources_UU') moreover have "Trgs U = Trgs U'" using U U' targets_UU' targets_char by (metis (full_types) arr_iff_has_target composable_def composable_iff_seq composite_of_arr_target elements_are_arr equals0I seq_char) ultimately show "T \<^sup>*\\\<^sup>* U \\<^sup>*\<^sub>0 T \<^sup>*\\\<^sup>* U'" using T U U' Coherent [of T U U'] arr_char by blast qed theorem coherent_normal_extends_to_paths: shows "coherent_normal_sub_rts Resid (Collect NPath)" .. lemma Cong\<^sub>0_append_Arr_NPath: assumes "T \ []" and "Arr (T @ U)" and "NPath U" shows "Cong\<^sub>0 (T @ U) T" using assms by (metis Arr.simps(1) Arr_appendE\<^sub>P NPath_implies_Arr append_is_composite_of arrI\<^sub>P arr_append_imp_seq composite_of_arr_normal mem_Collect_eq) lemma Cong_append_NPath_Arr: assumes "T \ []" and "Arr (U @ T)" and "NPath U" shows "U @ T \\<^sup>* T" using assms by (metis (full_types) Arr.simps(1) Con_Arr_self Con_append(2) Con_implies_Arr(2) Con_imp_eq_Srcs composite_of_normal_arr Srcs_Resid append_is_composite_of arr_char NPath_implies_Arr mem_Collect_eq seq_char) subsubsection "Permutation Congruence" text \ Here we show that \\<^sup>*\\<^sup>*\ coincides with ``permutation congruence'': the least congruence respecting composition that relates \[t, u \ t]\ and \[u, t \ u]\ whenever \t \ u\ and that relates \T @ [b]\ and \T\ whenever \b\ is an identity such that \seq T [b]\. \ inductive PCong where "Arr T \ PCong T T" | "PCong T U \ PCong U T" | "\PCong T U; PCong U V\ \ PCong T V" | "\seq T U; PCong T T'; PCong U U'\ \ PCong (T @ U) (T' @ U')" | "\seq T [b]; R.ide b\ \ PCong (T @ [b]) T" | "t \ u \ PCong [t, u \\ t] [u, t \\ u]" lemmas PCong.intros(3) [trans] lemma PCong_append_Ide: shows "\seq T B; Ide B\ \ PCong (T @ B) T" proof (induct B) show "\seq T []; Ide []\ \ PCong (T @ []) T" by auto fix b B T assume ind: "\seq T B; Ide B\ \ PCong (T @ B) T" assume seq: "seq T (b # B)" assume Ide: "Ide (b # B)" have "T @ (b # B) = (T @ [b]) @ B" by simp also have "PCong ... (T @ B)" apply (cases "B = []") using Ide PCong.intros(5) seq apply force using seq Ide PCong.intros(4) [of "T @ [b]" B T B] by (metis Arr.simps(1) Ide_imp_Ide_hd PCong.intros(1) PCong.intros(5) append_is_Nil_conv arr_append arr_append_imp_seq arr_char calculation list.distinct(1) list.sel(1) seq_char) also have "PCong (T @ B) T" proof (cases "B = []") show "B = [] \ ?thesis" using PCong.intros(1) seq seq_char by force assume B: "B \ []" have "seq T B" using B seq Ide by (metis Con_imp_eq_Srcs Ide_imp_Ide_hd Trgs_append \T @ b # B = (T @ [b]) @ B\ append_is_Nil_conv arr_append arr_append_imp_seq arr_char cong_cons_ideI(2) list.distinct(1) list.sel(1) not_arr_null null_char seq_char ide_implies_arr) thus ?thesis using seq Ide ind by (metis Arr.simps(1) Ide.elims(3) Ide.simps(3) seq_char) qed finally show "PCong (T @ (b # B)) T" by blast qed lemma PCong_imp_Cong: shows "PCong T U \ T \<^sup>*\\<^sup>* U" proof (induct rule: PCong.induct) show "\T. Arr T \ T \<^sup>*\\<^sup>* T" using cong_reflexive by blast show "\T U. \PCong T U; T \<^sup>*\\<^sup>* U\ \ U \<^sup>*\\<^sup>* T" by blast show "\T U V. \PCong T U; T \<^sup>*\\<^sup>* U; PCong U V; U \<^sup>*\\<^sup>* V\ \ T \<^sup>*\\<^sup>* V" using cong_transitive by blast show "\T U U' T'. \seq T U; PCong T T'; T \<^sup>*\\<^sup>* T'; PCong U U'; U \<^sup>*\\<^sup>* U'\ \ T @ U \<^sup>*\\<^sup>* T' @ U'" using cong_append by simp show "\T b. \seq T [b]; R.ide b\ \ T @ [b] \<^sup>*\\<^sup>* T" using cong_append_ideI(4) ide_char by force show "\t u. t \ u \ [t, u \\ t] \<^sup>*\\<^sup>* [u, t \\ u]" proof - have "\t u. t \ u \ [t, u \\ t] \<^sup>*\\<^sup>* [u, t \\ u]" proof - fix t u assume con: "t \ u" have "([t] @ [u \\ t]) \<^sup>*\\\<^sup>* ([u] @ [t \\ u]) = [(t \\ u) \\ (t \\ u), ((u \\ t) \\ (u \\ t)) \\ ((t \\ u) \\ (t \\ u))]" using con Resid_append2 [of "[t]" "[u \\ t]" "[u]" "[t \\ u]"] apply simp by (metis R.arr_resid_iff_con R.con_target R.conE R.con_sym R.prfx_implies_con R.prfx_reflexive R.cube) moreover have "Ide ..." using con by (metis Arr.simps(2) Arr.simps(3) Ide.simps(2) Ide.simps(3) R.arr_resid_iff_con R.con_sym R.resid_ide_arr R.prfx_reflexive calculation Con_imp_Arr_Resid) ultimately show"[t, u \\ t] \<^sup>*\\<^sup>* [u, t \\ u]" using ide_char by auto qed thus "\t u. t \ u \ [t, u \\ t] \<^sup>*\\<^sup>* [u, t \\ u]" using R.con_sym by blast qed qed lemma PCong_permute_single: shows "[t] \<^sup>*\\<^sup>* U \ PCong ([t] @ (U \<^sup>*\\\<^sup>* [t])) (U @ ([t] \<^sup>*\\\<^sup>* U))" proof (induct U arbitrary: t) show "\t. [t] \<^sup>*\\<^sup>* [] \ PCong ([t] @ [] \<^sup>*\\\<^sup>* [t]) ([] @ [t] \<^sup>*\\\<^sup>* [])" by auto fix t u U assume ind: "\t. [t] \<^sup>*\\\<^sup>* U \ [] \ PCong ([t] @( U \<^sup>*\\\<^sup>* [t])) (U @ ([t] \<^sup>*\\\<^sup>* U))" assume con: "[t] \<^sup>*\\<^sup>* u # U" show "PCong ([t] @ (u # U) \<^sup>*\\\<^sup>* [t]) ((u # U) @ [t] \<^sup>*\\\<^sup>* (u # U))" proof (cases "U = []") show "U = [] \ ?thesis" by (metis PCong.intros(6) Resid.simps(3) append_Cons append_eq_append_conv2 append_self_conv con_char con_def con con_sym_ax) assume U: "U \ []" show "PCong ([t] @ ((u # U) \<^sup>*\\\<^sup>* [t])) ((u # U) @ ([t] \<^sup>*\\\<^sup>* (u # U)))" proof - have "[t] @ ((u # U) \<^sup>*\\\<^sup>* [t]) = [t] @ ([u \\ t] @ (U \<^sup>*\\\<^sup>* [t \\ u]))" using Con_sym Resid_rec(2) U con by auto also have "... = ([t] @ [u \\ t]) @ (U \<^sup>*\\\<^sup>* [t \\ u])" by auto also have "PCong ... (([u] @ [t \\ u]) @ (U \<^sup>*\\\<^sup>* [t \\ u]))" proof - have "PCong ([t] @ [u \\ t]) ([u] @ [t \\ u])" using con by (simp add: Con_rec(3) PCong.intros(6) U) thus ?thesis by (metis Arr_Resid_single Con_implies_Arr(1) Con_rec(2) Con_sym PCong.intros(1,4) Srcs_Resid U append_is_Nil_conv append_is_composite_of arr_append_imp_seq arr_char calculation composite_of_unq_upto_cong con not_arr_null null_char ide_implies_arr seq_char) qed also have "([u] @ [t \\ u]) @ (U \<^sup>*\\\<^sup>* [t \\ u]) = [u] @ ([t \\ u] @ (U \<^sup>*\\\<^sup>* [t \\ u]))" by simp also have "PCong ... ([u] @ (U @ ([t \\ u] \<^sup>*\\\<^sup>* U)))" proof - have "PCong ([t \\ u] @ (U \<^sup>*\\\<^sup>* [t \\ u])) (U @ ([t \\ u] \<^sup>*\\\<^sup>* U))" using ind by (metis Resid_rec(3) U con) moreover have "seq [u] ([t \\ u] @ U \<^sup>*\\\<^sup>* [t \\ u])" proof show "Arr [u]" using Con_implies_Arr(2) Con_initial_right con by blast show "Arr ([t \\ u] @ U \<^sup>*\\\<^sup>* [t \\ u])" using Con_implies_Arr(1) U con Con_imp_Arr_Resid Con_rec(3) Con_sym by fastforce show "Trgs [u] \ Srcs ([t \\ u] @ U \<^sup>*\\\<^sup>* [t \\ u]) \ {}" by (metis Arr.simps(1) Arr.simps(2) Arr_has_Trg Con_implies_Arr(1) Int_absorb R.arr_resid_iff_con R.sources_resid Resid_rec(3) Srcs.simps(2) Srcs_append Trgs.simps(2) U \Arr [u]\ con) qed moreover have "PCong [u] [u]" using PCong.intros(1) calculation(2) seq_char by force ultimately show ?thesis using U arr_append arr_char con seq_char PCong.intros(4) [of "[u]" "[t \\ u] @ (U \<^sup>*\\\<^sup>* [t \\ u])" "[u]" "U @ ([t \\ u] \<^sup>*\\\<^sup>* U)"] by blast qed also have "([u] @ (U @ ([t \\ u] \<^sup>*\\\<^sup>* U))) = ((u # U) @ [t] \<^sup>*\\\<^sup>* (u # U))" by (metis Resid_rec(3) U append_Cons append_Nil con) finally show ?thesis by blast qed qed qed lemma PCong_permute: shows "T \<^sup>*\\<^sup>* U \ PCong (T @ (U \<^sup>*\\\<^sup>* T)) (U @ (T \<^sup>*\\\<^sup>* U))" proof (induct T arbitrary: U) show "\U. [] \<^sup>*\\\<^sup>* U \ [] \ PCong ([] @ U \<^sup>*\\\<^sup>* []) (U @ [] \<^sup>*\\\<^sup>* U)" by simp fix t T U assume ind: "\U. T \<^sup>*\\<^sup>* U \ PCong (T @ (U \<^sup>*\\\<^sup>* T)) (U @ (T \<^sup>*\\\<^sup>* U))" assume con: "t # T \<^sup>*\\<^sup>* U" show "PCong ((t # T) @ (U \<^sup>*\\\<^sup>* (t # T))) (U @ ((t # T) \<^sup>*\\\<^sup>* U))" proof (cases "T = []") assume T: "T = []" have "(t # T) @ (U \<^sup>*\\\<^sup>* (t # T)) = [t] @ (U \<^sup>*\\\<^sup>* [t])" using con T by simp also have "PCong ... (U @ ([t] \<^sup>*\\\<^sup>* U))" using PCong_permute_single T con by blast finally show ?thesis using T by fastforce next assume T: "T \ []" have "(t # T) @ (U \<^sup>*\\\<^sup>* (t # T)) = [t] @ (T @ (U \<^sup>*\\\<^sup>* (t # T)))" by simp also have "PCong ... ([t] @ (U \<^sup>*\\\<^sup>* [t]) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])))" using ind [of "U \<^sup>*\\\<^sup>* [t]"] by (metis Arr.simps(1) Con_imp_Arr_Resid Con_implies_Arr(2) Con_sym PCong.intros(1,4) Resid_cons(2) Srcs_Resid T arr_append arr_append_imp_seq calculation con not_arr_null null_char seq_char) also have "[t] @ (U \<^sup>*\\\<^sup>* [t]) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) = ([t] @ (U \<^sup>*\\\<^sup>* [t])) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))" by simp also have "PCong (([t] @ (U \<^sup>*\\\<^sup>* [t])) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t]))) ((U @ ([t] \<^sup>*\\\<^sup>* U)) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])))" by (metis Arr.simps(1) Con_cons(1) Con_imp_Arr_Resid Con_implies_Arr(2) PCong.intros(1,4) PCong_permute_single Srcs_Resid T Trgs_append arr_append arr_char con seq_char) also have "(U @ ([t] \<^sup>*\\\<^sup>* U)) @ (T \<^sup>*\\\<^sup>* (U \<^sup>*\\\<^sup>* [t])) = U @ ((t # T) \<^sup>*\\\<^sup>* U)" by (metis Resid.simps(2) Resid_cons(1) append.assoc con) finally show ?thesis by blast qed qed lemma Cong_imp_PCong: assumes "T \<^sup>*\\<^sup>* U" shows "PCong T U" proof - have "PCong T (T @ (U \<^sup>*\\\<^sup>* T))" using assms PCong.intros(2) PCong_append_Ide by (metis Con_implies_Arr(1) Ide.simps(1) Srcs_Resid ide_char Con_imp_Arr_Resid seq_char) also have "PCong (T @ (U \<^sup>*\\\<^sup>* T)) (U @ (T \<^sup>*\\\<^sup>* U))" using PCong_permute assms con_char prfx_implies_con by presburger also have "PCong (U @ (T \<^sup>*\\\<^sup>* U)) U" using assms PCong_append_Ide by (metis Con_imp_Arr_Resid Con_implies_Arr(1) Srcs_Resid arr_resid_iff_con ide_implies_arr con_char ide_char seq_char) finally show ?thesis by blast qed lemma Cong_iff_PCong: shows "T \<^sup>*\\<^sup>* U \ PCong T U" using PCong_imp_Cong Cong_imp_PCong by blast end section "Composite Completion" text \ The RTS of paths in an RTS factors via the coherent normal sub-RTS of identity paths into an extensional RTS with composites, which can be regarded as a ``composite completion'' of the original RTS. \ locale composite_completion = R: rts begin interpretation N: coherent_normal_sub_rts resid \Collect R.ide\ using R.rts_axioms R.identities_form_coherent_normal_sub_rts by auto sublocale P: paths_in_rts_with_coherent_normal resid \Collect R.ide\ .. sublocale quotient_by_coherent_normal P.Resid \Collect P.NPath\ .. notation P.Resid (infix "\<^sup>*\\\<^sup>*" 70) notation P.Con (infix "\<^sup>*\\<^sup>*" 50) notation P.Cong (infix "\<^sup>*\\<^sup>*" 50) notation P.Cong\<^sub>0 (infix "\<^sup>*\\<^sub>0\<^sup>*" 50) notation P.Cong_class ("\_\") notation Resid (infix "\\<^sup>*\\\<^sup>*\" 70) notation con (infix "\\<^sup>*\\<^sup>*\" 50) notation prfx (infix "\\<^sup>*\\<^sup>*\" 50) lemma NPath_char: shows "P.NPath T \ P.Ide T" using P.NPath_def P.Ide_implies_NPath by blast lemma Cong_eq_Cong\<^sub>0: shows "T \<^sup>*\\<^sup>* T' \ T \<^sup>*\\<^sub>0\<^sup>* T'" by (metis P.Cong_iff_cong P.ide_char P.ide_closed CollectD Collect_cong NPath_char) lemma Srcs_respects_Cong: assumes "T \<^sup>*\\<^sup>* T'" shows "P.Srcs T = P.Srcs T'" using assms by (meson P.Con_imp_eq_Srcs P.Cong\<^sub>0_imp_con P.con_char Cong_eq_Cong\<^sub>0) lemma sources_respects_Cong: assumes "T \<^sup>*\\<^sup>* T'" shows "P.sources T = P.sources T'" using assms by (meson P.Cong\<^sub>0_imp_coinitial Cong_eq_Cong\<^sub>0) lemma Trgs_respects_Cong: assumes "T \<^sup>*\\<^sup>* T'" shows "P.Trgs T = P.Trgs T'" proof - have "P.Trgs T = P.Trgs (T @ (T' \<^sup>*\\\<^sup>* T))" using assms NPath_char P.Arr.simps(1) P.Con_imp_Arr_Resid P.Con_sym P.Cong_def P.Con_Arr_self P.Con_implies_Arr(2) P.Resid_Ide(1) P.Srcs_Resid P.Trgs_append by (metis P.Cong\<^sub>0_imp_con P.con_char CollectD) also have "... = P.Trgs (T' @ (T \<^sup>*\\\<^sup>* T'))" using P.Cong\<^sub>0_imp_con P.con_char Cong_eq_Cong\<^sub>0 assms by force also have "... = P.Trgs T'" using assms NPath_char P.Arr.simps(1) P.Con_imp_Arr_Resid P.Con_sym P.Cong_def P.Con_Arr_self P.Con_implies_Arr(2) P.Resid_Ide(1) P.Srcs_Resid P.Trgs_append by (metis P.Cong\<^sub>0_imp_con P.con_char CollectD) finally show ?thesis by blast qed lemma targets_respects_Cong: assumes "T \<^sup>*\\<^sup>* T'" shows "P.targets T = P.targets T'" using assms P.Cong_imp_arr(1) P.Cong_imp_arr(2) P.arr_iff_has_target P.targets_char\<^sub>P Trgs_respects_Cong by force lemma ide_char\<^sub>C\<^sub>C: shows "ide \ \ arr \ \ (\T. T \ \ \ P.Ide T)" using NPath_char ide_char' by force lemma con_char\<^sub>C\<^sub>C: shows "\ \\<^sup>*\\<^sup>*\ \ \ arr \ \ arr \ \ P.Cong_class_rep \ \<^sup>*\\<^sup>* P.Cong_class_rep \" proof show "arr \ \ arr \ \ P.Cong_class_rep \ \<^sup>*\\<^sup>* P.Cong_class_rep \ \ \ \\<^sup>*\\<^sup>*\ \" using arr_char P.con_char by (meson P.rep_in_Cong_class con_char\<^sub>Q\<^sub>C\<^sub>N) show "\ \\<^sup>*\\<^sup>*\ \ \ arr \ \ arr \ \ P.Cong_class_rep \ \<^sup>*\\<^sup>* P.Cong_class_rep \" proof - assume con: "\ \\<^sup>*\\<^sup>*\ \" have 1: "arr \ \ arr \" using con coinitial_iff con_imp_coinitial by blast moreover have "P.Cong_class_rep \ \<^sup>*\\<^sup>* P.Cong_class_rep \" proof - obtain T U where TU: "T \ \ \ U \ \ \ P.Con T U" using con Resid_def by (meson P.con_char con_char\<^sub>Q\<^sub>C\<^sub>N) have "T \<^sup>*\\<^sup>* P.Cong_class_rep \ \ U \<^sup>*\\<^sup>* P.Cong_class_rep \" using TU 1 by (meson P.Cong_class_memb_Cong_rep arr_char) thus ?thesis using TU P.Cong_subst(1) [of T "P.Cong_class_rep \" U "P.Cong_class_rep \"] by (metis P.coinitial_iff P.con_char P.con_imp_coinitial sources_respects_Cong) qed ultimately show ?thesis by simp qed qed lemma con_char\<^sub>C\<^sub>C': shows "\ \\<^sup>*\\<^sup>*\ \ \ arr \ \ arr \ \ (\T U. T \ \ \ U \ \ \ T \<^sup>*\\<^sup>* U)" proof show "arr \ \ arr \ \ (\T U. T \ \ \ U \ \ \ T \<^sup>*\\<^sup>* U) \ \ \\<^sup>*\\<^sup>*\ \" using con_char\<^sub>C\<^sub>C by (simp add: P.rep_in_Cong_class arr_char) show "\ \\<^sup>*\\<^sup>*\ \ \ arr \ \ arr \ \ (\T U. T \ \ \ U \ \ \ T \<^sup>*\\<^sup>* U)" proof (intro conjI allI impI) assume 1: "\ \\<^sup>*\\<^sup>*\ \" show "arr \" using 1 con_implies_arr by simp show "arr \" using 1 con_implies_arr by simp fix T U assume 2: "T \ \ \ U \ \" show "T \<^sup>*\\<^sup>* U" using 1 2 P.Cong_class_memb_Cong_rep by (meson P.Cong\<^sub>0_subst_Con P.con_char Cong_eq_Cong\<^sub>0 arr_char con_char\<^sub>C\<^sub>C) qed qed lemma resid_char: shows "\ \\<^sup>*\\\<^sup>*\ \ = (if \ \\<^sup>*\\<^sup>*\ \ then \P.Cong_class_rep \ \<^sup>*\\\<^sup>* P.Cong_class_rep \\ else {})" by (metis P.con_char P.rep_in_Cong_class Resid_by_members arr_char arr_resid_iff_con con_char\<^sub>C\<^sub>C is_Cong_class_Resid) lemma src_char': shows "src \ = {A. arr \ \ P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A}" proof (cases "arr \") show "\ arr \ \ ?thesis" by (simp add: null_char src_def) assume \: "arr \" have 1: "\A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A" by (metis P.Arr.simps(1) P.Con_imp_eq_Srcs P.Cong\<^sub>0_imp_con P.Cong_class_memb_Cong_rep P.Cong_def P.con_char P.rep_in_Cong_class CollectD \ NPath_char P.Con_implies_Arr(1) arr_char) let ?A = "SOME A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A" have A: "P.Ide ?A \ P.Srcs (P.Cong_class_rep \) = P.Srcs ?A" using 1 someI_ex [of "\A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A"] by simp have a: "arr \?A\" using A P.ide_char P.is_Cong_classI arr_char by blast have ide_a: "ide \?A\" using a A P.Cong_class_def P.normal_is_Cong_closed NPath_char ide_char\<^sub>C\<^sub>C by auto have "sources \ = {\?A\}" proof - have "\ \\<^sup>*\\<^sup>*\ \?A\" by (metis (no_types, lifting) A P.Con_Ide_iff P.Cong_class_memb_Cong_rep P.Cong_imp_arr(1) P.arr_char P.arr_in_Cong_class P.ide_char P.ide_implies_arr P.rep_in_Cong_class Con_char a \ P.con_char null_char arr_char P.con_sym conI) hence "\?A\ \ sources \" using ide_a in_sourcesI by simp thus ?thesis using sources_char by auto qed moreover have "\?A\ = {A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A}" proof show "{A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A} \ \?A\" using A P.Cong_class_def P.Cong_closure_props(3) P.Ide_implies_Arr P.ide_closed P.ide_char by fastforce show "\?A\ \ {A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A}" using a A P.Cong_class_def Srcs_respects_Cong ide_a ide_char\<^sub>C\<^sub>C by blast qed ultimately show ?thesis using \ src_in_sources by force qed lemma src_char: shows "src \ = {A. arr \ \ P.Ide A \ (\T. T \ \ \ P.Srcs T = P.Srcs A)}" proof (cases "arr \") show "\ arr \ \ ?thesis" by (simp add: null_char src_def) assume \: "arr \" have "\T. T \ \ \ P.Srcs T = P.Srcs (P.Cong_class_rep \)" using \ P.Cong_class_memb_Cong_rep Srcs_respects_Cong arr_char by auto thus ?thesis using \ src_char' P.is_Cong_class_def arr_char by force qed lemma trg_char': shows "trg \ = {B. arr \ \ P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B}" proof (cases "arr \") show "\ arr \ \ ?thesis" by (metis (no_types, lifting) Collect_empty_eq arrI resid_arr_self resid_char) assume \: "arr \" have 1: "\B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B" by (metis P.Con_implies_Arr(2) P.Resid_Arr_self P.Srcs_Resid \ con_char\<^sub>C\<^sub>C arrE) define B where "B = (SOME B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B)" have B: "P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B" unfolding B_def using 1 someI_ex [of "\B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B"] by simp hence 2: "P.Ide B \ P.Con (P.Resid (P.Cong_class_rep \) (P.Cong_class_rep \)) B" using \ by (metis (no_types, lifting) P.Con_Ide_iff P.Ide_implies_Arr P.Resid_Arr_self P.Srcs_Resid arrE P.Con_implies_Arr(2) con_char\<^sub>C\<^sub>C) have b: "arr \B\" by (simp add: "2" P.ide_char P.is_Cong_classI arr_char) have ide_b: "ide \B\" by (meson "2" P.arr_in_Cong_class P.ide_char P.ide_closed b disjoint_iff ide_char P.ide_implies_arr) have "targets \ = {\B\}" proof - have "cong (\ \\<^sup>*\\\<^sup>*\ \) \B\" proof - have "\ \\<^sup>*\\\<^sup>*\ \ = \B\" by (metis (no_types, lifting) "2" P.Cong_class_eqI P.Cong_closure_props(3) P.Resid_Arr_Ide_ind P.Resid_Ide(1) NPath_char \ con_char\<^sub>C\<^sub>C resid_char P.Con_implies_Arr(2) P.Resid_Arr_self mem_Collect_eq) thus ?thesis using b cong_reflexive by presburger qed thus ?thesis using \ targets_char\<^sub>Q\<^sub>C\<^sub>N [of \] cong_char by auto qed moreover have "\B\ = {B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B}" proof show "{B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B} \ \B\" using B P.Cong_class_def P.Cong_closure_props(3) P.Ide_implies_Arr P.ide_closed P.ide_char by force show "\B\ \ {B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B}" proof - have "\B'. P.Cong B' B \ P.Ide B' \ P.Trgs (P.Cong_class_rep \) = P.Srcs B'" using B NPath_char P.normal_is_Cong_closed Srcs_respects_Cong by (metis P.Cong_closure_props(1) mem_Collect_eq) thus ?thesis using P.Cong_class_def by blast qed qed ultimately show ?thesis using \ trg_in_targets by force qed lemma trg_char: shows "trg \ = {B. arr \ \ P.Ide B \ (\T. T \ \ \ P.Trgs T = P.Srcs B)}" proof (cases "arr \") show "\ arr \ \ ?thesis" using trg_char' by presburger assume \: "arr \" have "\T. T \ \ \ P.Trgs T = P.Trgs (P.Cong_class_rep \)" using \ by (metis P.Cong_class_memb_Cong_rep Trgs_respects_Cong arr_char) thus ?thesis using \ trg_char' P.is_Cong_class_def arr_char by force qed lemma is_extensional_rts_with_composites: shows "extensional_rts_with_composites Resid" proof fix \ \ assume seq: "seq \ \" obtain T where T: "\ = \T\" using seq P.Cong_class_rep arr_char seq_def by blast obtain U where U: "\ = \U\" using seq P.Cong_class_rep arr_char seq_def by blast have 1: "P.Arr T \ P.Arr U" using seq T U P.Con_implies_Arr(2) P.Cong\<^sub>0_subst_right(1) P.Cong_class_def P.con_char seq_def by (metis Collect_empty_eq P.Cong_imp_arr(1) P.arr_char P.rep_in_Cong_class empty_iff arr_char) have 2: "P.Trgs T = P.Srcs U" proof - have "targets \ = sources \" using seq seq_def sources_char targets_char\<^sub>W\<^sub>E by force hence 3: "trg \ = src \" using seq arr_has_un_source arr_has_un_target by (metis seq_def src_in_sources trg_in_targets) hence "{B. P.Ide B \ P.Trgs (P.Cong_class_rep \) = P.Srcs B} = {A. P.Ide A \ P.Srcs (P.Cong_class_rep \) = P.Srcs A}" using seq seq_def src_char' [of \] trg_char' [of \] by force hence "P.Trgs (P.Cong_class_rep \) = P.Srcs (P.Cong_class_rep \)" using seq seq_def arr_char by (metis (mono_tags, lifting) "3" P.Cong_class_is_nonempty Collect_empty_eq arr_src_iff_arr mem_Collect_eq trg_char') thus ?thesis using seq seq_def arr_char T U P.Srcs_respects_Cong P.Trgs_respects_Cong P.Cong_class_memb_Cong_rep P.Cong_symmetric by (metis "1" P.arr_char P.arr_in_Cong_class Srcs_respects_Cong Trgs_respects_Cong) qed have "P.Arr (T @ U)" using 1 2 by simp moreover have "P.Ide (T \<^sup>*\\\<^sup>* (T @ U))" by (metis "1" P.Con_append(2) P.Con_sym P.Resid_Arr_self P.Resid_Ide_Arr_ind P.Resid_append(2) P.Trgs.simps(1) calculation P.Arr_has_Trg) moreover have "(T @ U) \<^sup>*\\\<^sup>* T \<^sup>*\\<^sup>* U" by (metis "1" P.Arr.simps(1) P.Con_sym P.Cong\<^sub>0_append_resid_NPath P.Cong\<^sub>0_cancel_left\<^sub>C\<^sub>S P.Ide.simps(1) calculation(2) Cong_eq_Cong\<^sub>0 NPath_char) ultimately have "composite_of \ \ \T @ U\" proof (unfold composite_of_def, intro conjI) show "prfx \ (P.Cong_class (T @ U))" proof - have "ide (\ \\<^sup>*\\\<^sup>*\ \T @ U\)" proof (unfold ide_char, intro conjI) have 3: "T \<^sup>*\\\<^sup>* (T @ U) \ \ \\<^sup>*\\\<^sup>*\ \T @ U\" proof - have "\ \\<^sup>*\\\<^sup>*\ \T @ U\ = \T \<^sup>*\\\<^sup>* (T @ U)\" by (metis "1" P.Ide.simps(1) P.arr_char P.arr_in_Cong_class P.con_char P.is_Cong_classI Resid_by_members T \P.Arr (T @ U)\ \P.Ide (T \<^sup>*\\<^sup>* (T @ U))\) thus ?thesis by (simp add: P.arr_in_Cong_class P.elements_are_arr NPath_char \P.Ide (T \<^sup>*\\<^sup>* (T @ U))\) qed show "arr (\ \\<^sup>*\\\<^sup>*\ \T @ U\)" using 3 arr_char is_Cong_class_Resid by blast show "\ \\<^sup>*\\\<^sup>*\ \T @ U\ \ Collect P.NPath \ {}" using 3 P.ide_closed P.ide_char \P.Ide (T \<^sup>*\\<^sup>* (T @ U))\ by blast qed thus ?thesis by blast qed show "\T @ U\ \\<^sup>*\\\<^sup>*\ \ \\<^sup>*\\<^sup>*\ \" proof - have 3: "((T @ U) \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* U \ (\T @ U\ \\<^sup>*\\\<^sup>*\ \) \\<^sup>*\\\<^sup>*\ \" proof - have "(\T @ U\ \\<^sup>*\\\<^sup>*\ \) \\<^sup>*\\\<^sup>*\ \ = \((T @ U) \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* U\" proof - have "\T @ U\ \\<^sup>*\\\<^sup>*\ \ = \(T @ U) \<^sup>*\\\<^sup>* T\" by (metis "1" P.Cong_imp_arr(1) P.arr_char P.arr_in_Cong_class P.is_Cong_classI T \P.Arr (T @ U)\ \(T @ U) \<^sup>*\\<^sup>* T \<^sup>*\\<^sup>* U\ Resid_by_members P.arr_resid_iff_con) moreover have "\(T @ U) \<^sup>*\\\<^sup>* T\ \\<^sup>*\\\<^sup>*\ \ = \((T @ U) \<^sup>*\\\<^sup>* T) \<^sup>*\\\<^sup>* U\" by (metis "1" P.Cong_class_eqI P.Cong_imp_arr(1) P.arr_char P.arr_in_Cong_class P.con_char P.is_Cong_classI arr_char arrE U \(T @ U) \<^sup>*\\<^sup>* T \<^sup>*\\<^sup>* U\ con_char\<^sub>C\<^sub>C' Resid_by_members) ultimately show ?thesis by auto qed thus ?thesis by (metis "1" P.Arr.simps(1) P.Cong\<^sub>0_reflexive P.Resid_append(2) P.arr_char P.arr_in_Cong_class P.elements_are_arr \P.Arr (T @ U)\) qed have "\T @ U\ \\<^sup>*\\\<^sup>*\ \ \\<^sup>*\\<^sup>*\ \" proof (unfold ide_char, intro conjI) show "arr ((\T @ U\ \\<^sup>*\\\<^sup>*\ \) \\<^sup>*\\\<^sup>*\ \)" using 3 arr_char is_Cong_class_Resid by blast show "(\T @ U\ \\<^sup>*\\\<^sup>*\ \) \\<^sup>*\\\<^sup>*\ \ \ Collect P.NPath \ {}" by (metis 1 3 P.Arr.simps(1) P.Resid_append(2) P.con_char IntI \P.Arr (T @ U)\ NPath_char P.Resid_Arr_self P.arr_char empty_iff mem_Collect_eq P.arrE) qed thus ?thesis by blast qed show "\ \\<^sup>*\\<^sup>*\ \T @ U\ \\<^sup>*\\\<^sup>*\ \" proof (unfold ide_char, intro conjI) have 3: "U \<^sup>*\\\<^sup>* ((T @ U) \<^sup>*\\\<^sup>* T) \ \ \\<^sup>*\\\<^sup>*\ (\T @ U\ \\<^sup>*\\\<^sup>*\ \)" proof - have "\ \\<^sup>*\\\<^sup>*\ (\T @ U\ \\<^sup>*\\\<^sup>*\ \) = \U \<^sup>*\\\<^sup>* ((T @ U) \<^sup>*\\\<^sup>* T)\" proof - have "\T @ U\ \\<^sup>*\\\<^sup>*\ \ = \(T @ U) \<^sup>*\\\<^sup>* T\" by (metis "1" P.Con_sym P.Ide.simps(1) P.arr_char P.arr_in_Cong_class P.con_char P.is_Cong_classI Resid_by_members T \P.Arr (T @ U)\ \P.Ide (T \<^sup>*\\<^sup>* (T @ U))\) moreover have "\ \\<^sup>*\\\<^sup>*\ (\T @ U\ \\<^sup>*\\\<^sup>*\ \) = \U \<^sup>*\\\<^sup>* ((T @ U) \<^sup>*\\\<^sup>* T)\" by (metis "1" P.Cong_class_eqI P.Cong_imp_arr(1) P.arr_char P.arr_in_Cong_class P.con_char P.is_Cong_classI prfx_implies_con U \(T @ U) \<^sup>*\\<^sup>* T \<^sup>*\\<^sup>* U\ \\T @ U\ \\<^sup>*\\<^sup>*\ \ \\<^sup>*\\<^sup>*\ \\ calculation con_char\<^sub>C\<^sub>C' Resid_by_members) ultimately show ?thesis by blast qed thus ?thesis by (metis "1" P.Arr.simps(1) P.Resid_append_ind P.arr_in_Cong_class P.con_char \P.Arr (T @ U)\ P.Con_Arr_self P.arr_resid_iff_con) qed show "arr (\ \\<^sup>*\\\<^sup>*\ (\T @ U\ \\<^sup>*\\\<^sup>*\ \))" by (metis "3" arr_resid_iff_con empty_iff resid_char) show "\ \\<^sup>*\\\<^sup>*\ (\T @ U\ \\<^sup>*\\\<^sup>*\ \) \ Collect P.NPath \ {}" by (metis "1" "3" P.Arr.simps(1) P.Cong\<^sub>0_append_resid_NPath P.Cong\<^sub>0_cancel_left\<^sub>C\<^sub>S P.Cong_imp_arr(1) P.arr_char NPath_char IntI \(T @ U) \<^sup>*\\<^sup>* T \<^sup>*\\<^sup>* U\ \P.Ide (T \<^sup>*\\<^sup>* (T @ U))\ empty_iff) qed qed thus "composable \ \" using composable_def by auto qed sublocale extensional_rts_with_composites Resid using is_extensional_rts_with_composites by simp subsection "Inclusion Map" abbreviation incl where "incl t \ \[t]\" text \ The inclusion into the composite completion preserves consistency and residuation. \ lemma incl_preserves_con: assumes "t \ u" shows "\[t]\ \\<^sup>*\\<^sup>*\ \[u]\" using assms by (meson P.Con_rec(1) P.arr_in_Cong_class P.con_char P.is_Cong_classI con_char\<^sub>Q\<^sub>C\<^sub>N P.con_implies_arr(1-2)) lemma incl_preserves_resid: shows "\[t \\ u]\ = \[t]\ \\<^sup>*\\\<^sup>*\ \[u]\" proof (cases "t \ u") show "t \ u \ ?thesis" proof - assume 1: "t \ u" have "P.is_Cong_class \[t]\ \ P.is_Cong_class \[u]\" using 1 con_char\<^sub>Q\<^sub>C\<^sub>N incl_preserves_con by presburger moreover have "[t] \ \[t]\ \ [u] \ \[u]\" using 1 by (meson P.Con_rec(1) P.arr_in_Cong_class P.con_char P.Con_implies_Arr(2) P.arr_char P.con_implies_arr(1)) moreover have "P.con [t] [u]" using 1 by (simp add: P.con_char) ultimately show ?thesis using Resid_by_members [of "\[t]\" "\[u]\" "[t]" "[u]"] by (simp add: "1") qed assume 1: "\ t \ u" have "\[t \\ u]\ = {}" using 1 R.arrI by (metis Collect_empty_eq P.Con_Arr_self P.Con_rec(1) P.Cong_class_def P.Cong_imp_arr(1) P.arr_char R.arr_resid_iff_con) also have "... = \[t]\ \\<^sup>*\\\<^sup>*\ \[u]\" by (metis (full_types) "1" Con_char CollectD P.Con_rec(1) P.Cong_class_def P.Cong_imp_arr(1) P.arr_in_Cong_class con_char\<^sub>C\<^sub>C' null_char conI) finally show ?thesis by simp qed lemma incl_reflects_con: assumes "\[t]\ \\<^sup>*\\<^sup>*\ \[u]\" shows "t \ u" by (metis P.Con_rec(1) P.Cong_class_def P.Cong_imp_arr(1) P.arr_in_Cong_class CollectD assms con_char\<^sub>C\<^sub>C' con_char\<^sub>Q\<^sub>C\<^sub>N) text \ The inclusion map is a simulation. \ sublocale incl: simulation resid Resid incl proof show "\t. \ R.arr t \ incl t = null" by (metis Collect_empty_eq P.Cong_class_def P.Cong_imp_arr(1) P.Ide.simps(2) P.Resid_rec(1) P.cong_reflexive P.elements_are_arr P.ide_char P.ide_closed P.not_arr_null P.null_char R.prfx_implies_con null_char R.con_implies_arr(1)) show "\t u. t \ u \ incl t \\<^sup>*\\<^sup>*\ incl u" using incl_preserves_con by blast show "\t u. t \ u \ incl (t \\ u) = incl t \\<^sup>*\\\<^sup>*\ incl u" using incl_preserves_resid by blast qed lemma inclusion_is_simulation: shows "simulation resid Resid incl" .. lemma incl_preserves_arr: assumes "R.arr a" shows "arr \[a]\" using assms incl_preserves_con by auto lemma incl_preserves_ide: assumes "R.ide a" shows "ide \[a]\" by (metis assms incl_preserves_con incl_preserves_resid R.ide_def ide_def) lemma cong_iff_eq_incl: assumes "R.arr t" and "R.arr u" shows "\[t]\ = \[u]\ \ t \ u" proof show "\[t]\ = \[u]\ \ t \ u" by (metis P.Con_rec(1) P.Ide.simps(2) P.Resid.simps(3) P.arr_in_Cong_class P.con_char R.arr_def R.cong_reflexive assms(1) ide_char\<^sub>C\<^sub>C incl_preserves_con incl_preserves_ide incl_preserves_resid incl_reflects_con P.arr_resid_iff_con) show "t \ u \ \[t]\ = \[u]\" using assms by (metis incl_preserves_resid extensional incl_preserves_ide) qed text \ The inclusion is surjective on identities. \ lemma img_incl_ide: shows "incl ` (Collect R.ide) = Collect ide" proof show "incl ` Collect R.ide \ Collect ide" by (simp add: image_subset_iff) show "Collect ide \ incl ` Collect R.ide" proof fix \ assume \: "\ \ Collect ide" obtain A where A: "A \ \" using \ ide_char by blast have "P.NPath A" by (metis A Ball_Collect \ ide_char' mem_Collect_eq) obtain a where a: "a \ P.Srcs A" using \P.NPath A\ by (meson P.NPath_implies_Arr equals0I P.Arr_has_Src) have "P.Cong\<^sub>0 A [a]" proof - have "P.Ide [a]" by (metis NPath_char P.Con_Arr_self P.Ide.simps(2) P.NPath_implies_Arr P.Resid_Ide(1) P.Srcs.elims R.in_sourcesE \P.NPath A\ a) thus ?thesis using a A by (metis P.Ide.simps(2) P.ide_char P.ide_closed \P.NPath A\ NPath_char P.Con_single_ide_iff P.Ide_implies_Arr P.Resid_Arr_Ide_ind P.Resid_Arr_Src) qed have "\ = \[a]\" by (metis A P.Cong\<^sub>0_imp_con P.Cong\<^sub>0_implies_Cong P.Cong\<^sub>0_transitive P.Cong_class_eqI P.ide_char P.resid_arr_ide Resid_by_members \ \A \<^sup>*\\<^sub>0\<^sup>* [a]\ \P.NPath A\ arr_char NPath_char ideE ide_implies_arr mem_Collect_eq) thus "\ \ incl ` Collect R.ide" using NPath_char P.Ide.simps(2) P.backward_stable \A \<^sup>*\\<^sub>0\<^sup>* [a]\ \P.NPath A\ by blast qed qed end subsection "Composite Completion of an Extensional RTS" locale composite_completion_of_extensional_rts = R: extensional_rts + composite_completion begin sublocale P: paths_in_weakly_extensional_rts resid .. notation comp (infixl "\\<^sup>*\\<^sup>*\" 55) text \ When applied to an extensional RTS, the composite completion construction does not identify any states that are distinct in the original RTS. \ lemma incl_injective_on_ide: shows "inj_on incl (Collect R.ide)" using R.extensional cong_iff_eq_incl by (intro inj_onI) auto text \ When applied to an extensional RTS, the composite completion construction is a bijection between the states of the original RTS and the states of its completion. \ lemma incl_bijective_on_ide: shows "bij_betw incl (Collect R.ide) (Collect ide)" using incl_injective_on_ide img_incl_ide bij_betw_def by blast end subsection "Freeness of Composite Completion" text \ In this section we show that the composite completion construction is free: any simulation from RTS \A\ to an extensional RTS with composites \B\ extends uniquely to a simulation on the composite completion of \A\. \ locale extension_of_simulation = A: paths_in_rts resid\<^sub>A + B: extensional_rts_with_composites resid\<^sub>B + F: simulation resid\<^sub>A resid\<^sub>B F for resid\<^sub>A :: "'a resid" (infix "\\\<^sub>A" 70) and resid\<^sub>B :: "'b resid" (infix "\\\<^sub>B" 70) and F :: "'a \ 'b" begin notation A.Resid (infix "\<^sup>*\\\<^sub>A\<^sup>*" 70) notation A.Resid1x (infix "\<^sup>1\\\<^sub>A\<^sup>*" 70) notation A.Residx1 (infix "\<^sup>*\\\<^sub>A\<^sup>1" 70) notation A.Con (infix "\<^sup>*\\<^sub>A\<^sup>*" 70) notation B.comp (infixl "\\<^sub>B" 55) notation B.con (infix "\\<^sub>B" 50) fun map where "map [] = B.null" | "map [t] = F t" | "map (t # T) = (if A.arr (t # T) then F t \\<^sub>B map T else B.null)" lemma map_o_incl_eq: shows "map (A.incl t) = F t" by (simp add: A.null_char F.extensional) lemma extensional: shows "\ A.arr T \ map T = B.null" using F.extensional A.arr_char by (metis A.Arr.simps(2) map.elims) lemma preserves_comp: shows "\T \ []; U \ []; A.Arr (T @ U)\ \ map (T @ U) = map T \\<^sub>B map U" proof (induct T arbitrary: U) show "\U. [] \ [] \ map ([] @ U) = map [] \\<^sub>B map U" by simp fix t and T U :: "'a list" assume ind: "\U. \T \ []; U \ []; A.Arr (T @ U)\ \ map (T @ U) = map T \\<^sub>B map U" assume U: "U \ []" assume Arr: "A.Arr ((t # T) @ U)" hence 1: "A.Arr (t # (T @ U))" by simp have 2: "A.Arr (t # T)" by (metis A.Con_Arr_self A.Con_append(1) A.Con_implies_Arr(1) Arr U append_is_Nil_conv list.distinct(1)) show "map ((t # T) @ U) = B.comp (map (t # T)) (map U)" proof (cases "T = []") show "T = [] \ ?thesis" by (metis (full_types) "1" A.arr_char U append_Cons append_Nil list.exhaust map.simps(2) map.simps(3)) assume T: "T \ []" have "map ((t # T) @ U) = map (t # (T @ U))" by simp also have "... = F t \\<^sub>B map (T @ U)" using T 1 by (metis A.arr_char Nil_is_append_conv list.exhaust map.simps(3)) also have "... = F t \\<^sub>B (map T \\<^sub>B map U)" using ind by (metis "1" A.Con_Arr_self A.Con_implies_Arr(1) A.Con_rec(4) T U append_is_Nil_conv) also have "... = F t \\<^sub>B map T \\<^sub>B map U" using B.comp_assoc\<^sub>E\<^sub>C by blast also have "... = map (t # T) \\<^sub>B map U" using T 2 by (metis A.arr_char list.exhaust map.simps(3)) finally show "map ((t # T) @ U) = map (t # T) \\<^sub>B map U" by simp qed qed lemma preserves_arr_ind: shows "\A.arr T; a \ A.Srcs T\ \ B.arr (map T) \ B.src (map T) = F a" proof (induct T arbitrary: a) show "\a. \A.arr []; a \ A.Srcs []\ \ B.arr (map []) \ B.src (map []) = F a" using A.arr_char by simp fix a t T assume a: "a \ A.Srcs (t # T)" assume tT: "A.arr (t # T)" assume ind: "\a. \A.arr T; a \ A.Srcs T\ \ B.arr (map T) \ B.src (map T) = F a" have 1: "a \ A.R.sources t" using a tT A.Con_imp_eq_Srcs A.Con_initial_right A.Srcs.simps(2) A.con_char by blast show "B.arr (map (t # T)) \ B.src (map (t # T)) = F a" proof (cases "T = []") show "T = [] \ ?thesis" by (metis "1" A.Arr.simps(2) A.arr_char B.arr_has_un_source B.src_in_sources F.preserves_reflects_arr F.preserves_sources image_subset_iff map.simps(2) tT) assume T: "T \ []" obtain a' where a': "a' \ A.R.targets t" using tT "1" A.R.resid_source_in_targets by auto have 2: "a' \ A.Srcs T" using a' tT by (metis A.Con_Arr_self A.R.sources_resid A.Srcs.simps(2) A.arr_char T A.Con_imp_eq_Srcs A.Con_rec(4)) have "B.arr (map (t # T)) \ B.arr (F t \\<^sub>B map T)" using tT T by (metis map.simps(3) neq_Nil_conv) also have 2: "... \ True" by (metis (no_types, lifting) "2" A.arr_char B.arr_comp\<^sub>E\<^sub>C B.arr_has_un_target B.trg_in_targets F.preserves_reflects_arr F.preserves_targets T a' A.Arr.elims(2) image_subset_iff ind list.sel(1) list.sel(3) tT) finally have "B.arr (map (t # T))" by simp moreover have "B.src (map (t # T)) = F a" proof - have "B.src (map (t # T)) = B.src (F t \\<^sub>B map T)" using tT T by (metis map.simps(3) neq_Nil_conv) also have "... = B.src (F t)" using "2" B.con_comp_iff by force also have "... = F a" by (meson "1" B.weakly_extensional_rts_axioms F.simulation_axioms simulation_to_weakly_extensional_rts.preserves_src simulation_to_weakly_extensional_rts_def) finally show ?thesis by simp qed ultimately show ?thesis by simp qed qed lemma preserves_arr: shows "A.arr T \ B.arr (map T)" using preserves_arr_ind A.arr_char A.Arr_has_Src by blast lemma preserves_src: assumes "A.arr T" and "a \ A.Srcs T" shows "B.src (map T) = F a" using assms preserves_arr_ind by simp lemma preserves_trg: shows "\A.arr T; b \ A.Trgs T\ \ B.trg (map T) = F b" proof (induct T) show "\A.arr []; b \ A.Trgs []\ \ B.trg (map []) = F b" by simp fix t T assume tT: "A.arr (t # T)" assume b: "b \ A.Trgs (t # T)" assume ind: "\A.arr T; b \ A.Trgs T\ \ B.trg (map T) = F b" show "B.trg (map (t # T)) = F b" proof (cases "T = []") show "T = [] \ ?thesis" using tT b by (metis A.Trgs.simps(2) B.arr_has_un_target B.trg_in_targets F.preserves_targets preserves_arr image_subset_iff map.simps(2)) assume T: "T \ []" have 1: "B.trg (map (t # T)) = B.trg (F t \\<^sub>B map T)" using tT T b by (metis map.simps(3) neq_Nil_conv) also have "... = B.trg (map T)" by (metis B.arr_trg_iff_arr B.composable_iff_arr_comp B.trg_comp calculation preserves_arr tT) also have "... = F b" using tT b ind by (metis A.Trgs.simps(3) T A.Arr.simps(3) A.arr_char list.exhaust) finally show ?thesis by simp qed qed lemma preserves_Resid1x_ind: shows "t \<^sup>1\\\<^sub>A\<^sup>* U \ A.R.null \ F t \\<^sub>B map U \ F (t \<^sup>1\\\<^sub>A\<^sup>* U) = F t \\\<^sub>B map U" proof (induct U arbitrary: t) show "\t. t \<^sup>1\\\<^sub>A\<^sup>* [] \ A.R.null \ F t \\<^sub>B map [] \ F (t \<^sup>1\\\<^sub>A\<^sup>* []) = F t \\\<^sub>B map []" by simp fix t u U assume uU: "t \<^sup>1\\\<^sub>A\<^sup>* (u # U) \ A.R.null" assume ind: "\t. t \<^sup>1\\\<^sub>A\<^sup>* U \ A.R.null \ F t \\<^sub>B map U \ F (t \<^sup>1\\\<^sub>A\<^sup>* U) = F t \\\<^sub>B map U" show "F t \\<^sub>B map (u # U) \ F (t \<^sup>1\\\<^sub>A\<^sup>* (u # U)) = F t \\\<^sub>B map (u # U)" proof show 1: "F t \\<^sub>B map (u # U)" proof (cases "U = []") show "U = [] \ ?thesis" using A.Resid1x.simps(2) map.simps(2) F.preserves_con uU by fastforce assume U: "U \ []" have 3: "[t] \<^sup>*\\\<^sub>A\<^sup>* [u] \ [] \ ([t] \<^sup>*\\\<^sub>A\<^sup>* [u]) \<^sup>*\\\<^sub>A\<^sup>* U \ []" using A.Con_cons(2) [of "[t]" U u] by (meson A.Resid1x_as_Resid' U not_Cons_self2 uU) hence 2: "F t \\<^sub>B F u \ F t \\\<^sub>B F u \\<^sub>B map U" by (metis A.Con_rec(1) A.Con_sym A.Con_sym1 A.Residx1_as_Resid A.Resid_rec(1) F.preserves_con F.preserves_resid ind) moreover have "B.seq (F u) (map U)" by (metis B.coinitial_iff\<^sub>W\<^sub>E B.con_imp_coinitial B.seqI\<^sub>W\<^sub>E B.src_resid calculation) ultimately have "F t \\<^sub>B map ([u] @ U)" using B.con_comp_iff\<^sub>E\<^sub>C(1) [of "F t" "F u" "map U"] B.con_sym preserves_comp by (metis 3 A.Con_cons(2) A.Con_implies_Arr(2) append.left_neutral append_Cons map.simps(2) not_Cons_self2) thus ?thesis by simp qed show "F (t \<^sup>1\\\<^sub>A\<^sup>* (u # U)) = F t \\\<^sub>B map (u # U)" proof (cases "U = []") show "U = [] \ ?thesis" using A.Resid1x.simps(2) F.preserves_resid map.simps(2) uU by fastforce assume U: "U \ []" have "F (t \<^sup>1\\\<^sub>A\<^sup>* (u # U)) = F ((t \\\<^sub>A u) \<^sup>1\\\<^sub>A\<^sup>* U)" using A.Resid1x_as_Resid' A.Resid_rec(3) U uU by metis also have "... = F (t \\\<^sub>A u) \\\<^sub>B map U" using uU U ind A.Con_rec(3) A.Resid1x_as_Resid [of "t \\\<^sub>A u" U] by (metis A.Resid1x.simps(3) list.exhaust) also have "... = (F t \\\<^sub>B F u) \\\<^sub>B map U" using uU U by (metis A.Resid1x_as_Resid' F.preserves_resid A.Con_rec(3)) also have "... = F t \\\<^sub>B (F u \\<^sub>B map U)" by (metis B.comp_null(2) B.composable_iff_comp_not_null B.con_compI(2) B.conI B.con_sym_ax B.mediating_transition B.null_is_zero(2) B.resid_comp(1)) also have "... = F t \\\<^sub>B map (u # U)" by (metis A.Resid1x_as_Resid' A.con_char U map.simps(3) neq_Nil_conv A.con_implies_arr(2) uU) finally show ?thesis by simp qed qed qed lemma preserves_Residx1_ind: shows "U \<^sup>*\\\<^sub>A\<^sup>1 t \ [] \ map U \\<^sub>B F t \ map (U \<^sup>*\\\<^sub>A\<^sup>1 t) = map U \\\<^sub>B F t" proof (induct U arbitrary: t) show "\t. [] \<^sup>*\\\<^sub>A\<^sup>1 t \ [] \ map [] \\<^sub>B F t \ map ([] \<^sup>*\\\<^sub>A\<^sup>1 t) = map [] \\\<^sub>B F t" by simp fix t u U assume ind: "\t. U \<^sup>*\\\<^sub>A\<^sup>1 t \ [] \ map U \\<^sub>B F t \ map (U \<^sup>*\\\<^sub>A\<^sup>1 t) = map U \\\<^sub>B F t" assume uU: "(u # U) \<^sup>*\\\<^sub>A\<^sup>1 t \ []" show "map (u # U) \\<^sub>B F t \ map ((u # U) \<^sup>*\\\<^sub>A\<^sup>1 t) = map (u # U) \\\<^sub>B F t" proof (cases "U = []") show "U = [] \ ?thesis" using A.Residx1.simps(2) F.preserves_con F.preserves_resid map.simps(2) uU by presburger assume U: "U \ []" show ?thesis proof show "map (u # U) \\<^sub>B F t" using uU U A.Con_sym1 B.con_sym preserves_Resid1x_ind by blast show "map ((u # U) \<^sup>*\\\<^sub>A\<^sup>1 t) = map (u # U) \\\<^sub>B F t" proof - have "map ((u # U) \<^sup>*\\\<^sub>A\<^sup>1 t) = map ((u \\\<^sub>A t) # U \<^sup>*\\\<^sub>A\<^sup>1 (t \\\<^sub>A u))" using uU U A.Residx1_as_Resid A.Resid_rec(2) by fastforce also have "... = F (u \\\<^sub>A t) \\<^sub>B map (U \<^sup>*\\\<^sub>A\<^sup>1 (t \\\<^sub>A u))" by (metis A.Residx1_as_Resid A.arr_char U A.Con_imp_Arr_Resid A.Con_rec(2) A.Resid_rec(2) list.exhaust map.simps(3) uU) also have "... = F (u \\\<^sub>A t) \\<^sub>B map U \\\<^sub>B F (t \\\<^sub>A u)" using uU U ind A.Con_rec(2) A.Residx1_as_Resid by force also have "... = (F u \\\<^sub>B F t) \\<^sub>B map U \\\<^sub>B (F t \\\<^sub>B F u)" using uU U by (metis A.Con_initial_right A.Con_rec(1) A.Con_sym1 A.Resid1x_as_Resid' A.Residx1_as_Resid F.preserves_resid) also have "... = (F u \\<^sub>B map U) \\\<^sub>B F t" by (metis B.comp_null(2) B.composable_iff_comp_not_null B.con_compI(2) B.con_sym B.mediating_transition B.null_is_zero(2) B.resid_comp(2) B.con_def) also have "... = map (u # U) \\\<^sub>B F t" by (metis A.Con_implies_Arr(2) A.Con_sym A.Residx1_as_Resid U A.arr_char map.simps(3) neq_Nil_conv uU) finally show ?thesis by simp qed qed qed qed lemma preserves_resid_ind: shows "A.con T U \ map T \\<^sub>B map U \ map (T \<^sup>*\\\<^sub>A\<^sup>* U) = map T \\\<^sub>B map U" proof (induct T arbitrary: U) show "\U. A.con [] U \ map [] \\<^sub>B map U \ map ([] \<^sup>*\\\<^sub>A\<^sup>* U) = map [] \\\<^sub>B map U" using A.con_char A.Resid.simps(1) by blast fix t T U assume tT: "A.con (t # T) U" assume ind: "\U. A.con T U \ map T \\<^sub>B map U \ map (T \<^sup>*\\\<^sub>A\<^sup>* U) = map T \\\<^sub>B map U" show "map (t # T) \\<^sub>B map U \ map ((t # T) \<^sup>*\\\<^sub>A\<^sup>* U) = map (t # T) \\\<^sub>B map U" proof (cases "T = []") assume T: "T = []" show ?thesis using T tT apply simp by (metis A.Resid1x_as_Resid A.Residx1_as_Resid A.con_char A.Con_sym A.Con_sym1 map.simps(2) preserves_Resid1x_ind) next assume T: "T \ []" have 1: "map (t # T) = F t \\<^sub>B map T" using tT T by (metis A.con_implies_arr(1) list.exhaust map.simps(3)) show ?thesis proof show 2: "B.con (map (t # T)) (map U)" using T tT by (metis "1" A.Con_cons(1) A.Residx1_as_Resid A.con_char A.not_arr_null A.null_char B.composable_iff_comp_not_null B.con_compI(2) B.con_sym B.not_arr_null preserves_arr ind preserves_Residx1_ind A.con_implies_arr(1-2)) show "map ((t # T) \<^sup>*\\\<^sub>A\<^sup>* U) = map (t # T) \\\<^sub>B map U" proof - have "map ((t # T) \<^sup>*\\\<^sub>A\<^sup>* U) = map (([t] \<^sup>*\\\<^sub>A\<^sup>* U) @ (T \<^sup>*\\\<^sub>A\<^sup>* (U \<^sup>*\\\<^sub>A\<^sup>* [t])))" by (metis A.Resid.simps(1) A.Resid_cons(1) A.con_char A.ex_un_null tT) also have "... = map ([t] \<^sup>*\\\<^sub>A\<^sup>* U) \\<^sub>B map (T \<^sup>*\\\<^sub>A\<^sup>* (U \<^sup>*\\\<^sub>A\<^sup>* [t]))" by (metis A.Arr.simps(1) A.Con_imp_Arr_Resid A.Con_implies_Arr(2) A.Con_sym A.Resid_cons(1-2) A.con_char T preserves_comp tT) also have "... = (map [t] \\\<^sub>B map U) \\<^sub>B map (T \<^sup>*\\\<^sub>A\<^sup>* (U \<^sup>*\\\<^sub>A\<^sup>* [t]))" by (metis A.Con_initial_right A.Con_sym A.Resid1x_as_Resid A.Residx1_as_Resid A.con_char A.Con_sym1 map.simps(2) preserves_Resid1x_ind tT) also have "... = (map [t] \\\<^sub>B map U) \\<^sub>B (map T \\\<^sub>B map (U \<^sup>*\\\<^sub>A\<^sup>* [t]))" using tT T ind by (metis A.Con_cons(1) A.Con_sym A.Resid.simps(1) A.con_char) also have "... = (map [t] \\\<^sub>B map U) \\<^sub>B (map T \\\<^sub>B (map U \\\<^sub>B map [t]))" using tT T by (metis A.Con_cons(1) A.Con_sym A.Resid.simps(2) A.Residx1_as_Resid A.con_char map.simps(2) preserves_Residx1_ind) also have "... = (F t \\\<^sub>B map U) \\<^sub>B (map T \\\<^sub>B (map U \\\<^sub>B F t))" using tT T by simp also have "... = map (t # T) \\\<^sub>B map U" using 1 2 B.resid_comp(2) by presburger finally show ?thesis by simp qed qed qed qed lemma preserves_con: assumes "A.con T U" shows "map T \\<^sub>B map U" using assms preserves_resid_ind by simp lemma preserves_resid: assumes "A.con T U" shows "map (T \<^sup>*\\\<^sub>A\<^sup>* U) = map T \\\<^sub>B map U" using assms preserves_resid_ind by simp sublocale simulation A.Resid resid\<^sub>B map using A.con_char preserves_con preserves_resid extensional by unfold_locales auto sublocale simulation_to_extensional_rts A.Resid resid\<^sub>B map .. lemma is_universal: assumes "rts_with_composites resid\<^sub>B" and "simulation resid\<^sub>A resid\<^sub>B F" shows "\!F'. simulation A.Resid resid\<^sub>B F' \ F' o A.incl = F" proof interpret B: rts_with_composites resid\<^sub>B using assms by auto interpret F: simulation resid\<^sub>A resid\<^sub>B F using assms by auto show "simulation A.Resid resid\<^sub>B map \ map \ A.incl = F" using map_o_incl_eq simulation_axioms by auto show "\F'. simulation A.Resid resid\<^sub>B F' \ F' o A.incl = F \ F' = map" proof fix F' T assume F': "simulation A.Resid resid\<^sub>B F' \ F' o A.incl = F" interpret F': simulation A.Resid resid\<^sub>B F' using F' by simp show "F' T = map T" proof (induct T) show "F' [] = map []" by (simp add: A.arr_char F'.extensional) fix t T assume ind: "F' T = map T" show "F' (t # T) = map (t # T)" proof (cases "A.Arr (t # T)") show "\ A.Arr (t # T) \ ?thesis" by (simp add: A.arr_char F'.extensional extensional) assume tT: "A.Arr (t # T)" show ?thesis proof (cases "T = []") show 2: "T = [] \ ?thesis" using F' tT by auto assume T: "T \ []" have "F' (t # T) = F' [t] \\<^sub>B map T" proof - have "F' (t # T) = F' ([t] @ T)" by simp also have "... = F' [t] \\<^sub>B F' T" proof - have "A.composite_of [t] T ([t] @ T)" using T tT by (metis (full_types) A.Arr.simps(2) A.Con_Arr_self A.append_is_composite_of A.Con_implies_Arr(1) A.Con_imp_eq_Srcs A.Con_rec(4) A.Resid_rec(1) A.Srcs_Resid A.seq_char A.R.arrI) thus ?thesis using F'.preserves_composites [of "[t]" T "[t] @ T"] B.comp_is_composite_of by auto qed also have "... = F' [t] \\<^sub>B map T" using T ind by simp finally show ?thesis by simp qed also have "... = (F' \ A.incl) t \\<^sub>B map T" using tT by (simp add: A.arr_char A.null_char F'.extensional) also have "... = F t \\<^sub>B map T" using F' by simp also have "... = map (t # T)" using T tT by (metis A.arr_char list.exhaust map.simps(3)) finally show ?thesis by simp qed qed qed qed qed end (* * TODO: Localize to context rts? *) lemma composite_completion_of_rts: assumes "rts A" shows "\(C :: 'a list resid) I. rts_with_composites C \ simulation A C I \ (\B (J :: 'a \ 'c). extensional_rts_with_composites B \ simulation A B J \ (\!J'. simulation C B J' \ J' o I = J))" proof (intro exI conjI) interpret A: rts A using assms by auto interpret P\<^sub>A: paths_in_rts A .. show "rts_with_composites P\<^sub>A.Resid" using P\<^sub>A.rts_with_composites_axioms by simp show "simulation A P\<^sub>A.Resid P\<^sub>A.incl" using P\<^sub>A.incl_is_simulation by simp show "\B (J :: 'a \ 'c). extensional_rts_with_composites B \ simulation A B J \ (\!J'. simulation P\<^sub>A.Resid B J' \ J' o P\<^sub>A.incl = J)" proof (intro allI impI) fix B :: "'c resid" and J assume 1: "extensional_rts_with_composites B \ simulation A B J" interpret B: extensional_rts_with_composites B using 1 by simp interpret J: simulation A B J using 1 by simp interpret J: extension_of_simulation A B J .. have "simulation P\<^sub>A.Resid B J.map" using J.simulation_axioms by simp moreover have "J.map o P\<^sub>A.incl = J" using J.map_o_incl_eq by auto moreover have "\J'. simulation P\<^sub>A.Resid B J' \ J' o P\<^sub>A.incl = J \ J' = J.map" using "1" B.rts_with_composites_axioms J.is_universal J.simulation_axioms calculation(2) by blast ultimately show "\!J'. simulation P\<^sub>A.Resid B J' \ J' \ P\<^sub>A.incl = J" by auto qed qed section "Constructions on RTS's" subsection "Products of RTS's" locale product_rts = R1: rts R1 + R2: rts R2 for R1 :: "'a1 resid" (infix "\\\<^sub>1" 70) and R2 :: "'a2 resid" (infix "\\\<^sub>2" 70) begin type_synonym ('aa1, 'aa2) arr = "'aa1 * 'aa2" abbreviation (input) Null :: "('a1, 'a2) arr" where "Null \ (R1.null, R2.null)" definition resid :: "('a1, 'a2) arr \ ('a1, 'a2) arr \ ('a1, 'a2) arr" where "resid t u = (if R1.con (fst t) (fst u) \ R2.con (snd t) (snd u) then (fst t \\\<^sub>1 fst u, snd t \\\<^sub>2 snd u) else Null)" notation resid (infix "\\" 70) sublocale partial_magma resid by unfold_locales (metis R1.con_implies_arr(1-2) R1.not_arr_null fst_conv resid_def) lemma is_partial_magma: shows "partial_magma resid" .. lemma null_char [simp]: shows "null = Null" by (metis R2.null_is_zero(1) R2.residuation_axioms ex_un_null null_is_zero(1) resid_def residuation.conE snd_conv) sublocale residuation resid proof show "\t u. t \\ u \ null \ u \\ t \ null" by (metis R1.con_def R1.con_sym null_char prod.inject resid_def R2.con_sym) show "\t u. t \\ u \ null \ (t \\ u) \\ (t \\ u) \ null" by (metis (no_types, lifting) R1.arrE R2.con_def R2.con_imp_arr_resid fst_conv null_char resid_def R1.arr_resid snd_conv) show "\v t u. (v \\ t) \\ (u \\ t) \ null \ (v \\ t) \\ (u \\ t) = (v \\ u) \\ (t \\ u)" proof - fix t u v assume 1: "(v \\ t) \\ (u \\ t) \ null" have "(fst v \\\<^sub>1 fst t) \\\<^sub>1 (fst u \\\<^sub>1 fst t) \ R1.null" by (metis 1 R1.not_arr_null fst_conv null_char null_is_zero(1-2) resid_def R1.arr_resid) moreover have "(snd v \\\<^sub>2 snd t) \\\<^sub>2 (snd u \\\<^sub>2 snd t) \ R2.null" by (metis 1 R2.not_arr_null snd_conv null_char null_is_zero(1-2) resid_def R2.arr_resid) ultimately show "(v \\ t) \\ (u \\ t) = (v \\ u) \\ (t \\ u)" using resid_def null_char R1.con_def R2.con_def R1.cube R2.cube apply simp by (metis (no_types, lifting) R1.conI R1.con_sym_ax R1.resid_reflects_con R2.con_sym_ax R2.null_is_zero(1)) qed qed lemma is_residuation: shows "residuation resid" .. lemma arr_char [iff]: shows "arr t \ R1.arr (fst t) \ R2.arr (snd t)" by (metis (no_types, lifting) R1.arr_def R2.arr_def R2.conE null_char resid_def residuation.arr_def residuation.con_def residuation_axioms snd_eqD) lemma ide_char [iff]: shows "ide t \ R1.ide (fst t) \ R2.ide (snd t)" by (metis (no_types, lifting) R1.residuation_axioms R2.residuation_axioms arr_char arr_def fst_conv null_char prod.collapse resid_def residuation.conE residuation.ide_def residuation.ide_implies_arr residuation_axioms snd_conv) lemma con_char [iff]: shows "con t u \ R1.con (fst t) (fst u) \ R2.con (snd t) (snd u)" by (simp add: R2.residuation_axioms con_def resid_def residuation.con_def) lemma trg_char: shows "trg t = (if arr t then (R1.trg (fst t), R2.trg (snd t)) else Null)" using R1.trg_def R2.trg_def resid_def trg_def by auto sublocale rts resid proof show "\t. arr t \ ide (trg t)" by (simp add: trg_char) show "\a t. \ide a; con t a\ \ t \\ a = t" by (simp add: R1.resid_arr_ide R2.resid_arr_ide resid_def) show "\a t. \ide a; con a t\ \ ide (a \\ t)" by (metis \\t a. \ide a; con t a\ \ t \ a = t\ con_sym cube ideE ideI residuation.con_def residuation_axioms) show "\t u. con t u \ \a. ide a \ con a t \ con a u" proof - fix t u assume tu: "con t u" obtain a1 where a1: "a1 \ R1.sources (fst t) \ R1.sources (fst u)" by (meson R1.con_imp_common_source all_not_in_conv con_char tu) obtain a2 where a2: "a2 \ R2.sources (snd t) \ R2.sources (snd u)" by (meson R2.con_imp_common_source all_not_in_conv con_char tu) have "ide (a1, a2) \ con (a1, a2) t \ con (a1, a2) u" using a1 a2 ide_char con_char by (metis R1.con_imp_common_source R1.in_sourcesE R1.sources_eqI R2.con_imp_common_source R2.in_sourcesE R2.sources_eqI con_sym fst_conv inf_idem snd_conv tu) thus "\a. ide a \ con a t \ con a u" by blast qed show "\t u v. \ide (t \\ u); con u v\ \ con (t \\ u) (v \\ u)" proof - fix t u v assume tu: "ide (t \\ u)" assume uv: "con u v" have "R1.ide (fst t \\\<^sub>1 fst u) \ R2.ide (snd t \\\<^sub>2 snd u)" using tu ide_char by (metis conI con_char fst_eqD ide_implies_arr not_arr_null resid_def snd_conv) moreover have "R1.con (fst u) (fst v) \ R2.con (snd u) (snd v)" using uv con_char by blast ultimately show "con (t \\ u) (v \\ u)" by (simp add: R1.con_target R1.con_sym R1.prfx_implies_con R2.con_target R2.con_sym R2.prfx_implies_con resid_def) qed qed lemma is_rts: shows "rts resid" .. lemma sources_char: shows "sources t = R1.sources (fst t) \ R2.sources (snd t)" by force lemma targets_char: shows "targets t = R1.targets (fst t) \ R2.targets (snd t)" proof show "targets t \ R1.targets (fst t) \ R2.targets (snd t)" using targets_def ide_char con_char resid_def trg_char trg_def by auto show "R1.targets (fst t) \ R2.targets (snd t) \ targets t" proof fix a assume a: "a \ R1.targets (fst t) \ R2.targets (snd t)" show "a \ targets t" proof show "ide a" using a ide_char by auto show "con (trg t) a" using a trg_char con_char [of "trg t" a] by (metis (no_types, lifting) SigmaE arr_char con_char con_implies_arr(1) fst_conv R1.in_targetsE R2.in_targetsE R1.arr_resid_iff_con R2.arr_resid_iff_con R1.trg_def R2.trg_def snd_conv) qed qed qed lemma prfx_char: shows "prfx t u \ R1.prfx (fst t) (fst u) \ R2.prfx (snd t) (snd u)" using R1.prfx_implies_con R2.prfx_implies_con resid_def by auto lemma cong_char: shows "cong t u \ R1.cong (fst t) (fst u) \ R2.cong (snd t) (snd u)" using prfx_char by auto end locale product_of_weakly_extensional_rts = R1: weakly_extensional_rts R1 + R2: weakly_extensional_rts R2 + product_rts begin sublocale weakly_extensional_rts resid proof show "\t u. \cong t u; ide t; ide u\ \ t = u" by (metis cong_char ide_char prod.exhaust_sel R1.weak_extensionality R2.weak_extensionality) qed lemma src_char: shows "src t = (if arr t then (R1.src (fst t), R2.src (snd t)) else null)" proof (cases "arr t") show "\ arr t \ ?thesis" using src_def by presburger assume t: "arr t" show ?thesis proof (intro src_eqI) show "ide (if arr t then (R1.src (fst t), R2.src (snd t)) else null)" using t by simp show "con (if arr t then (R1.src (fst t), R2.src (snd t)) else null) t" using t con_char arr_char apply (cases t) apply simp_all by (metis R1.con_imp_coinitial_ax R1.residuation_axioms R1.src_eqI R2.con_sym R2.in_sourcesE R2.src_in_sources residuation.arr_def) qed qed end locale product_of_extensional_rts = R1: extensional_rts R1 + R2: extensional_rts R2 + product_of_weakly_extensional_rts begin sublocale extensional_rts resid proof show "\t u. cong t u \ t = u" by (metis R1.extensional R2.extensional cong_char prod.collapse) qed end subsubsection "Product Simulations" locale product_simulation = A1: rts A1 + A2: rts A2 + B1: rts B1 + B2: rts B2 + A1xA2: product_rts A1 A2 + B1xB2: product_rts B1 B2 + F1: simulation A1 B1 F1 + F2: simulation A2 B2 F2 for A1 :: "'a1 resid" (infix "\\\<^sub>A\<^sub>1" 70) and A2 :: "'a2 resid" (infix "\\\<^sub>A\<^sub>2" 70) and B1 :: "'b1 resid" (infix "\\\<^sub>B\<^sub>1" 70) and B2 :: "'b2 resid" (infix "\\\<^sub>B\<^sub>2" 70) and F1 :: "'a1 \ 'b1" and F2 :: "'a2 \ 'b2" begin definition map where "map = (\a. if A1xA2.arr a then (F1 (fst a), F2 (snd a)) else B1xB2.null)" lemma map_simp [simp]: assumes "A1.arr a1" and "A2.arr a2" shows "map (a1, a2) = (F1 a1, F2 a2)" using assms map_def by auto sublocale simulation A1xA2.resid B1xB2.resid map proof show "\t. \ A1xA2.arr t \ map t = B1xB2.null" using map_def by auto show "\t u. A1xA2.con t u \ B1xB2.con (map t) (map u)" using A1xA2.con_char B1xB2.con_char A1.con_implies_arr A2.con_implies_arr by auto show "\t u. A1xA2.con t u \ map (A1xA2.resid t u) = B1xB2.resid (map t) (map u)" using A1xA2.resid_def B1xB2.resid_def A1.con_implies_arr A2.con_implies_arr by auto qed lemma is_simulation: shows "simulation A1xA2.resid B1xB2.resid map" .. end subsubsection "Binary Simulations" locale binary_simulation = A1: rts A1 + A2: rts A2 + A: product_rts A1 A2 + B: rts B + simulation A.resid B F for A1 :: "'a1 resid" (infixr "\\\<^sub>A\<^sub>1" 70) and A2 :: "'a2 resid" (infixr "\\\<^sub>A\<^sub>2" 70) and B :: "'b resid" (infixr "\\\<^sub>B" 70) and F :: "'a1 * 'a2 \ 'b" begin lemma fixing_ide_gives_simulation_1: assumes "A1.ide a1" shows "simulation A2 B (\t2. F (a1, t2))" proof show "\t2. \ A2.arr t2 \ F (a1, t2) = B.null" using assms extensional A.arr_char by simp show "\t2 u2. A2.con t2 u2 \ B.con (F (a1, t2)) (F (a1, u2))" using assms A.con_char preserves_con by auto show "\t2 u2. A2.con t2 u2 \ F (a1, t2 \\\<^sub>A\<^sub>2 u2) = F (a1, t2) \\\<^sub>B F (a1, u2)" using assms A.con_char A.resid_def preserves_resid by (metis A1.ideE fst_conv snd_conv) qed lemma fixing_ide_gives_simulation_2: assumes "A2.ide a2" shows "simulation A1 B (\t1. F (t1, a2))" proof show "\t1. \ A1.arr t1 \ F (t1, a2) = B.null" using assms extensional A.arr_char by simp show "\t1 u1. A1.con t1 u1 \ B.con (F (t1, a2)) (F (u1, a2))" using assms A.con_char preserves_con by auto show "\t1 u1. A1.con t1 u1 \ F (t1 \\\<^sub>A\<^sub>1 u1, a2) = F (t1, a2) \\\<^sub>B F (u1, a2)" using assms A.con_char A.resid_def preserves_resid by (metis A2.ideE fst_conv snd_conv) qed end subsection "Sub-RTS's" locale sub_rts = R: rts R for R :: "'a resid" (infix "\\\<^sub>R" 70) and Arr :: "'a \ bool" + assumes inclusion: "Arr t \ R.arr t" and sources_closed: "Arr t \ R.sources t \ Collect Arr" and resid_closed: "\Arr t; Arr u; R.con t u\ \ Arr (t \\\<^sub>R u)" begin definition resid (infix "\\" 70) where "t \\ u \ (if Arr t \ Arr u \ R.con t u then t \\\<^sub>R u else R.null)" sublocale partial_magma resid by unfold_locales (metis R.ex_un_null R.null_is_zero(2) resid_def) lemma is_partial_magma: shows "partial_magma resid" .. lemma null_char [simp]: shows "null = R.null" by (metis R.null_is_zero(1) ex_un_null null_is_zero(1) resid_def) sublocale residuation resid proof show "\t u. t \\ u \ null \ u \\ t \ null" by (metis R.con_sym R.con_sym_ax null_char resid_def) show "\t u. t \\ u \ null \ (t \\ u) \\ (t \\ u) \ null" by (metis R.arrE R.arr_resid R.not_arr_null null_char resid_closed resid_def) show "\v t u. (v \\ t) \\ (u \\ t) \ null \ (v \\ t) \\ (u \\ t) = (v \\ u) \\ (t \\ u)" by (metis R.cube R.ex_un_null R.null_is_zero(1) R.residuation_axioms null_is_zero(2) resid_closed resid_def residuation.conE residuation.conI) qed lemma is_residuation: shows "residuation resid" .. lemma arr_char [iff]: shows "arr t \ Arr t" proof show "arr t \ Arr t" by (metis arrE conE null_char resid_def) show "Arr t \ arr t" by (metis R.arrE R.conE conI con_implies_arr(2) inclusion null_char resid_def) qed lemma ide_char [iff]: shows "ide t \ Arr t \ R.ide t" by (metis R.ide_def arrE arr_char conE ide_def null_char resid_def) lemma con_char [iff]: shows "con t u \ Arr t \ Arr u \ R.con t u" using con_def resid_def by auto lemma trg_char: shows "trg t = (if arr t then R.trg t else null)" using R.trg_def arr_def resid_def trg_def by force sublocale rts resid proof show "\t. arr t \ ide (trg t)" by (metis R.ide_trg arrE arr_char arr_resid ide_char inclusion trg_char trg_def) show "\a t. \ide a; con t a\ \ t \\ a = t" by (simp add: R.resid_arr_ide resid_def) show "\a t. \ide a; con a t\ \ ide (a \\ t)" by (metis R.resid_ide_arr arr_resid_iff_con arr_char con_char ide_char resid_def) show "\t u. con t u \ \a. ide a \ con a t \ con a u" by (metis (full_types) R.con_imp_coinitial_ax R.con_sym R.in_sourcesI con_char ide_char in_mono mem_Collect_eq sources_closed) show "\t u v. \ide (t \\ u); con u v\ \ con (t \\ u) (v \\ u)" by (metis R.con_target arr_resid_iff_con con_char con_sym ide_char ide_implies_arr resid_closed resid_def) qed lemma is_rts: shows "rts resid" .. lemma sources_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S: shows "sources t = {a. Arr t \ a \ R.sources t}" using sources_closed by auto lemma targets_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S: shows "targets t = {b. Arr t \ b \ R.targets t}" proof show "targets t \ {b. Arr t \ b \ R.targets t}" proof fix b assume b: "b \ targets t" show "b \ {b. Arr t \ b \ R.targets t}" proof have "Arr t" using arr_iff_has_target b by force moreover have "Arr b" using b by blast moreover have "b \ R.targets t" by (metis R.in_targetsI b calculation(1) con_char in_targetsE arr_char ide_char trg_char) ultimately show "Arr t \ b \ R.targets t" by blast qed qed show "{b. Arr t \ b \ R.targets t} \ targets t" proof fix b assume b: "b \ {b. Arr t \ b \ R.targets t}" show "b \ targets t" proof (intro in_targetsI) show "ide b" using b by (metis R.arrE ide_char inclusion mem_Collect_eq R.sources_resid R.target_is_ide resid_closed sources_closed subset_eq) show "con (trg t) b" using b using \ide b\ ide_trg trg_char by auto qed qed qed lemma prfx_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S: shows "prfx t u \ Arr t \ Arr u \ R.prfx t u" by (metis R.prfx_implies_con con_char ide_char prfx_implies_con resid_closed resid_def) lemma cong_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S: shows "cong t u \ Arr t \ Arr u \ R.cong t u" using prfx_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S by force lemma inclusion_is_simulation: shows "simulation resid R (\t. if arr t then t else null)" using resid_closed resid_def by unfold_locales auto interpretation P\<^sub>R: paths_in_rts R .. interpretation P: paths_in_rts resid .. lemma path_reflection: shows "\P\<^sub>R.Arr T; set T \ Collect Arr\ \ P.Arr T" apply (induct T) apply simp proof - fix t T assume ind: "\P\<^sub>R.Arr T; set T \ Collect Arr\ \ P.Arr T" assume tT: "P\<^sub>R.Arr (t # T)" assume set: "set (t # T) \ Collect Arr" have 1: "R.arr t" using tT by (metis P\<^sub>R.Arr_imp_arr_hd list.sel(1)) show "P.Arr (t # T)" proof (cases "T = []") show "T = [] \ ?thesis" using 1 set by simp assume T: "T \ []" show ?thesis proof show "arr t" using 1 arr_char set by simp show "P.Arr T" using T tT P\<^sub>R.Arr_imp_Arr_tl by (metis ind insert_subset list.sel(3) list.simps(15) set) show "targets t \ P.Srcs T" proof - have "targets t \ R.targets t" using targets_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S by blast also have "... \ R.sources (hd T)" using T tT by (metis P\<^sub>R.Arr.simps(3) P\<^sub>R.Srcs_simp\<^sub>P list.collapse) also have "... \ P.Srcs T" using P.Arr_imp_arr_hd P.Srcs_simp\<^sub>P \P.Arr T\ sources_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S by force finally show ?thesis by blast qed qed qed qed end locale sub_weakly_extensional_rts = sub_rts + R: weakly_extensional_rts R begin sublocale weakly_extensional_rts resid apply unfold_locales using R.weak_extensionality cong_char\<^sub>S\<^sub>R\<^sub>T\<^sub>S by blast lemma is_weakly_extensional_rts: shows "weakly_extensional_rts resid" .. lemma src_char: shows "src t = (if arr t then R.src t else null)" proof (cases "arr t") show "\ arr t \ ?thesis" by (simp add: src_def) assume t: "arr t" show ?thesis proof (intro src_eqI) show "ide (if arr t then R.src t else null)" using t sources_closed inclusion R.src_in_sources by auto show "con (if arr t then R.src t else null) t" using t con_char by (metis (full_types) R.con_sym R.in_sourcesE R.src_in_sources \ide (if arr t then R.src t else null)\ arr_char ide_char inclusion) qed qed end text \ Here we justify the terminology ``normal sub-RTS'', which was introduced earlier, by showing that a normal sub-RTS really is a sub-RTS. \ lemma (in normal_sub_rts) is_sub_rts: shows "sub_rts resid (\t. t \ \)" using elements_are_arr ide_closed apply unfold_locales apply auto[2] by (meson R.con_imp_coinitial R.con_sym forward_stable) end